src/HOL/Hyperreal/MacLaurin_lemmas.ML
author obua
Tue, 11 May 2004 20:11:08 +0200
changeset 14738 83f1a514dcb4
child 15047 fa62de5862b9
permissions -rw-r--r--
changes made due to new Ring_and_Field theory
obua@14738
     1
(*  Title       : MacLaurin.thy
obua@14738
     2
    Author      : Jacques D. Fleuriot
obua@14738
     3
    Copyright   : 2001 University of Edinburgh
obua@14738
     4
    Description : MacLaurin series
obua@14738
     5
*)
obua@14738
     6
obua@14738
     7
Goal "sumr 0 n (%m. f (m + k)) = sumr 0 (n + k) f - sumr 0 k f";
obua@14738
     8
by (induct_tac "n" 1);
obua@14738
     9
by Auto_tac;
obua@14738
    10
qed "sumr_offset";
obua@14738
    11
obua@14738
    12
Goal "ALL f. sumr 0 n (%m. f (m + k)) = sumr 0 (n + k) f - sumr 0 k f";
obua@14738
    13
by (induct_tac "n" 1);
obua@14738
    14
by Auto_tac;
obua@14738
    15
qed "sumr_offset2";
obua@14738
    16
obua@14738
    17
Goal "sumr 0 (n + k) f = sumr 0 n (%m. f (m + k)) + sumr 0 k f";
obua@14738
    18
by (simp_tac (simpset() addsimps [sumr_offset]) 1);
obua@14738
    19
qed "sumr_offset3";
obua@14738
    20
obua@14738
    21
Goal "ALL n f. sumr 0 (n + k) f = sumr 0 n (%m. f (m + k)) + sumr 0 k f";
obua@14738
    22
by (simp_tac (simpset() addsimps [sumr_offset]) 1);
obua@14738
    23
qed "sumr_offset4";
obua@14738
    24
obua@14738
    25
Goal "0 < n ==> \
obua@14738
    26
