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