src/HOL/Hyperreal/MacLaurin.thy
author paulson
Wed, 28 Jul 2004 16:25:40 +0200
changeset 15081 32402f5624d1
parent 15079 2ef899e4526d
child 15131 c69542757a4d
permissions -rw-r--r--
abs notation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12224
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     1
(*  Title       : MacLaurin.thy
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     3
    Copyright   : 2001 University of Edinburgh
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     4
    Description : MacLaurin series
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
     5
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
12224
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     6
*)
02df7cbe7d25 even more theories from Jacques
paulson
parents:
diff changeset
     7
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
     8
theory MacLaurin = Log:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
     9
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    10
lemma sumr_offset: "sumr 0 n (%m. f (m+k)) = sumr 0 (n+k) f - sumr 0 k f"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    11
by (induct_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    12
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    13
lemma sumr_offset2: "\<forall>f. sumr 0 n (%m. f (m+k)) = sumr 0 (n+k) f - sumr 0 k f"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    14
by (induct_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    15
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    16
lemma sumr_offset3: "sumr 0 (n+k) f = sumr 0 n (%m. f (m+k)) + sumr 0 k f"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    17
by (simp  add: sumr_offset)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    18
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    19
lemma sumr_offset4: "\<forall>n f. sumr 0 (n+k) f = sumr 0 n (%m. f (m+k)) + sumr 0 k f"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    20
by (simp add: sumr_offset)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    21
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    22
lemma sumr_from_1_from_0: "0 < n ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    23
      sumr (Suc 0) (Suc n) (%n. (if even(n) then 0 else
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    24
             ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n) =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    25
      sumr 0 (Suc n) (%n. (if even(n) then 0 else
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    26
             ((- 1) ^ ((n - (Suc 0)) div 2))/(real (fact n))) * a ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    27
by (rule_tac n1 = 1 in sumr_split_add [THEN subst], auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    28
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    29
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    30
subsection{*Maclaurin's Theorem with Lagrange Form of Remainder*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    31
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    32
text{*This is a very long, messy proof even now that it's been broken down
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    33
into lemmas.*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    34
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    35
lemma Maclaurin_lemma:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    36
    "0 < h ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    37
     \<exists>B. f h = sumr 0 n (%m. (j m / real (fact m)) * (h^m)) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    38
               (B * ((h^n) / real(fact n)))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    39
by (rule_tac x = "(f h - sumr 0 n (%m. (j m / real (fact m)) * h^m)) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    40
                 real(fact n) / (h^n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    41
       in exI, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    42
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    43
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    44
lemma eq_diff_eq': "(x = y - z) = (y = x + (z::real))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    45
by arith
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    46
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    47
text{*A crude tactic to differentiate by proof.*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    48
ML
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    49
{*
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    50
exception DERIV_name;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    51
fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    52
|   get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    53
|   get_fun_name _ = raise DERIV_name;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    54
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    55
val deriv_rulesI = [DERIV_Id,DERIV_const,DERIV_cos,DERIV_cmult,
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    56
                    DERIV_sin, DERIV_exp, DERIV_inverse,DERIV_pow,
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    57
                    DERIV_add, DERIV_diff, DERIV_mult, DERIV_minus,
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    58
                    DERIV_inverse_fun,DERIV_quotient,DERIV_fun_pow,
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    59
                    DERIV_fun_exp,DERIV_fun_sin,DERIV_fun_cos,
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    60
                    DERIV_Id,DERIV_const,DERIV_cos];
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    61
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    62
val deriv_tac =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    63
  SUBGOAL (fn (prem,i) =>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    64
   (resolve_tac deriv_rulesI i) ORELSE
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    65
    ((rtac (read_instantiate [("f",get_fun_name prem)]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    66
                     DERIV_chain2) i) handle DERIV_name => no_tac));;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    67
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    68
val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i));
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    69
*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    70
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    71
lemma Maclaurin_lemma2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    72
      "[| \<forall>m t. m < n \<and> 0\<le>t \<and> t\<le>h \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    73
          n = Suc k;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    74
        difg =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    75
        (\<lambda>m t. diff m t -
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    76
               ((\<Sum>p = 0..<n - m. diff (m + p) 0 / real (fact p) * t ^ p) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    77
                B * (t ^ (n - m) / real (fact (n - m)))))|] ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    78
        \<forall>m t. m < n & 0 \<le> t & t \<le> h -->
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    79
                    DERIV (difg m) t :> difg (Suc m) t"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    80
apply clarify
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    81
apply (rule DERIV_diff)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    82
apply (simp (no_asm_simp))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    83
apply (tactic DERIV_tac)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    84
apply (tactic DERIV_tac)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    85
apply (rule_tac [2] lemma_DERIV_subst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    86
apply (rule_tac [2] DERIV_quotient)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    87
apply (rule_tac [3] DERIV_const)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    88
apply (rule_tac [2] DERIV_pow)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    89
  prefer 3 apply (simp add: fact_diff_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    90
 prefer 2 apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    91
apply (frule_tac m = m in less_add_one, clarify)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    92
apply (simp del: sumr_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    93
apply (insert sumr_offset4 [of 1])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    94
apply (simp del: sumr_Suc fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    95
apply (rule lemma_DERIV_subst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    96
apply (rule DERIV_add)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    97
apply (rule_tac [2] DERIV_const)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    98
apply (rule DERIV_sumr, clarify)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
    99
 prefer 2 apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   100
apply (simp (no_asm) add: divide_inverse mult_assoc del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   101
apply (rule DERIV_cmult)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   102
apply (rule lemma_DERIV_subst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   103
apply (best intro: DERIV_chain2 intro!: DERIV_intros)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   104
apply (subst fact_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   105
apply (subst real_of_nat_mult)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   106
apply (simp add: inverse_mult_distrib mult_ac)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   107
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   108
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   109
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   110
lemma Maclaurin_lemma3:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   111
     "[|\<forall>k t. k < Suc m \<and> 0\<le>t & t\<le>h \<longrightarrow> DERIV (difg k) t :> difg (Suc k) t;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   112
        \<forall>k<Suc m. difg k 0 = 0; DERIV (difg n) t :> 0;  n < m; 0 < t;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   113
        t < h|]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   114
     ==> \<exists>ta. 0 < ta & ta < t & DERIV (difg (Suc n)) ta :> 0"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   115
apply (rule Rolle, assumption, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   116
apply (drule_tac x = n and P="%k. k<Suc m --> difg k 0 = 0" in spec)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   117
apply (rule DERIV_unique)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   118
prefer 2 apply assumption
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   119
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   120
apply (subgoal_tac "\<forall>ta. 0 \<le> ta & ta \<le> t --> (difg (Suc n)) differentiable ta")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   121
apply (simp add: differentiable_def)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   122
apply (blast dest!: DERIV_isCont)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   123
apply (simp add: differentiable_def, clarify)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   124
apply (rule_tac x = "difg (Suc (Suc n)) ta" in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   125
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   126
apply (simp add: differentiable_def, clarify)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   127
apply (rule_tac x = "difg (Suc (Suc n)) x" in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   128
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   129
done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   130
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   131
lemma Maclaurin:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   132
   "[| 0 < h; 0 < n; diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   133
       \<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t |]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   134
    ==> \<exists>t. 0 < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   135
              t < h &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   136
              f h =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   137
              sumr 0 n (%m. (diff m 0 / real (fact m)) * h ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   138
              (diff n t / real (fact n)) * h ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   139
apply (case_tac "n = 0", force)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   140
apply (drule not0_implies_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   141
apply (erule exE)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   142
apply (frule_tac f=f and n=n and j="%m. diff m 0" in Maclaurin_lemma)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   143
apply (erule exE)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   144
apply (subgoal_tac "\<exists>g.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   145
     g = (%t. f t - (sumr 0 n (%m. (diff m 0 / real(fact m)) * t^m) + (B * (t^n / real(fact n)))))")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   146
 prefer 2 apply blast
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   147
apply (erule exE)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   148
apply (subgoal_tac "g 0 = 0 & g h =0")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   149
 prefer 2
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   150
 apply (simp del: sumr_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   151
 apply (cut_tac n = m and k = 1 in sumr_offset2)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   152
 apply (simp add: eq_diff_eq' del: sumr_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   153
apply (subgoal_tac "\<exists>difg. difg = (%m t. diff m t - (sumr 0 (n - m) (%p. (diff (m + p) 0 / real (fact p)) * (t ^ p)) + (B * ((t ^ (n - m)) / real (fact (n - m))))))")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   154
 prefer 2 apply blast
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   155
apply (erule exE)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   156
apply (subgoal_tac "difg 0 = g")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   157
 prefer 2 apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   158
apply (frule Maclaurin_lemma2, assumption+)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   159
apply (subgoal_tac "\<forall>ma. ma < n --> (\<exists>t. 0 < t & t < h & difg (Suc ma) t = 0) ")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   160
apply (drule_tac x = m and P="%m. m<n --> (\<exists>t. ?QQ m t)" in spec)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   161
apply (erule impE)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   162
apply (simp (no_asm_simp))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   163
apply (erule exE)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   164
apply (rule_tac x = t in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   165
apply (simp del: realpow_Suc fact_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   166
apply (subgoal_tac "\<forall>m. m < n --> difg m 0 = 0")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   167
 prefer 2
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   168
 apply clarify
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   169
 apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   170
 apply (frule_tac m = ma in less_add_one, clarify)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   171
 apply (simp del: sumr_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   172
apply (insert sumr_offset4 [of 1])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   173
apply (simp del: sumr_Suc fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   174
apply (subgoal_tac "\<forall>m. m < n --> (\<exists>t. 0 < t & t < h & DERIV (difg m) t :> 0) ")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   175
apply (rule allI, rule impI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   176
apply (drule_tac x = ma and P="%m. m<n --> (\<exists>t. ?QQ m t)" in spec)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   177
apply (erule impE, assumption)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   178
apply (erule exE)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   179
apply (rule_tac x = t in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   180
(* do some tidying up *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   181
apply (erule_tac [!] V= "difg = (%m t. diff m t - (sumr 0 (n - m) (%p. diff (m + p) 0 / real (fact p) * t ^ p) + B * (t ^ (n - m) / real (fact (n - m)))))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   182
       in thin_rl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   183
apply (erule_tac [!] V="g = (%t. f t - (sumr 0 n (%m. diff m 0 / real (fact m) * t ^ m) + B * (t ^ n / real (fact n))))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   184
       in thin_rl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   185
apply (erule_tac [!] V="f h = sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) + B * (h ^ n / real (fact n))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   186
       in thin_rl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   187
(* back to business *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   188
apply (simp (no_asm_simp))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   189
apply (rule DERIV_unique)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   190
prefer 2 apply blast
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   191
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   192
apply (rule allI, induct_tac "ma")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   193
apply (rule impI, rule Rolle, assumption, simp, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   194
apply (subgoal_tac "\<forall>t. 0 \<le> t & t \<le> h --> g differentiable t")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   195
apply (simp add: differentiable_def)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   196
apply (blast dest: DERIV_isCont)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   197
apply (simp add: differentiable_def, clarify)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   198
apply (rule_tac x = "difg (Suc 0) t" in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   199
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   200
apply (simp add: differentiable_def, clarify)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   201
apply (rule_tac x = "difg (Suc 0) x" in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   202
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   203
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   204
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   205
apply (frule Maclaurin_lemma3, assumption+, safe)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   206
apply (rule_tac x = ta in exI, force)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   207
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   208
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   209
lemma Maclaurin_objl:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   210
     "0 < h & 0 < n & diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   211
       (\<forall>m t. m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   212
    --> (\<exists>t. 0 < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   213
              t < h &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   214
              f h =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   215
              sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   216
              diff n t / real (fact n) * h ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   217
by (blast intro: Maclaurin)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   218
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   219
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   220
lemma Maclaurin2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   221
   "[| 0 < h; diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   222
       \<forall>m t.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   223
          m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t |]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   224
    ==> \<exists>t. 0 < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   225
              t \<le> h &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   226
              f h =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   227
              sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   228
              diff n t / real (fact n) * h ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   229
apply (case_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   230
apply (drule Maclaurin, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   231
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   232
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   233
lemma Maclaurin2_objl:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   234
     "0 < h & diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   235
       (\<forall>m t.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   236
          m < n & 0 \<le> t & t \<le> h --> DERIV (diff m) t :> diff (Suc m) t)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   237
    --> (\<exists>t. 0 < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   238
              t \<le> h &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   239
              f h =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   240
              sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   241
              diff n t / real (fact n) * h ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   242
by (blast intro: Maclaurin2)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   243
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   244
lemma Maclaurin_minus:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   245
   "[| h < 0; 0 < n; diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   246
       \<forall>m t. m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t |]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   247
    ==> \<exists>t. h < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   248
              t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   249
              f h =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   250
              sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   251
              diff n t / real (fact n) * h ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   252
apply (cut_tac f = "%x. f (-x)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   253
        and diff = "%n x. ((- 1) ^ n) * diff n (-x)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   254
        and h = "-h" and n = n in Maclaurin_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   255
apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   256
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   257
apply (subst minus_mult_right)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   258
apply (rule DERIV_cmult)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   259
apply (rule lemma_DERIV_subst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   260
apply (rule DERIV_chain2 [where g=uminus])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   261
apply (rule_tac [2] DERIV_minus, rule_tac [2] DERIV_Id)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   262
prefer 2 apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   263
apply force
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   264
apply (rule_tac x = "-t" in exI, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   265
apply (subgoal_tac "(\<Sum>m = 0..<n. -1 ^ m * diff m 0 * (-h)^m / real(fact m)) =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   266
                    (\<Sum>m = 0..<n. diff m 0 * h ^ m / real(fact m))")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   267
apply (rule_tac [2] sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   268
apply (auto simp add: divide_inverse power_mult_distrib [symmetric])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   269
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   270
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   271
lemma Maclaurin_minus_objl:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   272
     "(h < 0 & 0 < n & diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   273
       (\<forall>m t.
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   274
          m < n & h \<le> t & t \<le> 0 --> DERIV (diff m) t :> diff (Suc m) t))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   275
    --> (\<exists>t. h < t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   276
              t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   277
              f h =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   278
              sumr 0 n (%m. diff m 0 / real (fact m) * h ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   279
              diff n t / real (fact n) * h ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   280
by (blast intro: Maclaurin_minus)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   281
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   282
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   283
subsection{*More Convenient "Bidirectional" Version.*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   284
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   285
(* not good for PVS sin_approx, cos_approx *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   286
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   287
lemma Maclaurin_bi_le_lemma [rule_format]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   288
     "0 < n \<longrightarrow>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   289
       diff 0 0 =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   290
       (\<Sum>m = 0..<n. diff m 0 * 0 ^ m / real (fact m)) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   291
       diff n 0 * 0 ^ n / real (fact n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   292
by (induct_tac "n", auto)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   293
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   294
lemma Maclaurin_bi_le:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   295
   "[| diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   296
       \<forall>m t. m < n & abs t \<le> abs x --> DERIV (diff m) t :> diff (Suc m) t |]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   297
    ==> \<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   298
              f x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   299
              sumr 0 n (%m. diff m 0 / real (fact m) * x ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   300
              diff n t / real (fact n) * x ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   301
apply (case_tac "n = 0", force)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   302
apply (case_tac "x = 0")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   303
apply (rule_tac x = 0 in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   304
apply (force simp add: Maclaurin_bi_le_lemma)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   305
apply (cut_tac x = x and y = 0 in linorder_less_linear, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   306
txt{*Case 1, where @{term "x < 0"}*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   307
apply (cut_tac f = "diff 0" and diff = diff and h = x and n = n in Maclaurin_minus_objl, safe)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   308
apply (simp add: abs_if)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   309
apply (rule_tac x = t in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   310
apply (simp add: abs_if)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   311
txt{*Case 2, where @{term "0 < x"}*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   312
apply (cut_tac f = "diff 0" and diff = diff and h = x and n = n in Maclaurin_objl, safe)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   313
apply (simp add: abs_if)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   314
apply (rule_tac x = t in exI)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   315
apply (simp add: abs_if)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   316
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   317
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   318
lemma Maclaurin_all_lt:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   319
     "[| diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   320
         \<forall>m x. DERIV (diff m) x :> diff(Suc m) x;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   321
        x ~= 0; 0 < n
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   322
      |] ==> \<exists>t. 0 < abs t & abs t < abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   323
               f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   324
                     (diff n t / real (fact n)) * x ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   325
apply (rule_tac x = x and y = 0 in linorder_cases)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   326
prefer 2 apply blast
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   327
apply (drule_tac [2] diff=diff in Maclaurin)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   328
apply (drule_tac diff=diff in Maclaurin_minus, simp_all, safe)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   329
apply (rule_tac [!] x = t in exI, auto, arith+)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   330
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   331
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   332
lemma Maclaurin_all_lt_objl:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   333
     "diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   334
      (\<forall>m x. DERIV (diff m) x :> diff(Suc m) x) &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   335
      x ~= 0 & 0 < n
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   336
      --> (\<exists>t. 0 < abs t & abs t < abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   337
               f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   338
                     (diff n t / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   339
by (blast intro: Maclaurin_all_lt)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   340
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   341
lemma Maclaurin_zero [rule_format]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   342
     "x = (0::real)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   343
      ==> 0 < n -->
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   344
          sumr 0 n (%m. (diff m (0::real) / real (fact m)) * x ^ m) =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   345
          diff 0 0"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   346
by (induct n, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   347
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   348
lemma Maclaurin_all_le: "[| diff 0 = f;
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   349
        \<forall>m x. DERIV (diff m) x :> diff (Suc m) x
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   350
      |] ==> \<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   351
              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   352
                    (diff n t / real (fact n)) * x ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   353
apply (insert linorder_le_less_linear [of n 0])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   354
apply (erule disjE, force)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   355
apply (case_tac "x = 0")
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   356
apply (frule_tac diff = diff and n = n in Maclaurin_zero, assumption)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   357
apply (drule gr_implies_not0 [THEN not0_implies_Suc])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   358
apply (rule_tac x = 0 in exI, force)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   359
apply (frule_tac diff = diff and n = n in Maclaurin_all_lt, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   360
apply (rule_tac x = t in exI, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   361
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   362
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   363
lemma Maclaurin_all_le_objl: "diff 0 = f &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   364
      (\<forall>m x. DERIV (diff m) x :> diff (Suc m) x)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   365
      --> (\<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   366
              f x = sumr 0 n (%m. (diff m 0 / real (fact m)) * x ^ m) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   367
                    (diff n t / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   368
by (blast intro: Maclaurin_all_le)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   369
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   370
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   371
subsection{*Version for Exponential Function*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   372
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   373
lemma Maclaurin_exp_lt: "[| x ~= 0; 0 < n |]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   374
      ==> (\<exists>t. 0 < abs t &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   375
                abs t < abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   376
                exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   377
                        (exp t / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   378
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_lt_objl, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   379
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   380
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   381
lemma Maclaurin_exp_le:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   382
     "\<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   383
            exp x = sumr 0 n (%m. (x ^ m) / real (fact m)) +
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   384
                       (exp t / real (fact n)) * x ^ n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   385
by (cut_tac diff = "%n. exp" and f = exp and x = x and n = n in Maclaurin_all_le_objl, auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   386
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   387
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   388
subsection{*Version for Sine Function*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   389
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   390
lemma MVT2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   391
     "[| a < b; \<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x) |]
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   392
      ==> \<exists>z. a < z & z < b & (f b - f a = (b - a) * f'(z))"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   393
apply (drule MVT)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   394
apply (blast intro: DERIV_isCont)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   395
apply (force dest: order_less_imp_le simp add: differentiable_def)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   396
apply (blast dest: DERIV_unique order_less_imp_le)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   397
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   398
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   399
lemma mod_exhaust_less_4:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   400
     "m mod 4 = 0 | m mod 4 = 1 | m mod 4 = 2 | m mod 4 = (3::nat)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   401
by (case_tac "m mod 4", auto, arith)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   402
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   403
lemma Suc_Suc_mult_two_diff_two [rule_format, simp]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   404
     "0 < n --> Suc (Suc (2 * n - 2)) = 2*n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   405
by (induct_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   406
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   407
lemma lemma_Suc_Suc_4n_diff_2 [rule_format, simp]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   408
     "0 < n --> Suc (Suc (4*n - 2)) = 4*n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   409
by (induct_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   410
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   411
lemma Suc_mult_two_diff_one [rule_format, simp]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   412
      "0 < n --> Suc (2 * n - 1) = 2*n"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   413
by (induct_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   414
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   415
lemma Maclaurin_sin_expansion:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   416
     "\<exists>t. sin x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   417
       (sumr 0 n (%m. (if even m then 0
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   418
                       else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   419
                       x ^ m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   420
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   421
apply (cut_tac f = sin and n = n and x = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   422
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   423
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   424
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   425
apply (case_tac "n", clarify, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   426
apply (drule_tac x = 0 in spec, simp, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   427
apply (rule ccontr, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   428
apply (drule_tac x = x in spec, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   429
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   430
apply (rule_tac x = t in exI, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   431
apply (rule sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   432
apply (auto simp add: odd_Suc_mult_two_ex)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   433
apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   434
(*Could sin_zero_iff help?*)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   435
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   436
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   437
lemma Maclaurin_sin_expansion2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   438
     "\<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   439
       sin x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   440
       (sumr 0 n (%m. (if even m then 0
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   441
                       else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   442
                       x ^ m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   443
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   444
apply (cut_tac f = sin and n = n and x = x
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   445
        and diff = "%n x. sin (x + 1/2*real n * pi)" in Maclaurin_all_lt_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   446
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   447
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   448
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   449
apply (case_tac "n", clarify, simp, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   450
apply (rule ccontr, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   451
apply (drule_tac x = x in spec, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   452
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   453
apply (rule_tac x = t in exI, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   454
apply (rule sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   455
apply (auto simp add: odd_Suc_mult_two_ex)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   456
apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   457
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   458
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   459
lemma Maclaurin_sin_expansion3:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   460
     "[| 0 < n; 0 < x |] ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   461
       \<exists>t. 0 < t & t < x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   462
       sin x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   463
       (sumr 0 n (%m. (if even m then 0
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   464
                       else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   465
                       x ^ m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   466
      + ((sin(t + 1/2 * real(n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   467
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   468
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   469
apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   470
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   471
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   472
apply (rule_tac x = t in exI, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   473
apply (rule sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   474
apply (auto simp add: odd_Suc_mult_two_ex)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   475
apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   476
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   477
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   478
lemma Maclaurin_sin_expansion4:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   479
     "0 < x ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   480
       \<exists>t. 0 < t & t \<le> x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   481
       sin x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   482
       (sumr 0 n (%m. (if even m then 0
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   483
                       else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   484
                       x ^ m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   485
      + ((sin(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   486
apply (cut_tac f = sin and n = n and h = x and diff = "%n x. sin (x + 1/2*real (n) *pi)" in Maclaurin2_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   487
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   488
apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   489
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   490
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   491
apply (rule_tac x = t in exI, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   492
apply (rule sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   493
apply (auto simp add: odd_Suc_mult_two_ex)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   494
apply (auto simp add: even_mult_two_ex simp del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   495
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   496
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   497
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   498
subsection{*Maclaurin Expansion for Cosine Function*}
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   499
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   500
lemma sumr_cos_zero_one [simp]:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   501
     "sumr 0 (Suc n)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   502
         (%m. (if even m
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   503
               then (- 1) ^ (m div 2)/(real  (fact m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   504
               else 0) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   505
              0 ^ m) = 1"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   506
by (induct_tac "n", auto)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   507
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   508
lemma Maclaurin_cos_expansion:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   509
     "\<exists>t. abs t \<le> abs x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   510
       cos x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   511
       (sumr 0 n (%m. (if even m
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   512
                       then (- 1) ^ (m div 2)/(real (fact m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   513
                       else 0) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   514
                       x ^ m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   515
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   516
apply (cut_tac f = cos and n = n and x = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_all_lt_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   517
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   518
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   519
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   520
apply (case_tac "n", simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   521
apply (simp del: sumr_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   522
apply (rule ccontr, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   523
apply (drule_tac x = x in spec, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   524
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   525
apply (rule_tac x = t in exI, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   526
apply (rule sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   527
apply (auto simp add: odd_Suc_mult_two_ex)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   528
apply (auto simp add: even_mult_two_ex left_distrib cos_add simp del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   529
apply (simp add: mult_commute [of _ pi])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   530
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   531
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   532
lemma Maclaurin_cos_expansion2:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   533
     "[| 0 < x; 0 < n |] ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   534
       \<exists>t. 0 < t & t < x &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   535
       cos x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   536
       (sumr 0 n (%m. (if even m
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   537
                       then (- 1) ^ (m div 2)/(real (fact m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   538
                       else 0) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   539
                       x ^ m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   540
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   541
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   542
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   543
apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   544
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   545
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   546
apply (rule_tac x = t in exI, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   547
apply (rule sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   548
apply (auto simp add: odd_Suc_mult_two_ex)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   549
apply (auto simp add: even_mult_two_ex left_distrib cos_add simp del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   550
apply (simp add: mult_commute [of _ pi])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   551
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   552
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   553
lemma Maclaurin_minus_cos_expansion: "[| x < 0; 0 < n |] ==>
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   554
       \<exists>t. x < t & t < 0 &
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   555
       cos x =
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   556
       (sumr 0 n (%m. (if even m
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   557
                       then (- 1) ^ (m div 2)/(real (fact m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   558
                       else 0) *
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   559
                       x ^ m))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   560
      + ((cos(t + 1/2 * real (n) *pi) / real (fact n)) * x ^ n)"
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   561
apply (cut_tac f = cos and n = n and h = x and diff = "%n x. cos (x + 1/2*real (n) *pi)" in Maclaurin_minus_objl)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   562
apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   563
apply simp
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   564
apply (simp (no_asm))
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   565
apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   566
apply (rule_tac x = t in exI, simp)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   567
apply (rule sumr_fun_eq)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   568
apply (auto simp add: odd_Suc_mult_two_ex)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   569
apply (auto simp add: even_mult_two_ex left_distrib cos_add simp del: fact_Suc realpow_Suc)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   570
apply (simp add: mult_commute [of _ pi])
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   571
done
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   572
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   573
(* ------------------------------------------------------------------------- *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   574
(* Version for ln(1 +/- x). Where is it??                                    *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   575
(* ------------------------------------------------------------------------- *)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   576
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   577
lemma sin_bound_lemma:
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
   578
    "[|x = y; abs u \<le> (v::real) |] ==> \<bar>(x + u) - y\<bar> \<le> v"
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   579
by auto
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   580
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   581
lemma Maclaurin_sin_bound:
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   582
  "abs(sin x - sumr 0 n (%m. (if even m then 0 else ((- 1) ^ ((m - (Suc 0)) div 2)) / real (fact m)) *
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
   583
  x ^ m))  \<le> inverse(real (fact n)) * \<bar>x\<bar> ^ n"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   584
proof -
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   585
  have "!! x (y::real). x \<le> 1 \<Longrightarrow> 0 \<le> y \<Longrightarrow> x * y \<le> 1 * y"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   586
    by (rule_tac mult_right_mono,simp_all)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   587
  note est = this[simplified]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   588
  show ?thesis
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   589
    apply (cut_tac f=sin and n=n and x=x and
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   590
      diff = "%n x. if n mod 4 = 0 then sin(x) else if n mod 4 = 1 then cos(x) else if n mod 4 = 2 then -sin(x) else -cos(x)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   591
      in Maclaurin_all_le_objl)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   592
    apply safe
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   593
    apply simp
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   594
    apply (subst mod_Suc_eq_Suc_mod)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   595
    apply (cut_tac m=m in mod_exhaust_less_4, safe, simp+)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   596
    apply (rule DERIV_minus, simp+)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   597
    apply (rule lemma_DERIV_subst, rule DERIV_minus, rule DERIV_cos, simp)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   598
    apply (erule ssubst)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   599
    apply (rule sin_bound_lemma)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   600
    apply (rule sumr_fun_eq, safe)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   601
    apply (rule_tac f = "%u. u * (x^r)" in arg_cong)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   602
    apply (subst even_even_mod_4_iff)
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   603
    apply (cut_tac m=r in mod_exhaust_less_4, simp, safe)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   604
    apply (simp_all add:even_num_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   605
    apply (drule lemma_even_mod_4_div_2[simplified])
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   606
    apply(simp add: numeral_2_eq_2 divide_inverse)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   607
    apply (drule lemma_odd_mod_4_div_2)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   608
    apply (simp add: numeral_2_eq_2 divide_inverse)
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   609
    apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   610
                   simp add: est mult_pos_le mult_ac divide_inverse
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   611
                          power_abs [symmetric])
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   612
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   613
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 12224
diff changeset
   614
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 14738
diff changeset
   615
end