\     sumr (Suc 0) (Suc n) (%n. (if even(n) then 0 else \
obua@14738
    27
\            ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n) = \
obua@14738
    28
\     sumr 0 (Suc n) (%n. (if even(n) then 0 else \
obua@14738
    29
\            ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n)";
obua@14738
    30
by (res_inst_tac [("n1","1")] (sumr_split_add RS subst) 1);
obua@14738
    31
by Auto_tac;
obua@14738
    32
qed "sumr_from_1_from_0";
obua@14738
    33
obua@14738
    34
(*---------------------------------------------------------------------------*)
obua@14738
    35
(* Maclaurin's theorem with Lagrange form of remainder                       *)
obua@14738
    36
(*---------------------------------------------------------------------------*)
obua@14738
    37
obua@14738
    38
(* Annoying: Proof is now even longer due mostly to 
obua@14738
    39
   change in behaviour of simplifier  since Isabelle99 *)
obua@14738
    40
Goal " [| 0 < h; 0 < n; diff 0 = f; \
obua@14738
    41
\      ALL m t. \
obua@14738
    42
\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t |] \
obua@14738
    43
\   ==> EX t. 0 < t & \
obua@14738
    44
\             t < h & \
obua@14738
    45
\             f h = \
obua@14738
    46
\             sumr 0 n (%m. (diff m 0 / real (fact m)) * h ^ m) + \
obua@14738
    47
\             (diff n t / real (fact n)) * h ^ n";
obua@14738
    48
by (case_tac "n = 0" 1);
obua@14738
    49
by (Force_tac 1);
obua@14738
    50
by (dtac not0_implies_Suc 1);
obua@14738
    51
by (etac exE 1);
obua@14738
    52
by (subgoal_tac 
obua@14738
    53
     "EX B. f h = sumr 0 n (%m. (diff m 0 / real (fact m)) * (h ^ m)) \
obua@14738
    54
\                  + (B * ((h ^ n) / real (fact n)))" 1);
obua@14738
    55
obua@14738
    56
by (simp_tac (HOL_ss addsimps [real_add_commute, real_divide_def,
obua@14738
    57
    ARITH_PROVE "(x = z + (y::real)) = (x - y = z)"]) 2);
obua@14738
    58
by (res_inst_tac 
obua@14738
    59
  [("x","(f(h) - sumr 0 n (%m. (diff(m)(0) / real (fact m)) * (h ^ m))) \
obua@14738
    60
\        * real (fact n) / (h ^ n)")] exI 2);
obua@14738
    61
by (simp_tac (HOL_ss addsimps [real_mult_assoc,real_divide_def]) 2);
obua@14738
    62
 by (rtac (CLAIM "x = (1::real) ==>  a = a * (x::real)") 2);
obua@14738
    63
by (asm_simp_tac (HOL_ss addsimps 
obua@14738
    64
    [CLAIM "(a::real) * (b * (c * d)) = (d * a) * (b * c)"]
obua@14738
    65
     delsimps [realpow_Suc]) 2);
obua@14738
    66
by (stac left_inverse 2);
obua@14738
    67
by (stac left_inverse 3);
obua@14738
    68
by (rtac (real_not_refl2 RS not_sym) 2);
obua@14738
    69
by (etac zero_less_power 2);
obua@14738
    70
by (rtac real_of_nat_fact_not_zero 2);
obua@14738
    71
by (Simp_tac 2);
obua@14738
    72
by (etac exE 1);
obua@14738
    73
by (cut_inst_tac [("b","%t. f t - \
obua@14738
    74
\      (sumr 0 n (%m. (diff m 0 / real (fact m)) * (t ^ m)) + \
obua@14738
    75
\                       (B * ((t ^ n) / real (fact n))))")] 
obua@14738
    76
    (CLAIM "EX g. g = b") 1);
obua@14738
    77
by (etac exE 1);
obua@14738
    78
by (subgoal_tac "g 0 = 0 & g h =0" 1);
obua@14738
    79
by (asm_simp_tac (simpset() addsimps 
obua@14738
    80
    [ARITH_PROVE "(x - y = z) = (x = z + (y::real))"]
obua@14738
    81
    delsimps [sumr_Suc]) 2);
obua@14738
    82
by (cut_inst_tac [("n","m"),("k","1")] sumr_offset2 2);
obua@14738
    83
by (asm_full_simp_tac (simpset() addsimps 
obua@14738
    84
    [ARITH_PROVE "(x = y - z) = (y = x + (z::real))"]
obua@14738
    85
    delsimps [sumr_Suc]) 2);
obua@14738
    86
by (cut_inst_tac [("b","%m t. diff m t - \
obua@14738
    87
\      (sumr 0 (n - m) (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) \
obua@14738
    88
\       + (B * ((t ^ (n - m)) / real (fact(n - m)))))")] 
obua@14738
    89
    (CLAIM "EX difg. difg = b") 1);
obua@14738
    90
by (etac exE 1);
obua@14738
    91
by (subgoal_tac "difg 0 = g" 1);
obua@14738
    92
by (asm_simp_tac (simpset() delsimps [realpow_Suc,fact_Suc]) 2);
obua@14738
    93
by (subgoal_tac "ALL m t. m < n & 0 <= t & t <= h --> \
obua@14738
    94
\                   DERIV (difg m) t :> difg (Suc m) t" 1);
obua@14738
    95
by (Clarify_tac 2);
obua@14738
    96
by (rtac DERIV_diff 2);
obua@14738
    97
by (Asm_simp_tac 2);
obua@14738
    98
by DERIV_tac;
obua@14738
    99
by DERIV_tac;
obua@14738
   100
by (rtac lemma_DERIV_subst 3);
obua@14738
   101
by (rtac DERIV_quotient 3);
obua@14738
   102
by (rtac DERIV_const 4);
obua@14738
   103
by (rtac DERIV_pow 3);
obua@14738
   104
by (asm_simp_tac (simpset() addsimps [inverse_mult_distrib,
obua@14738
   105
    CLAIM_SIMP "(a::real) * b * c * (d * e) = a * b * (c * d) * e" 
obua@14738
   106
    mult_ac,fact_diff_Suc]) 4);
obua@14738
   107
by (Asm_simp_tac 3);
obua@14738
   108
by (forw_inst_tac [("m","ma")] less_add_one 2);
obua@14738
   109
by (Clarify_tac 2);
obua@14738
   110
by (asm_simp_tac (simpset() addsimps 
obua@14738
   111
    [CLAIM "Suc m = ma + d + 1 ==> m - ma = d"]
obua@14738
   112
    delsimps [sumr_Suc]) 2);
obua@14738
   113
by (asm_simp_tac (simpset() addsimps [(simplify (simpset() delsimps [sumr_Suc]) 
obua@14738
   114
          (read_instantiate [("k","1")] sumr_offset4))] 
obua@14738
   115
    delsimps [sumr_Suc,fact_Suc,realpow_Suc]) 2);
obua@14738
   116
by (rtac lemma_DERIV_subst 2);
obua@14738
   117
by (rtac DERIV_add 2);
obua@14738
   118
by (rtac DERIV_const 3);
obua@14738
   119
by (rtac DERIV_sumr 2);
obua@14738
   120
by (Clarify_tac 2);
obua@14738
   121
by (Simp_tac 3);
obua@14738
   122
by (simp_tac (simpset() addsimps [real_divide_def,real_mult_assoc] 
obua@14738
   123
    delsimps [fact_Suc,realpow_Suc]) 2);
obua@14738
   124
by (rtac DERIV_cmult 2);
obua@14738
   125
by (rtac lemma_DERIV_subst 2);
obua@14738
   126
by DERIV_tac;
obua@14738
   127
by (stac fact_Suc 2);
obua@14738
   128
by (stac real_of_nat_mult 2);
obua@14738
   129
by (simp_tac (simpset() addsimps [inverse_mult_distrib] @
obua@14738
   130
    mult_ac) 2);
obua@14738
   131
by (subgoal_tac "ALL ma. ma < n --> \
obua@14738
   132
\        (EX t. 0 < t & t < h & difg (Suc ma) t = 0)" 1);
obua@14738
   133
by (rotate_tac 11 1);
obua@14738
   134
by (dres_inst_tac [("x","m")] spec 1);
obua@14738
   135
by (etac impE 1);
obua@14738
   136
by (Asm_simp_tac 1);
obua@14738
   137
by (etac exE 1);
obua@14738
   138
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   139
by (asm_full_simp_tac (simpset() addsimps 
obua@14738
   140
     [ARITH_PROVE "(x - y = 0) = (y = (x::real))"] 
obua@14738
   141
      delsimps [realpow_Suc,fact_Suc]) 1);
obua@14738
   142
by (subgoal_tac "ALL m. m < n --> difg m 0 = 0" 1);
obua@14738
   143
by (Clarify_tac 2);
obua@14738
   144
by (Asm_simp_tac 2);
obua@14738
   145
by (forw_inst_tac [("m","ma")] less_add_one 2);
obua@14738
   146
by (Clarify_tac 2);
obua@14738
   147
by (asm_simp_tac (simpset() delsimps [sumr_Suc]) 2);
obua@14738
   148
by (asm_simp_tac (simpset() addsimps [(simplify (simpset() delsimps [sumr_Suc]) 
obua@14738
   149
          (read_instantiate [("k","1")] sumr_offset4))] 
obua@14738
   150
    delsimps [sumr_Suc,fact_Suc,realpow_Suc]) 2);
obua@14738
   151
by (subgoal_tac "ALL m. m < n --> (EX t. 0 < t & t < h & \
obua@14738
   152
\                DERIV (difg m) t :> 0)" 1);
obua@14738
   153
by (rtac allI 1 THEN rtac impI 1);
obua@14738
   154
by (rotate_tac 12 1);
obua@14738
   155
by (dres_inst_tac [("x","ma")] spec 1);
obua@14738
   156
by (etac impE 1 THEN assume_tac 1);
obua@14738
   157
by (etac exE 1);
obua@14738
   158
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   159
(* do some tidying up *)
obua@14738
   160
by (ALLGOALS(thin_tac "difg = \
obua@14738
   161
\          (%m t. diff m t - \
obua@14738
   162
\                 (sumr 0 (n - m) \
obua@14738
   163
\                   (%p. diff (m + p) 0 / real (fact p) * t ^ p) + \
obua@14738
   164
\                  B * (t ^ (n - m) / real (fact (n - m)))))"));
obua@14738
   165
by (ALLGOALS(thin_tac "g = \
obua@14738
   166
\          (%t. f t - \
obua@14738
   167
\               (sumr 0 n (%m. diff m 0 / real  (fact m) * t ^ m) + \
obua@14738
   168
\                B * (t ^ n / real (fact n))))"));
obua@14738
   169
by (ALLGOALS(thin_tac "f h = \
obua@14738
   170
\          sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
obua@14738
   171
\          B * (h ^ n / real (fact n))"));
obua@14738
   172
(* back to business *)
obua@14738
   173
by (Asm_simp_tac 1);
obua@14738
   174
by (rtac DERIV_unique 1);
obua@14738
   175
by (Blast_tac 2);
obua@14738
   176
by (Force_tac 1);
obua@14738
   177
by (rtac allI 1 THEN induct_tac "ma" 1);
obua@14738
   178
by (rtac impI 1 THEN rtac Rolle 1);
obua@14738
   179
by (assume_tac 1);
obua@14738
   180
by (Asm_full_simp_tac 1);
obua@14738
   181
by (Asm_full_simp_tac 1);
obua@14738
   182
by (subgoal_tac "ALL t. 0 <= t & t <= h --> g differentiable t" 1);
obua@14738
   183
by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
obua@14738
   184
by (blast_tac (claset() addDs [DERIV_isCont]) 1);
obua@14738
   185
by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
obua@14738
   186
by (Clarify_tac 1);
obua@14738
   187
by (res_inst_tac [("x","difg (Suc 0) t")] exI 1);
obua@14738
   188
by (Force_tac 1);
obua@14738
   189
by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);
obua@14738
   190
by (Clarify_tac 1);
obua@14738
   191
by (res_inst_tac [("x","difg (Suc 0) x")] exI 1);
obua@14738
   192
by (Force_tac 1);
obua@14738
   193
by (Step_tac 1);
obua@14738
   194
by (Force_tac 1);
obua@14738
   195
by (subgoal_tac "EX ta. 0 < ta & ta < t & \
obua@14738
   196
\                DERIV difg (Suc n) ta :> 0" 1);
obua@14738
   197
by (rtac Rolle 2 THEN assume_tac 2);
obua@14738
   198
by (Asm_full_simp_tac 2);
obua@14738
   199
by (rotate_tac 2 2);
obua@14738
   200
by (dres_inst_tac [("x","n")] spec 2);
obua@14738
   201
by (ftac (ARITH_PROVE "n < m  ==> n < Suc m") 2);
obua@14738
   202
by (rtac DERIV_unique 2);
obua@14738
   203
by (assume_tac 3);
obua@14738
   204
by (Force_tac 2);
obua@14738
   205
by (subgoal_tac 
obua@14738
   206
    "ALL ta. 0 <= ta & ta <= t --> (difg (Suc n)) differentiable ta" 2);
obua@14738
   207
by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
obua@14738
   208
by (blast_tac (claset() addSDs [DERIV_isCont]) 2);
obua@14738
   209
by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
obua@14738
   210
by (Clarify_tac 2);
obua@14738
   211
by (res_inst_tac [("x","difg (Suc (Suc n)) ta")] exI 2);
obua@14738
   212
by (Force_tac 2);
obua@14738
   213
by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 2);
obua@14738
   214
by (Clarify_tac 2);
obua@14738
   215
by (res_inst_tac [("x","difg (Suc (Suc n)) x")] exI 2);
obua@14738
   216
by (Force_tac 2);
obua@14738
   217
by (Step_tac 1);
obua@14738
   218
by (res_inst_tac [("x","ta")] exI 1);
obua@14738
   219
by (Force_tac 1);
obua@14738
   220
qed "Maclaurin";
obua@14738
   221
obua@14738
   222
Goal "0 < h & 0 < n & diff 0 = f & \
obua@14738
   223
\      (ALL m t. \
obua@14738
   224
\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t) \
obua@14738
   225
\   --> (EX t. 0 < t & \
obua@14738
   226
\             t < h & \
obua@14738
   227
\             f h = \
obua@14738
   228
\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
obua@14738
   229
\             diff n t / real (fact n) * h ^ n)";
obua@14738
   230
by (blast_tac (claset() addIs [Maclaurin]) 1);
obua@14738
   231
qed "Maclaurin_objl";
obua@14738
   232
obua@14738
   233
Goal " [| 0 < h; diff 0 = f; \
obua@14738
   234
\      ALL m t. \
obua@14738
   235
\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t |] \
obua@14738
   236
\   ==> EX t. 0 < t & \
obua@14738
   237
\             t <= h & \
obua@14738
   238
\             f h = \
obua@14738
   239
\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
obua@14738
   240
\             diff n t / real (fact n) * h ^ n";
obua@14738
   241
by (case_tac "n" 1);
obua@14738
   242
by Auto_tac;
obua@14738
   243
by (dtac Maclaurin 1 THEN Auto_tac);
obua@14738
   244
qed "Maclaurin2";
obua@14738
   245
obua@14738
   246
Goal "0 < h & diff 0 = f & \
obua@14738
   247
\      (ALL m t. \
obua@14738
   248
\         m < n & 0 <= t & t <= h --> DERIV (diff m) t :> diff (Suc m) t) \
obua@14738
   249
\   --> (EX t. 0 < t & \
obua@14738
   250
\             t <= h & \
obua@14738
   251
\             f h = \
obua@14738
   252
\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
obua@14738
   253
\             diff n t / real (fact n) * h ^ n)";
obua@14738
   254
by (blast_tac (claset() addIs [Maclaurin2]) 1);
obua@14738
   255
qed "Maclaurin2_objl";
obua@14738
   256
obua@14738
   257
Goal " [| h < 0; 0 < n; diff 0 = f; \
obua@14738
   258
\      ALL m t. \
obua@14738
   259
\         m < n & h <= t & t <= 0 --> DERIV (diff m) t :> diff (Suc m) t |] \
obua@14738
   260
\   ==> EX t. h < t & \
obua@14738
   261
\             t < 0 & \
obua@14738
   262
\             f h = \
obua@14738
   263
\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
obua@14738
   264
\             diff n t / real (fact n) * h ^ n";
obua@14738
   265
by (cut_inst_tac [("f","%x. f (-x)"),
obua@14738
   266
                 ("diff","%n x. ((- 1) ^ n) * diff n (-x)"),
obua@14738
   267
                 ("h","-h"),("n","n")] Maclaurin_objl 1);
obua@14738
   268
by (Asm_full_simp_tac 1);
obua@14738
   269
by (etac impE 1 THEN Step_tac 1);
obua@14738
   270
by (stac minus_mult_right 1);
obua@14738
   271
by (rtac DERIV_cmult 1);
obua@14738
   272
by (rtac lemma_DERIV_subst 1);
obua@14738
   273
by (rtac (read_instantiate [("g","uminus")] DERIV_chain2) 1);
obua@14738
   274
by (rtac DERIV_minus 2 THEN rtac DERIV_Id 2);
obua@14738
   275
by (Force_tac 2);
obua@14738
   276
by (Force_tac 1);
obua@14738
   277
by (res_inst_tac [("x","-t")] exI 1);
obua@14738
   278
by Auto_tac;
obua@14738
   279
by (rtac (CLAIM "[| x = x'; y = y' |] ==> x + y = x' + (y'::real)") 1);
obua@14738
   280
by (rtac sumr_fun_eq 1);
obua@14738
   281
by (Asm_full_simp_tac 1);
obua@14738
   282
by (auto_tac (claset(),simpset() addsimps [real_divide_def,
obua@14738
   283
    CLAIM "((a * b) * c) * d = (b * c) * (a * (d::real))",
obua@14738
   284
    power_mult_distrib RS sym]));
obua@14738
   285
qed "Maclaurin_minus";
obua@14738
   286
obua@14738
   287
Goal "(h < 0 & 0 < n & diff 0 = f & \
obua@14738
   288
\      (ALL m t. \
obua@14738
   289
\         m < n & h <= t & t <= 0 --> DERIV (diff m) t :> diff (Suc m) t))\
obua@14738
   290
\   --> (EX t. h < t & \
obua@14738
   291
\             t < 0 & \
obua@14738
   292
\             f h = \
obua@14738
   293
\             sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + \
obua@14738
   294
\             diff n t / real (fact n) * h ^ n)";
obua@14738
   295
by (blast_tac (claset() addIs [Maclaurin_minus]) 1);
obua@14738
   296
qed "Maclaurin_minus_objl";
obua@14738
   297
obua@14738
   298
(* ------------------------------------------------------------------------- *)
obua@14738
   299
(* More convenient "bidirectional" version.                                  *)
obua@14738
   300
(* ------------------------------------------------------------------------- *)
obua@14738
   301
obua@14738
   302
(* not good for PVS sin_approx, cos_approx *)
obua@14738
   303
Goal " [| diff 0 = f; \
obua@14738
   304
\      ALL m t. \
obua@14738
   305
\         m < n & abs t <= abs x --> DERIV (diff m) t :> diff (Suc m) t |] \
obua@14738
   306
\   ==> EX t. abs t <= abs x & \
obua@14738
   307
\             f x = \
obua@14738
   308
\             sumr 0 n (%m. diff m 0 / real (fact m) * x ^ m) + \
obua@14738
   309
\             diff n t / real (fact n) * x ^ n";
obua@14738
   310
by (case_tac "n = 0" 1);
obua@14738
   311
by (Force_tac 1);
obua@14738
   312
by (case_tac "x = 0" 1);
obua@14738
   313
by (res_inst_tac [("x","0")] exI 1);
obua@14738
   314
by (Asm_full_simp_tac 1);
obua@14738
   315
by (res_inst_tac [("P","0 < n")] impE 1);
obua@14738
   316
by (assume_tac 2 THEN assume_tac 2);
obua@14738
   317
by (induct_tac "n" 1);
obua@14738
   318
by (Simp_tac 1);
obua@14738
   319
by Auto_tac;
obua@14738
   320
by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
obua@14738
   321
by Auto_tac;
obua@14738
   322
by (cut_inst_tac [("f","diff 0"),
obua@14738
   323
                 ("diff","diff"),
obua@14738
   324
                 ("h","x"),("n","n")] Maclaurin_objl 2);
obua@14738
   325
by (Step_tac 2);
obua@14738
   326
by (blast_tac (claset() addDs 
obua@14738
   327
    [ARITH_PROVE "[|(0::real) <= t;t <= x |] ==> abs t <= abs x"]) 2);
obua@14738
   328
by (res_inst_tac [("x","t")] exI 2);
obua@14738
   329
by (force_tac (claset() addIs 
obua@14738
   330
    [ARITH_PROVE "[| 0 < t; (t::real) < x|] ==> abs t <= abs x"],simpset()) 2);
obua@14738
   331
by (cut_inst_tac [("f","diff 0"),
obua@14738
   332
                 ("diff","diff"),
obua@14738
   333
                 ("h","x"),("n","n")] Maclaurin_minus_objl 1);
obua@14738
   334
by (Step_tac 1);
obua@14738
   335
by (blast_tac (claset() addDs 
obua@14738
   336
    [ARITH_PROVE "[|x <= t;t <= (0::real) |] ==> abs t <= abs x"]) 1);
obua@14738
   337
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   338
by (force_tac (claset() addIs 
obua@14738
   339
    [ARITH_PROVE "[| x < t; (t::real) < 0|] ==> abs t <= abs x"],simpset()) 1);
obua@14738
   340
qed "Maclaurin_bi_le";
obua@14738
   341
obua@14738
   342
Goal "[| diff 0 = f; \
obua@14738
   343
\        ALL m x. DERIV (diff m) x :> diff(Suc m) x; \ 
obua@14738
   344
\       x ~= 0; 0 < n \
obua@14738
   345
\     |] ==> EX t. 0 < abs t & abs t < abs x & \
obua@14738
   346
\              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
obua@14738
   347
\                    (diff n t / real (fact n)) * x ^ n";
obua@14738
   348
by (res_inst_tac [("x","x"),("y","0")] linorder_cases 1);
obua@14738
   349
by (Blast_tac 2);
obua@14738
   350
by (dtac Maclaurin_minus 1);
obua@14738
   351
by (dtac Maclaurin 5);
obua@14738
   352
by (TRYALL(assume_tac));
obua@14738
   353
by (Blast_tac 1);
obua@14738
   354
by (Blast_tac 2);
obua@14738
   355
by (Step_tac 1);
obua@14738
   356
by (ALLGOALS(res_inst_tac [("x","t")] exI));
obua@14738
   357
by (Step_tac 1);
obua@14738
   358
by (ALLGOALS(arith_tac));
obua@14738
   359
qed "Maclaurin_all_lt";
obua@14738
   360
obua@14738
   361
Goal "diff 0 = f & \
obua@14738
   362
\     (ALL m x. DERIV (diff m) x :> diff(Suc m) x) & \
obua@14738
   363
\     x ~= 0 & 0 < n \
obua@14738
   364
\     --> (EX t. 0 < abs t & abs t < abs x & \
obua@14738
   365
\              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
obua@14738
   366
\                    (diff n t / real (fact n)) * x ^ n)";
obua@14738
   367
by (blast_tac (claset() addIs [Maclaurin_all_lt]) 1);
obua@14738
   368
qed "Maclaurin_all_lt_objl";
obua@14738
   369
obua@14738
   370
Goal "x = (0::real)  \
obua@14738
   371
\     ==> 0 < n --> \
obua@14738
   372
\         sumr 0 n (%m. (diff m (0::real) / real (fact m)) * x ^ m) = \
obua@14738
   373
\         diff 0 0";
obua@14738
   374
by (Asm_simp_tac 1);
obua@14738
   375
by (induct_tac "n" 1);
obua@14738
   376
by Auto_tac; 
obua@14738
   377
qed_spec_mp "Maclaurin_zero";
obua@14738
   378
obua@14738
   379
Goal "[| diff 0 = f; \
obua@14738
   380
\       ALL m x. DERIV (diff m) x :> diff (Suc m) x \
obua@14738
   381
\     |] ==> EX t. abs t <= abs x & \
obua@14738
   382
\             f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
obua@14738
   383
\                   (diff n t / real (fact n)) * x ^ n";
obua@14738
   384
by (cut_inst_tac [("n","n"),("m","0")] 
obua@14738
   385
       (ARITH_PROVE "n <= m | m < (n::nat)") 1);
obua@14738
   386
by (etac disjE 1);
obua@14738
   387
by (Force_tac 1);
obua@14738
   388
by (case_tac "x = 0" 1);
obua@14738
   389
by (forw_inst_tac [("diff","diff"),("n","n")] Maclaurin_zero 1);
obua@14738
   390
by (assume_tac 1);
obua@14738
   391
by (dtac (gr_implies_not0 RS  not0_implies_Suc) 1);
obua@14738
   392
by (res_inst_tac [("x","0")] exI 1);
obua@14738
   393
by (Force_tac 1);
obua@14738
   394
by (forw_inst_tac [("diff","diff"),("n","n")] Maclaurin_all_lt 1);
obua@14738
   395
by (TRYALL(assume_tac));
obua@14738
   396
by (Step_tac 1);
obua@14738
   397
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   398
by Auto_tac;
obua@14738
   399
qed "Maclaurin_all_le";
obua@14738
   400
obua@14738
   401
Goal "diff 0 = f & \
obua@14738
   402
\     (ALL m x. DERIV (diff m) x :> diff (Suc m) x)  \
obua@14738
   403
\     --> (EX t. abs t <= abs x & \
obua@14738
   404
\             f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) + \
obua@14738
   405
\                   (diff n t / real (fact n)) * x ^ n)";
obua@14738
   406
by (blast_tac (claset() addIs [Maclaurin_all_le]) 1);
obua@14738
   407
qed "Maclaurin_all_le_objl";
obua@14738
   408
obua@14738
   409
(* ------------------------------------------------------------------------- *)
obua@14738
   410
(* Version for exp.                                                          *)
obua@14738
   411
(* ------------------------------------------------------------------------- *)
obua@14738
   412
obua@14738
   413
Goal "[| x ~= 0; 0 < n |] \
obua@14738
   414
\     ==> (EX t. 0 < abs t & \
obua@14738
   415
\               abs t < abs x & \
obua@14738
   416
\               exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + \
obua@14738
   417
\                       (exp t / real (fact n)) * x ^ n)";
obua@14738
   418
by (cut_inst_tac [("diff","%n. exp"),("f","exp"),("x","x"),("n","n")] 
obua@14738
   419
    Maclaurin_all_lt_objl 1);
obua@14738
   420
by Auto_tac;
obua@14738
   421
qed "Maclaurin_exp_lt";
obua@14738
   422
obua@14738
   423
Goal "EX t. abs t <= abs x & \
obua@14738
   424
\           exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) + \
obua@14738
   425
\                      (exp t / real (fact n)) * x ^ n";
obua@14738
   426
by (cut_inst_tac [("diff","%n. exp"),("f","exp"),("x","x"),("n","n")] 
obua@14738
   427
    Maclaurin_all_le_objl 1);
obua@14738
   428
by Auto_tac;
obua@14738
   429
qed "Maclaurin_exp_le";
obua@14738
   430
obua@14738
   431
(* ------------------------------------------------------------------------- *)
obua@14738
   432
(* Version for sin function                                                  *)
obua@14738
   433
(* ------------------------------------------------------------------------- *)
obua@14738
   434
obua@14738
   435
Goal "[| a < b; ALL x. a <= x & x <= b --> DERIV f x :> f'(x) |] \
obua@14738
   436
\     ==> EX z. a < z & z < b & (f b - f a = (b - a) * f'(z))";
obua@14738
   437
by (dtac MVT 1);
obua@14738
   438
by (blast_tac (claset() addIs [DERIV_isCont]) 1);
obua@14738
   439
by (force_tac (claset() addDs [order_less_imp_le],
obua@14738
   440
    simpset() addsimps [differentiable_def]) 1);
obua@14738
   441
by (blast_tac (claset() addDs [DERIV_unique,order_less_imp_le]) 1);
obua@14738
   442
qed "MVT2";
obua@14738
   443
obua@14738
   444
Goal "d < (4::nat) ==> d = 0 | d = 1 | d = 2 | d = 3";
obua@14738
   445
by (case_tac "d" 1 THEN Auto_tac);
obua@14738
   446
qed "lemma_exhaust_less_4";
obua@14738
   447
obua@14738
   448
bind_thm ("real_mult_le_lemma",
obua@14738
   449
          simplify (simpset()) (inst "b" "1" mult_right_mono));
obua@14738
   450
obua@14738
   451
obua@14738
   452
Goal "0 < n --> Suc (Suc (2 * n - 2)) = 2*n";
obua@14738
   453
by (induct_tac "n" 1);
obua@14738
   454
by Auto_tac;
obua@14738
   455
qed_spec_mp "Suc_Suc_mult_two_diff_two";
obua@14738
   456
Addsimps [Suc_Suc_mult_two_diff_two];
obua@14738
   457
obua@14738
   458
Goal "0 < n --> Suc (Suc (4*n - 2)) = 4*n";
obua@14738
   459
by (induct_tac "n" 1);
obua@14738
   460
by Auto_tac;
obua@14738
   461
qed_spec_mp "lemma_Suc_Suc_4n_diff_2";
obua@14738
   462
Addsimps [lemma_Suc_Suc_4n_diff_2];
obua@14738
   463
obua@14738
   464
Goal "0 < n --> Suc (2 * n - 1) = 2*n";
obua@14738
   465
by (induct_tac "n" 1);
obua@14738
   466
by Auto_tac;
obua@14738
   467
qed_spec_mp "Suc_mult_two_diff_one";
obua@14738
   468
Addsimps [Suc_mult_two_diff_one];
obua@14738
   469
obua@14738
   470
Goal "EX t. sin x = \
obua@14738
   471
\      (sumr 0 n (%m. (if even m then 0 \
obua@14738
   472
\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
obua@14738
   473
\                      x ^ m)) \
obua@14738
   474
\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
obua@14738
   475
by (cut_inst_tac [("f","sin"),("n","n"),("x","x"),
obua@14738
   476
       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
obua@14738
   477
       Maclaurin_all_lt_objl 1);
obua@14738
   478
by (Safe_tac);
obua@14738
   479
by (Simp_tac 1);
obua@14738
   480
by (Simp_tac 1);
obua@14738
   481
by (case_tac "n" 1);
obua@14738
   482
by (Clarify_tac 1); 
obua@14738
   483
by (Asm_full_simp_tac 1);
obua@14738
   484
by (dres_inst_tac [("x","0")] spec 1 THEN Asm_full_simp_tac 1);
obua@14738
   485
by (Asm_full_simp_tac 1);
obua@14738
   486
by (rtac ccontr 1);
obua@14738
   487
by (Asm_full_simp_tac 1);
obua@14738
   488
by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
obua@14738
   489
by (dtac ssubst 1 THEN assume_tac 2);
obua@14738
   490
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   491
by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
obua@14738
   492
by (rtac sumr_fun_eq 1);
obua@14738
   493
by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
obua@14738
   494
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
obua@14738
   495
(*Could sin_zero_iff help?*)
obua@14738
   496
qed "Maclaurin_sin_expansion";
obua@14738
   497
obua@14738
   498
Goal "EX t. abs t <= abs x &  \
obua@14738
   499
\      sin x = \
obua@14738
   500
\      (sumr 0 n (%m. (if even m then 0 \
obua@14738
   501
\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
obua@14738
   502
\                      x ^ m)) \
obua@14738
   503
\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
obua@14738
   504
obua@14738
   505
by (cut_inst_tac [("f","sin"),("n","n"),("x","x"),
obua@14738
   506
       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
obua@14738
   507
       Maclaurin_all_lt_objl 1);
obua@14738
   508
by (Step_tac 1);
obua@14738
   509
by (Simp_tac 1);
obua@14738
   510
by (Simp_tac 1);
obua@14738
   511
by (case_tac "n" 1);
obua@14738
   512
by (Clarify_tac 1); 
obua@14738
   513
by (Asm_full_simp_tac 1);
obua@14738
   514
by (Asm_full_simp_tac 1);
obua@14738
   515
by (rtac ccontr 1);
obua@14738
   516
by (Asm_full_simp_tac 1);
obua@14738
   517
by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
obua@14738
   518
by (dtac ssubst 1 THEN assume_tac 2);
obua@14738
   519
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   520
by (rtac conjI 1);
obua@14738
   521
by (arith_tac 1);
obua@14738
   522
by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
obua@14738
   523
by (rtac sumr_fun_eq 1);
obua@14738
   524
by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
obua@14738
   525
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
obua@14738
   526
qed "Maclaurin_sin_expansion2";
obua@14738
   527
obua@14738
   528
Goal "[| 0 < n; 0 < x |] ==> \
obua@14738
   529
\      EX t. 0 < t & t < x & \
obua@14738
   530
\      sin x = \
obua@14738
   531
\      (sumr 0 n (%m. (if even m then 0 \
obua@14738
   532
\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
obua@14738
   533
\                      x ^ m)) \
obua@14738
   534
\     + ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)";
obua@14738
   535
by (cut_inst_tac [("f","sin"),("n","n"),("h","x"),
obua@14738
   536
       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
obua@14738
   537
       Maclaurin_objl 1);
obua@14738
   538
by (Step_tac 1);
obua@14738
   539
by (Asm_full_simp_tac 1);
obua@14738
   540
by (Simp_tac 1);
obua@14738
   541
by (dtac ssubst 1 THEN assume_tac 2);
obua@14738
   542
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   543
by (rtac conjI 1 THEN rtac conjI 2);
obua@14738
   544
by (assume_tac 1 THEN assume_tac 1);
obua@14738
   545
by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
obua@14738
   546
by (rtac sumr_fun_eq 1);
obua@14738
   547
by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
obua@14738
   548
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
obua@14738
   549
qed "Maclaurin_sin_expansion3";
obua@14738
   550
obua@14738
   551
Goal "0 < x ==> \
obua@14738
   552
\      EX t. 0 < t & t <= x & \
obua@14738
   553
\      sin x = \
obua@14738
   554
\      (sumr 0 n (%m. (if even m then 0 \
obua@14738
   555
\                      else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) * \
obua@14738
   556
\                      x ^ m)) \
obua@14738
   557
\     + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
obua@14738
   558
by (cut_inst_tac [("f","sin"),("n","n"),("h","x"),
obua@14738
   559
       ("diff","%n x. sin(x + 1/2*real (n)*pi)")] 
obua@14738
   560
       Maclaurin2_objl 1);
obua@14738
   561
by (Step_tac 1);
obua@14738
   562
by (Asm_full_simp_tac 1);
obua@14738
   563
by (Simp_tac 1);
obua@14738
   564
by (dtac ssubst 1 THEN assume_tac 2);
obua@14738
   565
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   566
by (rtac conjI 1 THEN rtac conjI 2);
obua@14738
   567
by (assume_tac 1 THEN assume_tac 1);
obua@14738
   568
by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
obua@14738
   569
by (rtac sumr_fun_eq 1);
obua@14738
   570
by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
obua@14738
   571
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex] delsimps [fact_Suc,realpow_Suc]));
obua@14738
   572
qed "Maclaurin_sin_expansion4";
obua@14738
   573
obua@14738
   574
(*-----------------------------------------------------------------------------*)
obua@14738
   575
(* Maclaurin expansion for cos                                                 *)
obua@14738
   576
(*-----------------------------------------------------------------------------*)
obua@14738
   577
obua@14738
   578
Goal "sumr 0 (Suc n) \
obua@14738
   579
\        (%m. (if even m \
obua@14738
   580
\              then (- 1) ^ (m div 2)/(real  (fact m)) \
obua@14738
   581
\              else 0) * \
obua@14738
   582
\             0 ^ m) = 1";
obua@14738
   583
by (induct_tac "n" 1);
obua@14738
   584
by Auto_tac;
obua@14738
   585
qed "sumr_cos_zero_one";
obua@14738
   586
Addsimps [sumr_cos_zero_one];
obua@14738
   587
obua@14738
   588
Goal "EX t. abs t <= abs x & \
obua@14738
   589
\      cos x = \
obua@14738
   590
\      (sumr 0 n (%m. (if even m \
obua@14738
   591
\                      then (- 1) ^ (m div 2)/(real (fact m)) \
obua@14738
   592
\                      else 0) * \
obua@14738
   593
\                      x ^ m)) \
obua@14738
   594
\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
obua@14738
   595
by (cut_inst_tac [("f","cos"),("n","n"),("x","x"),
obua@14738
   596
       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
obua@14738
   597
       Maclaurin_all_lt_objl 1);
obua@14738
   598
by (Step_tac 1);
obua@14738
   599
by (Simp_tac 1);
obua@14738
   600
by (Simp_tac 1);
obua@14738
   601
by (case_tac "n" 1);
obua@14738
   602
by (Asm_full_simp_tac 1);
obua@14738
   603
by (asm_full_simp_tac (simpset() delsimps [sumr_Suc]) 1);
obua@14738
   604
by (rtac ccontr 1);
obua@14738
   605
by (Asm_full_simp_tac 1);
obua@14738
   606
by (dres_inst_tac [("x","x")] spec 1 THEN Asm_full_simp_tac 1);
obua@14738
   607
by (dtac ssubst 1 THEN assume_tac 2);
obua@14738
   608
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   609
by (rtac conjI 1);
obua@14738
   610
by (arith_tac 1);
obua@14738
   611
by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
obua@14738
   612
by (rtac sumr_fun_eq 1);
obua@14738
   613
by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
obua@14738
   614
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps 
obua@14738
   615
    [fact_Suc,realpow_Suc]));
obua@14738
   616
by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
obua@14738
   617
qed "Maclaurin_cos_expansion";
obua@14738
   618
obua@14738
   619
Goal "[| 0 < x; 0 < n |] ==> \
obua@14738
   620
\      EX t. 0 < t & t < x & \
obua@14738
   621
\      cos x = \
obua@14738
   622
\      (sumr 0 n (%m. (if even m \
obua@14738
   623
\                      then (- 1) ^ (m div 2)/(real (fact m)) \
obua@14738
   624
\                      else 0) * \
obua@14738
   625
\                      x ^ m)) \
obua@14738
   626
\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
obua@14738
   627
by (cut_inst_tac [("f","cos"),("n","n"),("h","x"),
obua@14738
   628
       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
obua@14738
   629
       Maclaurin_objl 1);
obua@14738
   630
by (Step_tac 1);
obua@14738
   631
by (Asm_full_simp_tac 1);
obua@14738
   632
by (Simp_tac 1);
obua@14738
   633
by (dtac ssubst 1 THEN assume_tac 2);
obua@14738
   634
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   635
by (rtac conjI 1 THEN rtac conjI 2);
obua@14738
   636
by (assume_tac 1 THEN assume_tac 1);
obua@14738
   637
by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
obua@14738
   638
by (rtac sumr_fun_eq 1);
obua@14738
   639
by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
obua@14738
   640
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps [fact_Suc,realpow_Suc]));
obua@14738
   641
by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
obua@14738
   642
qed "Maclaurin_cos_expansion2";
obua@14738
   643
obua@14738
   644
Goal "[| x < 0; 0 < n |] ==> \
obua@14738
   645
\      EX t. x < t & t < 0 & \
obua@14738
   646
\      cos x = \
obua@14738
   647
\      (sumr 0 n (%m. (if even m \
obua@14738
   648
\                      then (- 1) ^ (m div 2)/(real (fact m)) \
obua@14738
   649
\                      else 0) * \
obua@14738
   650
\                      x ^ m)) \
obua@14738
   651
\     + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)";
obua@14738
   652
by (cut_inst_tac [("f","cos"),("n","n"),("h","x"),
obua@14738
   653
       ("diff","%n x. cos(x + 1/2*real (n)*pi)")] 
obua@14738
   654
       Maclaurin_minus_objl 1);
obua@14738
   655
by (Step_tac 1);
obua@14738
   656
by (Asm_full_simp_tac 1);
obua@14738
   657
by (Simp_tac 1);
obua@14738
   658
by (dtac ssubst 1 THEN assume_tac 2);
obua@14738
   659
by (res_inst_tac [("x","t")] exI 1);
obua@14738
   660
by (rtac conjI 1 THEN rtac conjI 2);
obua@14738
   661
by (assume_tac 1 THEN assume_tac 1);
obua@14738
   662
by (rtac (CLAIM "[|x = y; x' = y'|] ==> x + x' = y + (y'::real)") 1);
obua@14738
   663
by (rtac sumr_fun_eq 1);
obua@14738
   664
by (auto_tac (claset(),simpset() addsimps [odd_Suc_mult_two_ex]));
obua@14738
   665
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex,left_distrib,cos_add]  delsimps [fact_Suc,realpow_Suc]));
obua@14738
   666
by (auto_tac (claset(),simpset() addsimps [real_mult_commute]));
obua@14738
   667
qed "Maclaurin_minus_cos_expansion";
obua@14738
   668
obua@14738
   669
(* ------------------------------------------------------------------------- *)
obua@14738
   670
(* Version for ln(1 +/- x). Where is it??                                    *)
obua@14738
   671
(* ------------------------------------------------------------------------- *)
obua@14738
   672