src/HOL/Hyperreal/Transcendental.ML
author paulson
Thu, 01 Jan 2004 10:06:32 +0100
changeset 14334 6137d24eef79
parent 14331 8dbbb7cf3637
child 14336 8f731d3cd65b
permissions -rw-r--r--
tweaking of lemmas in RealDef, RealOrd
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     1
(*  Title       : Transcendental.ML
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     3
    Copyright   : 1998,1999  University of Cambridge
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     4
                  1999 University of Edinburgh
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     5
    Description : Power Series
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     6
*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     7
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
     8
fun multr_by_tac x i = 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
     9
       let val cancel_thm = 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
    10
           CLAIM "[| (0::real)<z; x*z<y*z |] ==> x<y" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
    11
       in
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
    12
           res_inst_tac [("z",x)] cancel_thm i 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
    13
       end;
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
    14
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
    15
val mult_less_cancel_left = thm"mult_less_cancel_left";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
    16
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    17
Goalw [root_def] "root (Suc n) 0 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    18
by (safe_tac (claset() addSIs [some_equality,realpow_zero] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    19
    addSEs [realpow_zero_zero]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    20
qed "real_root_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    21
Addsimps [real_root_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    22
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    23
Goalw [root_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    24
     "0 < x ==> (root(Suc n) x) ^ (Suc n) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    25
by (dres_inst_tac [("n","n")] realpow_pos_nth2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    26
by (auto_tac (claset() addIs [someI2],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    27
qed "real_root_pow_pos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    28
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    29
Goal "0 <= x ==> (root(Suc n) x) ^ (Suc n) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    30
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    31
              addDs [real_root_pow_pos],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    32
qed "real_root_pow_pos2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    33
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    34
Goalw [root_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    35
     "0 < x ==> root(Suc n) (x ^ (Suc n)) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    36
by (rtac some_equality 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    37
by (forw_inst_tac [("n","n")] realpow_gt_zero 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
    38
by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    39
by (res_inst_tac [("R1.0","u"),("R2.0","x")] real_linear_less2 1);
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 13958
diff changeset
    40
by (dres_inst_tac [("n1","n"),("x","u")] (zero_less_Suc RSN  (3, realpow_less)) 1);
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 13958
diff changeset
    41
by (dres_inst_tac [("n1","n"),("x","x")] (zero_less_Suc RSN (3, realpow_less)) 4);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    42
by (auto_tac (claset(),simpset() addsimps [real_less_not_refl])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    43
qed "real_root_pos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    44
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    45
Goal "0 <= x ==> root(Suc n) (x ^ (Suc n)) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    46
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    47
              real_root_pos],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    48
qed "real_root_pos2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    49
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    50
Goalw [root_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    51
     "0 < x ==> 0 <= root(Suc n) x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    52
by (dres_inst_tac [("n","n")] realpow_pos_nth2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    53
by (Safe_tac THEN rtac someI2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    54
by (auto_tac (claset() addSIs [order_less_imp_le] 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
    55
    addDs [realpow_gt_zero],simpset() addsimps [zero_less_mult_iff]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    56
qed "real_root_pos_pos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    57
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    58
Goal "0 <= x ==> 0 <= root(Suc n) x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    59
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    60
              addDs [real_root_pos_pos],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    61
qed "real_root_pos_pos_le";  
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    62
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    63
Goalw [root_def] "root (Suc n) 1 = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    64
by (rtac some_equality 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    65
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    66
by (rtac ccontr 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    67
by (res_inst_tac [("R1.0","u"),("R2.0","1")] real_linear_less2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    68
by (dres_inst_tac [("n","n")] realpow_Suc_less_one 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    69
by (dres_inst_tac [("n","n")] realpow_Suc_gt_one 4);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    70
by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    71
qed "real_root_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    72
Addsimps [real_root_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    73
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    74
(*----------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    75
(* Square root                                                          *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    76
(*----------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    77
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    78
(*lcp: needed now because 2 is a binary numeral!*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    79
Goal "root 2 = root (Suc (Suc 0))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    80
by (simp_tac (simpset() delsimps [numeral_0_eq_0, numeral_1_eq_1]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    81
	                addsimps [numeral_0_eq_0 RS sym]) 1);  
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    82
qed "root_2_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    83
Addsimps [root_2_eq];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    84
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    85
Goalw [sqrt_def] "sqrt 0 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    86
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    87
qed "real_sqrt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    88
Addsimps [real_sqrt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    89
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    90
Goalw [sqrt_def] "sqrt 1 = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    91
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    92
qed "real_sqrt_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    93
Addsimps [real_sqrt_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    94
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    95
Goalw [sqrt_def] "(sqrt(x) ^ 2 = x) = (0 <= x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    96
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    97
by (cut_inst_tac [("r","root 2 x")] realpow_two_le 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    98
by (stac numeral_2_eq_2 2); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
    99
by (rtac real_root_pow_pos2 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   100
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   101
qed "real_sqrt_pow2_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   102
Addsimps [real_sqrt_pow2_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   103
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   104
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   105
Addsimps [realpow_two_le_add_order RS (real_sqrt_pow2_iff RS iffD2)];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   106
Addsimps [simplify (simpset()) (realpow_two_le_add_order RS 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   107
           (real_sqrt_pow2_iff RS iffD2))];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   108
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   109
Goalw [sqrt_def] "0 < x ==> sqrt(x) ^ 2 = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   110
by (stac numeral_2_eq_2 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   111
by (etac real_root_pow_pos 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   112
qed "real_sqrt_gt_zero_pow2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   113
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   114
Goal "(sqrt(abs(x)) ^ 2 = abs x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   115
by (rtac (real_sqrt_pow2_iff RS iffD2) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   116
by (arith_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   117
qed "real_sqrt_abs_abs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   118
Addsimps [real_sqrt_abs_abs];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   119
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   120
Goalw [sqrt_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   121
      "0 <= x ==> sqrt(x) ^ 2 = sqrt(x ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   122
by (stac numeral_2_eq_2 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   123
by (auto_tac (claset() addIs [real_root_pow_pos2 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   124
    RS ssubst, real_root_pos2 RS ssubst],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   125
     simpset() delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   126
qed "real_pow_sqrt_eq_sqrt_pow";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   127
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   128
Goal "0 <= x ==> sqrt(x) ^ 2 = sqrt(abs(x ^ 2))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   129
by (asm_full_simp_tac (simpset() addsimps [real_pow_sqrt_eq_sqrt_pow]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   130
by (stac numeral_2_eq_2 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   131
by (asm_full_simp_tac (simpset()delsimps [realpow_Suc]) 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   132
qed "real_pow_sqrt_eq_sqrt_abs_pow";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   133
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   134
Goal "0 <= x ==> sqrt(x) ^ 2 = sqrt(abs(x) ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   135
by (asm_full_simp_tac (simpset() addsimps [real_pow_sqrt_eq_sqrt_abs_pow]) 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   136
by (stac numeral_2_eq_2 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   137
by (asm_full_simp_tac (simpset()delsimps [realpow_Suc]) 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   138
qed "real_pow_sqrt_eq_sqrt_abs_pow2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   139
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   140
Goal "0 <= x ==> sqrt(x) ^ 2 = abs(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   141
by (rtac (real_sqrt_abs_abs RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   142
by (res_inst_tac [("x1","x")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   143
     (real_pow_sqrt_eq_sqrt_abs_pow2 RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   144
by (rtac (real_pow_sqrt_eq_sqrt_pow RS sym) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   145
by (assume_tac 1 THEN arith_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   146
qed "real_sqrt_pow_abs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   147
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   148
Goal "(~ (0::real) < x*x) = (x = 0)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   149
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   150
by (rtac ccontr 1);
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   151
by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   152
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   153
by (ftac (real_mult_order) 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   154
by (asm_full_simp_tac (simpset() addsimps [zero_less_mult_iff]) 1); 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   155
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   156
qed "not_real_square_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   157
Addsimps [not_real_square_gt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   158
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   159
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   160
(* proof used to be simpler *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   161
Goalw [sqrt_def,root_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   162
      "[| 0 < x; 0 < y |] ==>sqrt(x*y) =  sqrt(x) * sqrt(y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   163
by (dres_inst_tac [("n","1")] realpow_pos_nth2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   164
by (dres_inst_tac [("n","1")] realpow_pos_nth2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   165
by (asm_full_simp_tac (simpset() delsimps [realpow_Suc]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   166
                                 addsimps [numeral_2_eq_2]) 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   167
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   168
by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   169
by (Asm_full_simp_tac 1 THEN Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   170
by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   171
by (Asm_full_simp_tac 1 THEN Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   172
by (res_inst_tac [("a","xa * x")] someI2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   173
by (auto_tac (claset() addEs [real_less_asym],
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   174
    simpset() addsimps mult_ac@[realpow_mult RS sym,realpow_two_disj,
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   175
    realpow_gt_zero, real_mult_order] delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   176
qed "real_sqrt_mult_distrib";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   177
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   178
Goal "[|0<=x; 0<=y |] ==> sqrt(x*y) =  sqrt(x) * sqrt(y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   179
by (auto_tac (claset() addIs [ real_sqrt_mult_distrib],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   180
    simpset() addsimps [real_le_less]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   181
qed "real_sqrt_mult_distrib2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   182
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   183
Goal "(r * r = 0) = (r = (0::real))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   184
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   185
qed "real_mult_self_eq_zero_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   186
Addsimps [real_mult_self_eq_zero_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   187
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   188
Goalw [sqrt_def,root_def] "0 < x ==> 0 < sqrt(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   189
by (stac numeral_2_eq_2 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   190
by (dtac realpow_pos_nth2 1 THEN Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   191
by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   192
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   193
qed "real_sqrt_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   194
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   195
Goal "0 <= x ==> 0 <= sqrt(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   196
by (auto_tac (claset() addIs [real_sqrt_gt_zero],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   197
    simpset() addsimps [real_le_less]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   198
qed "real_sqrt_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   199
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   200
Goal "0 <= sqrt (x ^ 2 + y ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   201
by (auto_tac (claset() addSIs [real_sqrt_ge_zero],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   202
qed "real_sqrt_sum_squares_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   203
Addsimps [real_sqrt_sum_squares_ge_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   204
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   205
Goal "0 <= sqrt ((x ^ 2 + y ^ 2)*(xa ^ 2 + ya ^ 2))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   206
by (auto_tac (claset() addSIs [real_sqrt_ge_zero],simpset() 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   207
    addsimps [zero_le_mult_iff]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   208
qed "real_sqrt_sum_squares_mult_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   209
Addsimps [real_sqrt_sum_squares_mult_ge_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   210
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   211
Goal "sqrt ((x ^ 2 + y ^ 2) * (xa ^ 2 + ya ^ 2)) ^ 2 = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   212
\     (x ^ 2 + y ^ 2) * (xa ^ 2 + ya ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   213
by (auto_tac (claset(),simpset() addsimps [real_sqrt_pow2_iff,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   214
    zero_le_mult_iff] delsimps [realpow_Suc]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   215
qed "real_sqrt_sum_squares_mult_squared_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   216
Addsimps [real_sqrt_sum_squares_mult_squared_eq];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   217
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   218
Goal "sqrt(x ^ 2) = abs(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   219
by (rtac (abs_realpow_two RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   220
by (rtac (real_sqrt_abs_abs RS subst) 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   221
by (stac real_pow_sqrt_eq_sqrt_pow 1);
13097
c9c7f23d0ceb Had to update proof for some strange reason
nipkow
parents: 12486
diff changeset
   222
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, abs_mult]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   223
qed "real_sqrt_abs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   224
Addsimps [real_sqrt_abs];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   225
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   226
Goal "sqrt(x*x) = abs(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   227
by (rtac (realpow_two RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   228
by (stac (numeral_2_eq_2 RS sym) 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   229
by (rtac real_sqrt_abs 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   230
qed "real_sqrt_abs2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   231
Addsimps [real_sqrt_abs2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   232
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   233
Goal "0 < x ==> 0 < sqrt(x) ^ 2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   234
by (asm_full_simp_tac (simpset() addsimps [real_sqrt_gt_zero_pow2]) 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   235
qed "real_sqrt_pow2_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   236
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   237
Goal "0 < x ==> sqrt x ~= 0";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   238
by (ftac real_sqrt_pow2_gt_zero 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   239
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, real_less_not_refl]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   240
qed "real_sqrt_not_eq_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   241
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   242
Goal "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   243
by (ftac real_sqrt_not_eq_zero 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   244
by (cut_inst_tac [("n1","2"),("r1","sqrt x")] (realpow_inverse RS sym) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   245
by (auto_tac (claset() addDs [real_sqrt_gt_zero_pow2],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   246
qed "real_inv_sqrt_pow2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   247
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   248
Goal "[| 0 <= x; sqrt(x) = 0|] ==> x = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   249
by (dtac real_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   250
by (auto_tac (claset() addDs [real_sqrt_not_eq_zero],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   251
qed "real_sqrt_eq_zero_cancel";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   252
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   253
Goal "0 <= x ==> ((sqrt x = 0) = (x = 0))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   254
by (auto_tac (claset(),simpset() addsimps [real_sqrt_eq_zero_cancel]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   255
qed "real_sqrt_eq_zero_cancel_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   256
Addsimps [real_sqrt_eq_zero_cancel_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   257
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   258
Goal "x <= sqrt(x ^ 2 + y ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   259
by (subgoal_tac "x <= 0 | 0 <= x" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   260
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   261
by (rtac real_le_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   262
by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   263
by (res_inst_tac [("n","1")] realpow_increasing 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   264
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   265
				 delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   266
qed "real_sqrt_sum_squares_ge1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   267
Addsimps [real_sqrt_sum_squares_ge1];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   268
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   269
Goal "y <= sqrt(z ^ 2 + y ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   270
by (simp_tac (simpset() addsimps [real_add_commute] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   271
    delsimps [realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   272
qed "real_sqrt_sum_squares_ge2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   273
Addsimps [real_sqrt_sum_squares_ge2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   274
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   275
Goal "1 <= x ==> 1 <= sqrt x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   276
by (res_inst_tac [("n","1")] realpow_increasing 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   277
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym, real_sqrt_gt_zero_pow2,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   278
    real_sqrt_ge_zero] delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   279
qed "real_sqrt_ge_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   280
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   281
(*-------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   282
(* Exponential function                                                    *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   283
(*-------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   284
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   285
Goal "summable (%n. inverse (real (fact n)) * x ^ n)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   286
by (cut_facts_tac [zero_less_one RS real_dense] 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   287
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   288
by (cut_inst_tac [("x","r")] reals_Archimedean3 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   289
by Auto_tac;
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   290
by (dres_inst_tac [("x","abs x")] spec 1 THEN Safe_tac);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   291
by (res_inst_tac [("N","n"),("c","r")] ratio_test 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   292
by (auto_tac (claset(),
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   293
     simpset() addsimps [abs_mult,mult_assoc RS sym] delsimps [fact_Suc]));
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   294
by (rtac mult_right_mono 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   295
by (res_inst_tac [("b1","abs x")] (mult_commute RS ssubst) 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   296
by (stac fact_Suc 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   297
by (stac real_of_nat_mult 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   298
by (auto_tac (claset(),simpset() addsimps [abs_mult,inverse_mult_distrib]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   299
by (auto_tac (claset(), simpset() addsimps 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   300
     [mult_assoc RS sym, abs_eqI2, positive_imp_inverse_positive]));
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   301
by (rtac order_less_imp_le 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   302
by (res_inst_tac [("z1","real (Suc na)")] (real_mult_less_iff1
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   303
    RS iffD1) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   304
by (auto_tac (claset(),simpset() addsimps [real_not_refl2 RS not_sym,
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   305
    mult_assoc,abs_inverse]));
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   306
by (etac order_less_trans 1);
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   307
by (auto_tac (claset(),simpset() addsimps [mult_less_cancel_left]@mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   308
qed "summable_exp";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   309
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   310
Addsimps [real_of_nat_fact_gt_zero,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   311
    real_of_nat_fact_ge_zero,inv_real_of_nat_fact_gt_zero,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   312
    inv_real_of_nat_fact_ge_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   313
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   314
Goalw [real_divide_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   315
     "summable (%n. \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   316
\          (if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   317
\          else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   318
\               x ^ n)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   319
by (res_inst_tac [("g","(%n. inverse (real (fact n)) * abs(x) ^ n)")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   320
    summable_comparison_test 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   321
by (rtac summable_exp 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   322
by (res_inst_tac [("x","0")] exI 1);
12330
c69bee072501 *** empty log message ***
nipkow
parents: 12196
diff changeset
   323
by (auto_tac (claset(), simpset() addsimps [realpow_abs RS sym,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   324
    abs_mult,zero_le_mult_iff]));
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   325
by (auto_tac (claset() addIs [mult_right_mono],
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   326
    simpset() addsimps [positive_imp_inverse_positive,abs_eqI2]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   327
qed "summable_sin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   328
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   329
Goalw [real_divide_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   330
      "summable (%n. \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   331
\          (if even n then \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   332
\          (- 1) ^ (n div 2)/(real (fact n)) else 0) * x ^ n)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   333
by (res_inst_tac [("g","(%n. inverse (real (fact n)) * abs(x) ^ n)")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   334
    summable_comparison_test 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   335
by (rtac summable_exp 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   336
by (res_inst_tac [("x","0")] exI 1);
13097
c9c7f23d0ceb Had to update proof for some strange reason
nipkow
parents: 12486
diff changeset
   337
by (auto_tac (claset(), simpset() addsimps [realpow_abs RS sym,abs_mult,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   338
    zero_le_mult_iff]));
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   339
by (auto_tac (claset() addSIs [mult_right_mono],
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   340
    simpset() addsimps [positive_imp_inverse_positive,abs_eqI2]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   341
qed "summable_cos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   342
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   343
Goal "(if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   344
\      else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   345
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   346
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   347
val lemma_STAR_sin = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   348
Addsimps [lemma_STAR_sin];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   349
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   350
Goal "0 < n --> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   351
\     (- 1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   352
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   353
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   354
val lemma_STAR_cos = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   355
Addsimps [lemma_STAR_cos];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   356
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   357
Goal "0 < n --> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   358
\     (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   359
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   360
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   361
val lemma_STAR_cos1 = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   362
Addsimps [lemma_STAR_cos1];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   363
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   364
Goal "sumr 1 n (%n. if even n \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   365
\                   then (- 1) ^ (n div 2)/(real (fact n)) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   366
\                         0 ^ n \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   367
\                   else 0) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   368
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   369
by (case_tac "n" 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   370
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   371
val lemma_STAR_cos2 = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   372
Addsimps [lemma_STAR_cos2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   373
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   374
Goalw [exp_def] "(%n. inverse (real (fact n)) * x ^ n) sums exp(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   375
by (rtac (summable_exp RS summable_sums) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   376
qed "exp_converges";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   377
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   378
Goalw [sin_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   379
      "(%n. (if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   380
\           else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   381
\                x ^ n) sums sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   382
by (rtac (summable_sin RS summable_sums) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   383
qed "sin_converges";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   384
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   385
Goalw [cos_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   386
      "(%n. (if even n then \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   387
\          (- 1) ^ (n div 2)/(real (fact n)) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   388
\          else 0) * x ^ n) sums cos(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   389
by (rtac (summable_cos RS summable_sums) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   390
qed "cos_converges";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   391
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   392
Goal "p <= n --> y ^ (Suc n - p) = ((y::real) ^ (n - p)) * y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   393
by (induct_tac "n" 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   394
by (subgoal_tac "p = Suc n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   395
by (Asm_simp_tac 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   396
by (dtac sym 1 THEN asm_full_simp_tac (simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   397
    [Suc_diff_le,real_mult_commute,realpow_Suc RS sym] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   398
    delsimps [realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   399
qed_spec_mp "lemma_realpow_diff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   400
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   401
(*--------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   402
(* Properties of power series                                               *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   403
(*--------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   404
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   405
Goal "sumr 0 (Suc n) (%p. (x ^ p) * y ^ ((Suc n) - p)) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   406
\     y * sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p)))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   407
by (auto_tac (claset(),simpset() addsimps [sumr_mult] delsimps [sumr_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   408
by (rtac sumr_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   409
by (strip_tac 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   410
by (stac lemma_realpow_diff 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   411
by (auto_tac (claset(),simpset() addsimps mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   412
qed "lemma_realpow_diff_sumr";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   413
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   414
Goal "x ^ (Suc n) - y ^ (Suc n) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   415
\     (x - y) * sumr 0 (Suc n) (%p. (x ^ p) * (y ^(n - p)))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   416
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   417
by (auto_tac (claset(),simpset() delsimps [sumr_Suc]));
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   418
by (stac sumr_Suc 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   419
by (dtac sym 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   420
by (auto_tac (claset(),simpset() addsimps [lemma_realpow_diff_sumr,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   421
    right_distrib,real_diff_def] @ 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   422
    mult_ac delsimps [sumr_Suc]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   423
qed "lemma_realpow_diff_sumr2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   424
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   425
Goal "sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p))) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   426
\     sumr 0 (Suc n) (%p. (x ^ (n - p)) * (y ^ p))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   427
by (case_tac "x = y" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   428
by (auto_tac (claset(),simpset() addsimps [real_mult_commute,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   429
    realpow_add RS sym] delsimps [sumr_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   430
by (res_inst_tac [("c1","x - y")] (real_mult_left_cancel RS iffD1) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   431
by (rtac (minus_minus RS subst) 2);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   432
by (stac minus_mult_left 2);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   433
by (auto_tac (claset(),simpset() addsimps [lemma_realpow_diff_sumr2 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   434
    RS sym] delsimps [sumr_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   435
qed "lemma_realpow_rev_sumr";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   436
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   437
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   438
(* Power series has a `circle` of convergence,                              *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   439
(* i.e. if it sums for x, then it sums absolutely for z with |z| < |x|.     *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   440
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   441
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   442
Goalw [real_divide_def] "1/(x::real) = inverse x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   443
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   444
qed "real_divide_eq_inverse";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   445
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   446
Goal "[| summable (%n. f(n) * (x ^ n)); abs(z) < abs(x) |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   447
\     ==> summable (%n. abs(f(n)) * (z ^ n))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   448
by (dtac summable_LIMSEQ_zero 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   449
by (dtac convergentI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   450
by (asm_full_simp_tac (simpset() addsimps [Cauchy_convergent_iff RS sym]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   451
by (dtac Cauchy_Bseq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   452
by (asm_full_simp_tac (simpset() addsimps [Bseq_def]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   453
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   454
by (res_inst_tac [("g","%n. K * abs(z ^ n) * inverse (abs(x ^ n))")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   455
    summable_comparison_test 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   456
by (res_inst_tac [("x","0")] exI 1 THEN Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   457
by (subgoal_tac "0 < abs (x ^ n)" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   458
by (res_inst_tac [("z","abs (x ^ n)")] (CLAIM_SIMP 
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14294
diff changeset
   459
    "[| (0::real) <z; x*z<=y*z |] ==> x<=y" [mult_le_cancel_left]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   460
by (auto_tac (claset(),
12330
c69bee072501 *** empty log message ***
nipkow
parents: 12196
diff changeset
   461
    simpset() addsimps [real_mult_assoc,realpow_abs]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   462
by (dres_inst_tac [("x","0")] spec 2 THEN Force_tac 2); 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   463
by (auto_tac (claset(),simpset() addsimps [abs_mult,realpow_abs] @ mult_ac));
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14288
diff changeset
   464
by (res_inst_tac [("a2","z ^ n")] (abs_ge_zero RS real_le_imp_less_or_eq
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   465
    RS disjE) 1 THEN dtac sym 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   466
by (auto_tac (claset() addSIs [mult_right_mono],
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   467
    simpset() addsimps [mult_assoc RS sym, realpow_abs,summable_def]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   468
by (res_inst_tac [("x","K * inverse(1 - (abs(z) * inverse(abs x)))")] exI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   469
by (auto_tac (claset() addSIs [sums_mult],simpset() addsimps [real_mult_assoc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   470
by (subgoal_tac 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   471
    "abs(z ^ n) * inverse(abs x ^ n) = (abs(z) * inverse(abs x)) ^ n" 1);
12330
c69bee072501 *** empty log message ***
nipkow
parents: 12196
diff changeset
   472
by (auto_tac (claset(),simpset() addsimps [realpow_abs RS sym]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   473
by (subgoal_tac "x ~= 0" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   474
by (subgoal_tac "x ~= 0" 3);
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14288
diff changeset
   475
by (auto_tac (claset(),
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14288
diff changeset
   476
     simpset() delsimps [abs_inverse]
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14288
diff changeset
   477
      addsimps [abs_inverse RS sym, realpow_not_zero, abs_mult RS sym,
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14288
diff changeset
   478
    realpow_inverse, realpow_mult RS sym]));
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14288
diff changeset
   479
by (auto_tac (claset() addSIs [geometric_sums],
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14288
diff changeset
   480
   simpset() addsimps
12330
c69bee072501 *** empty log message ***
nipkow
parents: 12196
diff changeset
   481
    [realpow_abs,real_divide_eq_inverse RS sym]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   482
by (res_inst_tac [("z","abs(x)")] (CLAIM_SIMP 
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14294
diff changeset
   483
    "[|(0::real)<z; x*z<y*z |] ==> x<y" [mult_less_cancel_left]) 1);
13097
c9c7f23d0ceb Had to update proof for some strange reason
nipkow
parents: 12486
diff changeset
   484
by (auto_tac (claset(),simpset() addsimps [abs_mult RS sym,real_mult_assoc]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   485
qed "powser_insidea";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   486
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   487
Goal "[| summable (%n. f(n) * (x ^ n)); abs(z) < abs(x) |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   488
\     ==> summable (%n. f(n) * (z ^ n))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   489
by (dres_inst_tac [("z","abs z")] powser_insidea 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   490
by (auto_tac (claset() addIs [summable_rabs_cancel],
12330
c69bee072501 *** empty log message ***
nipkow
parents: 12196
diff changeset
   491
    simpset() addsimps [realpow_abs RS sym,abs_mult RS sym]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   492
qed "powser_inside";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   493
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   494
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   495
(*               Differentiation of power series                            *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   496
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   497
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   498
(* Lemma about distributing negation over it *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   499
Goalw [diffs_def] "diffs (%n. - c n) = (%n. - diffs c n)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   500
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   501
qed "diffs_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   502
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   503
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   504
(* Show that we can shift the terms down one                                *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   505
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   506
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   507
Goal "sumr 0 n (%n. (diffs c)(n) * (x ^ n)) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   508
\     sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) + \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   509
\     (real n * c(n) * x ^ (n - Suc 0))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   510
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   511
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc, 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   512
    real_add_assoc RS sym,diffs_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   513
qed "lemma_diffs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   514
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   515
Goal "sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   516
\     sumr 0 n (%n. (diffs c)(n) * (x ^ n)) - \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   517
\     (real n * c(n) * x ^ (n - Suc 0))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   518
by (auto_tac (claset(),simpset() addsimps [lemma_diffs]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   519
qed "lemma_diffs2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   520
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   521
Goal "summable (%n. (diffs c)(n) * (x ^ n)) ==> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   522
\     (%n. real n * c(n) * (x ^ (n - Suc 0))) sums \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   523
\        (suminf(%n. (diffs c)(n) * (x ^ n)))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   524
by (ftac summable_LIMSEQ_zero 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   525
by (subgoal_tac "(%n. real n * c(n) * (x ^ (n - Suc 0))) ----> 0" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   526
by (rtac LIMSEQ_imp_Suc 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   527
by (dtac summable_sums 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   528
by (auto_tac (claset(),simpset() addsimps [sums_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   529
by (thin_tac "(%n. diffs c n * x ^ n) ----> 0" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   530
by (rotate_tac 1 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   531
by (dtac LIMSEQ_diff 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   532
by (auto_tac (claset(),simpset() addsimps [lemma_diffs2 RS sym,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   533
    symmetric diffs_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   534
by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   535
qed "diffs_equiv";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   536
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   537
(* -------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   538
(* Term-by-term differentiability of power series                           *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   539
(* -------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   540
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   541
Goal "sumr 0 m (%p. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   542
\       sumr 0 m (%p. (z ^ p) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   543
\       (((z + h) ^ (m - p)) - (z ^ (m - p))))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   544
by (rtac sumr_subst 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   545
by (auto_tac (claset(),simpset() addsimps [right_distrib,
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   546
    real_diff_def,realpow_add RS sym] 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   547
    @ mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   548
qed "lemma_termdiff1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   549
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   550
(* proved elsewhere? *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   551
Goal "m < n --> (EX d. n = m + d + Suc 0)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   552
by (induct_tac "m" 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   553
by (case_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   554
by (case_tac "d" 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   555
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   556
qed_spec_mp "less_add_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   557
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   558
Goal " h ~= 0 ==> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   559
\       (((z + h) ^ n) - (z ^ n)) * inverse h - \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   560
\           real n * (z ^ (n - Suc 0)) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   561
\        h * sumr 0 (n - Suc 0) (%p. (z ^ p) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   562
\              sumr 0 ((n - Suc 0) - p) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   563
\                (%q. ((z + h) ^ q) * (z ^ (((n - 2) - p) - q))))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   564
by (rtac (real_mult_left_cancel RS iffD1) 1 THEN Asm_simp_tac 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   565
by (asm_full_simp_tac (simpset() addsimps [right_diff_distrib]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   566
    @ mult_ac) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   567
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   568
by (case_tac "n" 1 THEN auto_tac (claset(),simpset() 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   569
    addsimps [lemma_realpow_diff_sumr2,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   570
    right_diff_distrib RS sym,real_mult_assoc] 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   571
    delsimps [realpow_Suc,sumr_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   572
by (auto_tac (claset(),simpset() addsimps [lemma_realpow_rev_sumr]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   573
    delsimps [sumr_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   574
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc,sumr_diff_mult_const,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   575
    left_distrib,CLAIM "(a + b) - (c + d) = a - c + b - (d::real)",
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   576
    lemma_termdiff1,sumr_mult]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   577
by (auto_tac (claset() addSIs [sumr_subst],simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   578
    [real_diff_def,real_add_assoc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   579
by (fold_tac [real_diff_def] THEN dtac less_add_one 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   580
by (auto_tac (claset(),simpset() addsimps [sumr_mult,lemma_realpow_diff_sumr2] 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   581
    @ mult_ac delsimps [sumr_Suc,realpow_Suc]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   582
qed "lemma_termdiff2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   583
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   584
Goal "[| h ~= 0; abs z <= K; abs (z + h) <= K |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   585
\     ==> abs (((z + h) ^ n - z ^ n) * inverse h - real n * z ^ (n - Suc 0)) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   586
\         <= real n * real (n - Suc 0) * K ^ (n - 2) * abs h";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   587
by (stac lemma_termdiff2 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   588
by (asm_full_simp_tac (simpset() addsimps [abs_mult,real_mult_commute]) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   589
by (stac real_mult_commute 2); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   590
by (rtac (sumr_rabs RS real_le_trans) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   591
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   592
by (rtac (real_mult_commute RS subst) 2);
13097
c9c7f23d0ceb Had to update proof for some strange reason
nipkow
parents: 12486
diff changeset
   593
by (auto_tac (claset() addSIs [sumr_bound2],simpset() addsimps [abs_mult]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   594
by (case_tac "n" 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   595
by (dtac less_add_one 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   596
by (auto_tac (claset(),simpset() addsimps [realpow_add,real_add_assoc RS sym,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   597
    CLAIM_SIMP "(a * b) * c = a * (c * (b::real))" mult_ac] 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   598
    delsimps [sumr_Suc]));
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   599
by (auto_tac (claset() addSIs [mult_mono],simpset()delsimps [sumr_Suc])); 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   600
by (auto_tac (claset() addSIs [realpow_le2],simpset() addsimps 
13097
c9c7f23d0ceb Had to update proof for some strange reason
nipkow
parents: 12486
diff changeset
   601
    [realpow_abs] delsimps [sumr_Suc] ));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   602
by (res_inst_tac [("j","real (Suc d) * (K ^ d)")] real_le_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   603
by (subgoal_tac "0 <= K" 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   604
by (arith_tac 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   605
by (dres_inst_tac [("n","d")] realpow_ge_zero2 2);
13097
c9c7f23d0ceb Had to update proof for some strange reason
nipkow
parents: 12486
diff changeset
   606
by (auto_tac (claset(),simpset() delsimps [sumr_Suc] ));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   607
by (rtac (sumr_rabs RS real_le_trans) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   608
by (rtac sumr_bound2 1 THEN 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   609
    auto_tac (claset() addSDs [less_add_one]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   610
    addSIs [mult_mono], simpset() addsimps [abs_mult, realpow_add]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   611
by (auto_tac (claset() addSIs [realpow_le2,realpow_ge_zero],simpset() addsimps 
13097
c9c7f23d0ceb Had to update proof for some strange reason
nipkow
parents: 12486
diff changeset
   612
    [realpow_abs]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   613
by (ALLGOALS(arith_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   614
qed "lemma_termdiff3";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   615
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   616
Goalw [LIM_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   617
  "[| 0 < k; \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   618
\     (ALL h. 0 < abs(h) & abs(h) < k --> abs(f h) <= K * abs(h)) |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   619
\  ==> f -- 0 --> 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   620
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   621
by (subgoal_tac "0 <= K" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   622
by (dres_inst_tac [("x","k*inverse 2")] spec 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   623
by (ftac real_less_half_sum 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   624
by (dtac real_gt_half_sum 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   625
by (auto_tac (claset(),simpset() addsimps [abs_eqI2]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   626
by (res_inst_tac [("z","k/2")] (CLAIM_SIMP 
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14294
diff changeset
   627
    "[| (0::real) <z; x*z<=y*z |] ==> x<=y" [mult_le_cancel_left]) 2);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   628
by (auto_tac (claset() addIs [abs_ge_zero RS real_le_trans],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   629
by (dtac real_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   630
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   631
by (subgoal_tac "0 < (r * inverse K) * inverse 2" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   632
by (REPEAT(rtac (real_mult_order) 2));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   633
by (dres_inst_tac [("d1.0","r * inverse K * inverse 2"),("d2.0","k")]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   634
    real_lbound_gt_zero 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   635
by (auto_tac (claset(),simpset() addsimps [positive_imp_inverse_positive,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   636
    zero_less_mult_iff]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   637
by (rtac real_le_trans 2 THEN assume_tac 3 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   638
by (res_inst_tac [("x","e")] exI 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   639
by (res_inst_tac [("y","K * abs x")] order_le_less_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   640
by (res_inst_tac [("R2.0","K * e")] real_less_trans 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   641
by (res_inst_tac [("z","inverse K")] (CLAIM_SIMP 
14305
f17ca9f6dc8c tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents: 14294
diff changeset
   642
    "[|(0::real) <z; z*x<z*y |] ==> x<y" [mult_less_cancel_left]) 3);
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   643
by (asm_full_simp_tac (simpset() addsimps [mult_assoc RS sym]) 4);
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   644
by (Force_tac 1);
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   645
by (asm_full_simp_tac (simpset() addsimps [mult_less_cancel_left]) 1); 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   646
by (auto_tac (claset(),simpset() addsimps mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   647
qed "lemma_termdiff4";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   648
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   649
Goal "[| 0 < k; \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   650
\           summable f; \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   651
\           ALL h. 0 < abs(h) & abs(h) < k --> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   652
\                   (ALL n. abs(g(h) (n::nat)) <= (f(n) * abs(h))) |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   653
\        ==> (%h. suminf(g h)) -- 0 --> 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   654
by (dtac summable_sums 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   655
by (subgoal_tac "ALL h. 0 < abs h & abs h < k --> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   656
\         abs(suminf (g h)) <= suminf f * abs h" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   657
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   658
by (subgoal_tac "summable (%n. f n * abs h)" 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   659
by (simp_tac (simpset() addsimps [summable_def]) 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   660
by (res_inst_tac [("x","suminf f * abs h")] exI 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   661
by (dres_inst_tac [("c","abs h")] sums_mult 3);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   662
by (asm_full_simp_tac (simpset() addsimps mult_ac) 3);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   663
by (subgoal_tac "summable (%n. abs(g(h::real)(n::nat)))" 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   664
by (res_inst_tac [("g","%n. f(n::nat) * abs(h)")] summable_comparison_test 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   665
by (res_inst_tac [("x","0")] exI 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   666
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   667
by (res_inst_tac [("j","suminf(%n. abs(g h n))")] real_le_trans 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   668
by (auto_tac (claset() addIs [summable_rabs,summable_le],simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   669
    [sums_summable RS suminf_mult]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   670
by (auto_tac (claset() addSIs [lemma_termdiff4],simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   671
    [(sums_summable RS suminf_mult) RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   672
qed "lemma_termdiff5";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   673
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   674
(* FIXME: Long proof *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   675
Goalw [deriv_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   676
    "[| summable(%n. c(n) * (K ^ n)); \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   677
\       summable(%n. (diffs c)(n) * (K ^ n)); \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   678
\       summable(%n. (diffs(diffs c))(n) * (K ^ n)); \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   679
\       abs(x) < abs(K) |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   680
\    ==> DERIV (%x. suminf (%n. c(n) * (x ^ n)))  x :> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   681
\            suminf (%n. (diffs c)(n) * (x ^ n))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   682
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   683
by (res_inst_tac [("g","%h. suminf(%n. ((c(n) * ((x + h) ^ n)) - \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   684
\                 (c(n) * (x ^ n))) * inverse h)")] LIM_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   685
by (asm_full_simp_tac (simpset() addsimps [LIM_def]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   686
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   687
by (res_inst_tac [("x","abs K - abs x")] exI 1);
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   688
by (auto_tac (claset(),simpset() addsimps [less_diff_eq]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   689
by (dtac (abs_triangle_ineq RS order_le_less_trans) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   690
by (res_inst_tac [("y","0")] order_le_less_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   691
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   692
by (subgoal_tac "(%n. (c n) * (x ^ n)) sums \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   693
\           (suminf(%n. (c n) * (x ^ n))) & \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   694
\       (%n. (c n) * ((x + xa) ^ n)) sums \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   695
\           (suminf(%n. (c n) * ((x + xa) ^ n)))" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   696
by (auto_tac (claset() addSIs [summable_sums],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   697
by (rtac powser_inside 2 THEN rtac powser_inside 4);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   698
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   699
by (EVERY1[rotate_tac 8, dtac sums_diff, assume_tac]);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   700
by (dres_inst_tac [("x","(%n. c n * (xa + x) ^ n - c n * x ^ n)"),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   701
    ("c","inverse xa")] sums_mult 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   702
by (rtac (sums_unique RS sym) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   703
by (asm_full_simp_tac (simpset() addsimps [real_diff_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   704
    real_divide_def] @ add_ac @ mult_ac) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   705
by (rtac LIM_zero_cancel 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   706
by (res_inst_tac [("g","%h. suminf (%n. c(n) * (((((x + h) ^ n) - \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   707
\   (x ^ n)) * inverse  h) - (real n * (x ^ (n - Suc 0)))))")] LIM_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   708
by (asm_full_simp_tac (simpset() addsimps [LIM_def]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   709
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   710
by (res_inst_tac [("x","abs K - abs x")] exI 1);
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   711
by (auto_tac (claset(),simpset() addsimps [less_diff_eq]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   712
by (dtac (abs_triangle_ineq RS order_le_less_trans) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   713
by (res_inst_tac [("y","0")] order_le_less_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   714
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   715
by (subgoal_tac "summable(%n. (diffs c)(n) * (x ^ n))" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   716
by (rtac powser_inside 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   717
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   718
by (dres_inst_tac [("c","c"),("x","x")] diffs_equiv 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   719
by (ftac sums_unique 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   720
by (subgoal_tac "(%n. (c n) * (x ^ n)) sums \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   721
\           (suminf(%n. (c n) * (x ^ n))) & \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   722
\       (%n. (c n) * ((x + xa) ^ n)) sums \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   723
\           (suminf(%n. (c n) * ((x + xa) ^ n)))" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   724
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   725
by (auto_tac (claset() addSIs [summable_sums],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   726
by (rtac powser_inside 2 THEN rtac powser_inside 4);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   727
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   728
by (forw_inst_tac [("x","(%n. c n * (xa + x) ^ n)"),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   729
    ("y","(%n. c n * x ^ n)")] sums_diff 1 THEN assume_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   730
by (asm_full_simp_tac (simpset() addsimps [[sums_summable,sums_summable] 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   731
    MRS suminf_diff,right_diff_distrib RS sym]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   732
by (forw_inst_tac [("x","(%n. c n * ((xa + x) ^ n - x ^ n))"),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   733
    ("c","inverse xa")] sums_mult 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   734
by (asm_full_simp_tac (simpset() addsimps [sums_summable RS suminf_mult2]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   735
by (forw_inst_tac [("x","(%n. inverse xa * (c n * ((xa + x) ^ n - x ^ n)))"),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   736
    ("y","(%n. real n * c n * x ^ (n - Suc 0))")] sums_diff 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   737
by (assume_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   738
by (rtac (ARITH_PROVE "z - y = x ==> - x = (y::real) - z") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   739
by (asm_full_simp_tac (simpset() addsimps [[sums_summable,sums_summable] 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   740
    MRS suminf_diff] @ add_ac @ mult_ac ) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   741
by (res_inst_tac [("f","suminf")] arg_cong 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   742
by (rtac ext 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   743
by (asm_full_simp_tac (simpset() addsimps [real_diff_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   744
     right_distrib] @ add_ac @ mult_ac) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   745
(* 46 *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   746
by (dtac real_dense 1 THEN Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   747
by (ftac (real_less_sum_gt_zero) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   748
by (dres_inst_tac [("f","%n. abs(c n) * real n * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   749
\                    real (n - Suc 0) * (r ^ (n - 2))"),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   750
                   ("g","%h n. c(n) * (((((x + h) ^ n) - (x ^ n)) * inverse  h) - \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   751
\                    (real  n * (x ^ (n - Suc 0))))")] lemma_termdiff5 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   752
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   753
by (subgoal_tac "summable(%n. abs(diffs(diffs c) n) * (r ^ n))" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   754
by (res_inst_tac [("x","K")] powser_insidea 2 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   755
by (subgoal_tac "abs r = r" 2 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   756
by (res_inst_tac [("j1","abs x")] (real_le_trans RS abs_eqI1) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   757
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   758
by (asm_full_simp_tac (simpset() addsimps [diffs_def,abs_mult,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   759
    real_mult_assoc RS sym]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   760
by (subgoal_tac "ALL n. real (Suc n) * real (Suc(Suc n)) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   761
\   abs(c(Suc(Suc n))) * (r ^ n) = diffs(diffs (%n. abs(c n))) n * (r ^ n)" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   762
by (dres_inst_tac [("P","summable")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   763
    (CLAIM "[|ALL n. f(n) = g(n); P(%n. f n)|] ==> P(%n. g(n))") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   764
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   765
by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 2
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   766
    THEN asm_full_simp_tac (simpset() addsimps [diffs_def]) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   767
by (dtac diffs_equiv 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   768
by (dtac sums_summable 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   769
by (asm_full_simp_tac (simpset() addsimps [diffs_def] @ mult_ac) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   770
by (subgoal_tac "(%n. real n * (real (Suc n) * (abs(c(Suc n)) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   771
\                 (r ^ (n - Suc 0))))) = (%n. diffs(%m. real (m - Suc 0) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   772
\                  abs(c m) * inverse r) n * (r ^ n))" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   773
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   774
by (rtac ext 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   775
by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   776
by (case_tac "n" 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   777
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   778
(* 69 *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   779
by (dtac (abs_ge_zero RS order_le_less_trans) 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   780
by (asm_full_simp_tac (simpset() addsimps mult_ac) 2);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   781
by (dtac diffs_equiv 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   782
by (dtac sums_summable 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   783
by (res_inst_tac [("a","summable (%n. real n * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   784
\    (real (n - Suc 0) * abs (c n) * inverse r) * r ^ (n - Suc 0))")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   785
    (CLAIM "(a = b) ==> a ==> b") 1  THEN assume_tac 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   786
by (res_inst_tac [("f","summable")] arg_cong 1 THEN rtac ext 1);
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   787
by (dtac (abs_ge_zero RS order_le_less_trans) 2);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   788
by (asm_full_simp_tac (simpset() addsimps mult_ac) 2);
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   789
(* 77 *)
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   790
by (case_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   791
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   792
by (case_tac "nat" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   793
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   794
by (dtac (abs_ge_zero RS order_le_less_trans) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   795
by (auto_tac (claset(),simpset() addsimps [CLAIM_SIMP 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   796
    "(a::real) * (b * (c * d)) = a * (b * c) * d"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   797
     mult_ac]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   798
by (dtac (abs_ge_zero RS order_le_less_trans) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   799
by (asm_full_simp_tac (simpset() addsimps [abs_mult,real_mult_assoc]) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   800
by (rtac mult_left_mono 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   801
by (rtac (add_commute RS subst) 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   802
by (simp_tac (simpset() addsimps [mult_assoc RS sym]) 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   803
by (rtac lemma_termdiff3 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   804
by (auto_tac (claset() addIs [(abs_triangle_ineq RS real_le_trans)],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   805
    simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   806
by (arith_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   807
qed "termdiffs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   808
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   809
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   810
(* Formal derivatives of exp, sin, and cos series                           *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   811
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   812
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   813
Goalw [diffs_def]  
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   814
      "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   815
by (rtac ext 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   816
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   817
by (stac real_of_nat_mult 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   818
by (stac inverse_mult_distrib 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   819
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   820
qed "exp_fdiffs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   821
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   822
Goalw [diffs_def,real_divide_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   823
      "diffs(%n. if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   824
\          else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   825
\      = (%n. if even n then \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   826
\                (- 1) ^ (n div 2)/(real (fact n)) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   827
\             else 0)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   828
by (rtac ext 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   829
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   830
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   831
by (stac even_Suc 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   832
by (stac inverse_mult_distrib 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   833
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   834
qed "sin_fdiffs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   835
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   836
Goalw  [diffs_def,real_divide_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   837
       "diffs(%n. if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   838
\          else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) n \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   839
\      = (if even n then \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   840
\                (- 1) ^ (n div 2)/(real (fact n)) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   841
\             else 0)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   842
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   843
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   844
by (stac even_Suc 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   845
by (stac inverse_mult_distrib 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   846
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   847
qed "sin_fdiffs2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   848
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   849
(* thms in EvenOdd needed *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   850
Goalw [diffs_def,real_divide_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   851
      "diffs(%n. if even n then \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   852
\                (- 1) ^ (n div 2)/(real (fact n)) else 0) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   853
\      = (%n. - (if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   854
\          else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n))))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   855
by (rtac ext 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   856
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   857
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   858
by (stac even_Suc 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   859
by (stac inverse_mult_distrib 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   860
by (res_inst_tac [("z1","real (Suc n)")] (real_mult_commute RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   861
by (res_inst_tac [("z1","inverse(real (Suc n))")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   862
     (real_mult_commute RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   863
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   864
    odd_not_even RS sym,odd_Suc_mult_two_ex]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   865
qed "cos_fdiffs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   866
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   867
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   868
Goalw [diffs_def,real_divide_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   869
      "diffs(%n. if even n then \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   870
\                (- 1) ^ (n div 2)/(real (fact n)) else 0) n\
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   871
\      = - (if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   872
\          else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n)))";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   873
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   874
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   875
by (stac even_Suc 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   876
by (stac inverse_mult_distrib 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   877
by (res_inst_tac [("z1","real (Suc n)")] (real_mult_commute RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   878
by (res_inst_tac [("z1","inverse (real (Suc n))")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   879
     (real_mult_commute RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   880
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   881
    odd_not_even RS sym,odd_Suc_mult_two_ex]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   882
qed "cos_fdiffs2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   883
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   884
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   885
(* Now at last we can get the derivatives of exp, sin and cos               *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   886
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   887
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   888
Goal "- sin x = suminf(%n. - ((if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   889
\    else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   890
by (auto_tac (claset() addSIs [sums_unique,sums_minus,sin_converges],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   891
    simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   892
qed "lemma_sin_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   893
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   894
Goal "exp = (%x. suminf (%n. inverse (real (fact n)) * x ^ n))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   895
by (auto_tac (claset() addSIs [ext],simpset() addsimps [exp_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   896
val lemma_exp_ext = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   897
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   898
Goalw [exp_def] "DERIV exp x :> exp(x)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   899
by (stac lemma_exp_ext 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   900
by (subgoal_tac "DERIV (%u. suminf (%n. inverse (real (fact n)) * u ^ n)) x \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   901
\    :> suminf (%n. diffs (%n. inverse (real (fact n))) n * x ^ n)" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   902
by (res_inst_tac [("K","1 + abs(x)")] termdiffs 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   903
by (auto_tac (claset() addIs [exp_converges RS sums_summable],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   904
    simpset() addsimps [exp_fdiffs]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   905
by (arith_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   906
qed "DERIV_exp";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   907
Addsimps [DERIV_exp];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   908
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   909
Goal "sin = (%x. suminf(%n. (if even n then 0 \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   910
\           else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   911
\                x ^ n))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   912
by (auto_tac (claset() addSIs [ext],simpset() addsimps [sin_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   913
val lemma_sin_ext = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   914
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   915
Goal "cos = (%x. suminf(%n. (if even n then \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   916
\          (- 1) ^ (n div 2)/(real (fact n)) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   917
\          else 0) * x ^ n))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   918
by (auto_tac (claset() addSIs [ext],simpset() addsimps [cos_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   919
val lemma_cos_ext = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   920
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   921
Goalw [cos_def] "DERIV sin x :> cos(x)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   922
by (stac lemma_sin_ext 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   923
by (auto_tac (claset(),simpset() addsimps [sin_fdiffs2 RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   924
by (res_inst_tac [("K","1 + abs(x)")] termdiffs 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   925
by (auto_tac (claset() addIs [sin_converges, cos_converges, sums_summable] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   926
    addSIs [sums_minus RS sums_summable],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   927
    simpset() addsimps [cos_fdiffs,sin_fdiffs]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   928
by (arith_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   929
qed "DERIV_sin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   930
Addsimps [DERIV_sin];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   931
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   932
Goal "DERIV cos x :> -sin(x)";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
   933
by (stac lemma_cos_ext 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   934
by (auto_tac (claset(),simpset() addsimps [lemma_sin_minus,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   935
    cos_fdiffs2 RS sym,minus_mult_left]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   936
by (res_inst_tac [("K","1 + abs(x)")] termdiffs 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   937
by (auto_tac (claset() addIs [sin_converges,cos_converges, sums_summable] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   938
    addSIs [sums_minus RS sums_summable],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   939
    simpset() addsimps [cos_fdiffs,sin_fdiffs,diffs_minus]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   940
by (arith_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   941
qed "DERIV_cos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   942
Addsimps [DERIV_cos];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   943
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   944
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   945
(* Properties of the exponential function                                   *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   946
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   947
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   948
Goalw [exp_def] "exp 0 = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   949
by (rtac (CLAIM_SIMP "sumr 0 1 (%n. inverse (real (fact n)) * 0 ^ n) = 1" 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   950
           [real_of_nat_one] RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   951
by (rtac ((series_zero RS sums_unique) RS sym) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   952
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   953
by (case_tac "m" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   954
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   955
qed "exp_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   956
Addsimps [exp_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   957
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   958
Goal "0 <= x ==> (1 + x) <= exp(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   959
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   960
by (rewtac exp_def);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   961
by (rtac real_le_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   962
by (res_inst_tac [("n","2"),("f","(%n. inverse (real (fact n)) * x ^ n)")]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   963
    series_pos_le 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   964
by (auto_tac (claset() addIs [summable_exp],simpset() 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   965
    addsimps [numeral_2_eq_2,realpow_ge_zero,zero_le_mult_iff]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   966
qed "exp_ge_add_one_self";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   967
Addsimps [exp_ge_add_one_self];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   968
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   969
Goal "0 < x ==> 1 < exp x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   970
by (rtac order_less_le_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   971
by (rtac exp_ge_add_one_self 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   972
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   973
qed "exp_gt_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   974
Addsimps [exp_gt_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   975
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   976
Goal "DERIV (%x. exp (x + y)) x :> exp(x + y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   977
by (auto_tac (claset(),simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   978
    [CLAIM_SIMP "(%x. exp (x + y)) = exp o (%x. x + y)" [ext]]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   979
by (rtac (real_mult_1_right RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   980
by (rtac DERIV_chain 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   981
by (rtac (add_zero_right RS subst) 2);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   982
by (rtac DERIV_add 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   983
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   984
qed "DERIV_exp_add_const";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   985
Addsimps [DERIV_exp_add_const];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   986
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   987
Goal "DERIV (%x. exp (-x)) x :> - exp(-x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   988
by (auto_tac (claset(),simpset() addsimps
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   989
    [CLAIM_SIMP "(%x. exp(-x)) = exp o (%x. - x)" [ext]]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   990
by (rtac (real_mult_1_right RS subst) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   991
by (rtac (minus_mult_left RS subst) 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
   992
by (stac minus_mult_right 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   993
by (rtac DERIV_chain 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   994
by (rtac DERIV_minus 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   995
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   996
qed "DERIV_exp_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   997
Addsimps [DERIV_exp_minus];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   998
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
   999
Goal "DERIV (%x. exp (x + y) * exp (- x)) x :> 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1000
by (cut_inst_tac [("x","x"),("y2","y")] ([DERIV_exp_add_const,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1001
    DERIV_exp_minus] MRS DERIV_mult) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1002
by (auto_tac (claset(),simpset() addsimps mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1003
qed "DERIV_exp_exp_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1004
Addsimps [DERIV_exp_exp_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1005
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1006
Goal "exp(x + y)*exp(-x) = exp(y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1007
by (cut_inst_tac [("x","x"),("y2","y"),("y","0")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1008
    ((CLAIM "ALL x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0") RS
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1009
      DERIV_isconst_all) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1010
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1011
qed "exp_add_mult_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1012
Addsimps [exp_add_mult_minus];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1013
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1014
Goal "exp(x)*exp(-x) = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1015
by (cut_inst_tac [("x","x"),("y","0")] exp_add_mult_minus 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1016
by (Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1017
qed "exp_mult_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1018
Addsimps [exp_mult_minus];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1019
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1020
Goal "exp(-x)*exp(x) = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1021
by (simp_tac (simpset() addsimps [real_mult_commute]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1022
qed "exp_mult_minus2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1023
Addsimps [exp_mult_minus2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1024
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1025
Goal "exp(-x) = inverse(exp(x))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1026
by (auto_tac (claset() addIs [real_inverse_unique],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1027
qed "exp_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1028
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1029
Goal "exp(x + y) = exp(x) * exp(y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1030
by (cut_inst_tac [("x1","x"),("y1","y"),("z","exp x")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1031
    (exp_add_mult_minus RS (CLAIM "x = y ==> z * y = z * (x::real)")) 1);
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
  1032
by (asm_full_simp_tac HOL_ss 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1033
by (asm_full_simp_tac (simpset() delsimps [exp_add_mult_minus] 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1034
    addsimps mult_ac) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1035
qed "exp_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1036
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1037
Goal "0 <= exp x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1038
by (res_inst_tac [("t","x")] (real_sum_of_halves RS subst) 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1039
by (stac exp_add 1 THEN Auto_tac);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1040
qed "exp_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1041
Addsimps [exp_ge_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1042
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1043
Goal "exp x ~= 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1044
by (cut_inst_tac [("x","x")] exp_mult_minus2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1045
by (auto_tac (claset(),simpset() delsimps [exp_mult_minus2]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1046
qed "exp_not_eq_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1047
Addsimps [exp_not_eq_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1048
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1049
Goal "0 < exp x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1050
by (simp_tac (simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1051
    [CLAIM_SIMP "(x < y) = (x <= y & y ~= (x::real))" [real_le_less]]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1052
qed "exp_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1053
Addsimps [exp_gt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1054
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1055
Goal "0 < inverse(exp x)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1056
by (auto_tac (claset() addIs [positive_imp_inverse_positive],simpset()));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1057
qed "inv_exp_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1058
Addsimps [inv_exp_gt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1059
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1060
Goal "abs(exp x) = exp x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1061
by (auto_tac (claset(),simpset() addsimps [abs_eqI2]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1062
qed "abs_exp_cancel";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1063
Addsimps [abs_exp_cancel];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1064
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1065
Goal "exp(real n * x) = exp(x) ^ n";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1066
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1067
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1068
    right_distrib,exp_add,real_mult_commute]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1069
qed "exp_real_of_nat_mult";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1070
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1071
Goalw [real_diff_def,real_divide_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1072
       "exp(x - y) = exp(x)/(exp y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1073
by (simp_tac (simpset() addsimps [exp_add,exp_minus]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1074
qed "exp_diff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1075
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1076
Goal "x < y ==> exp x < exp y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1077
by (dtac ((real_less_sum_gt_zero) RS exp_gt_one) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1078
by (multr_by_tac "inverse(exp x)" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1079
by (auto_tac (claset(),simpset() addsimps [exp_add,exp_minus]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1080
qed "exp_less_mono";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1081
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1082
Goal "exp x < exp y ==> x < y";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1083
by (EVERY1[rtac ccontr, dtac (linorder_not_less RS iffD1), dtac real_le_imp_less_or_eq]);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1084
by (auto_tac (claset() addDs [exp_less_mono],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1085
qed "exp_less_cancel";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1086
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1087
Goal "(exp(x) < exp(y)) = (x < y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1088
by (auto_tac (claset() addIs [exp_less_mono,exp_less_cancel],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1089
qed "exp_less_cancel_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1090
AddIffs [exp_less_cancel_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1091
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1092
Goalw [real_le_def] "(exp(x) <= exp(y)) = (x <= y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1093
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1094
qed "exp_le_cancel_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1095
AddIffs [exp_le_cancel_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1096
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1097
Goal "(exp x = exp y) = (x = y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1098
by (auto_tac (claset(),simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1099
    [CLAIM "(x = (y::real)) = (x <= y & y <= x)"]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1100
qed "exp_inj_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1101
AddIffs [exp_inj_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1102
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1103
Goal "1 <= y ==> EX x. 0 <= x & x <= y - 1 & exp(x) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1104
by (rtac IVT 1);
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1105
by (auto_tac (claset() addIs [DERIV_exp RS DERIV_isCont],
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1106
              simpset() addsimps [le_diff_eq]));
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1107
by (dtac (CLAIM_SIMP "x <= y ==> (0::real) <= y - x" [le_diff_eq]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1108
by (dtac exp_ge_add_one_self 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1109
by (Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1110
qed "lemma_exp_total";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1111
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1112
Goal "0 < y ==> EX x. exp x = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1113
by (res_inst_tac [("R1.0","1"),("R2.0","y")] real_linear_less2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1114
by (dtac (order_less_imp_le RS lemma_exp_total) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1115
by (res_inst_tac [("x","0")] exI 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1116
by (ftac real_inverse_gt_one 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1117
by (dtac (order_less_imp_le RS lemma_exp_total) 4);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1118
by (Step_tac 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1119
by (res_inst_tac [("x","-x")] exI 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1120
by (auto_tac (claset(),simpset() addsimps [exp_minus]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1121
qed "exp_total";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1122
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1123
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1124
(* Properties of the logarithmic function                                   *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1125
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1126
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1127
Goalw [ln_def] "ln(exp x) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1128
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1129
qed "ln_exp";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1130
Addsimps [ln_exp];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1131
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1132
Goal "(exp(ln x) = x) = (0 < x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1133
by (auto_tac (claset() addDs [exp_total],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1134
by (dtac subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1135
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1136
qed "exp_ln_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1137
Addsimps [exp_ln_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1138
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1139
Goal "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1140
by (rtac (exp_inj_iff RS iffD1) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1141
by (ftac (real_mult_order) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1142
by (auto_tac (claset(),simpset() addsimps [exp_add,exp_ln_iff RS sym] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1143
    delsimps [exp_inj_iff,exp_ln_iff]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1144
qed "ln_mult";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1145
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1146
Goal "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)";
13601
fd3e3d6b37b2 Adapted to new simplifier.
berghofe
parents: 13097
diff changeset
  1147
by (auto_tac (claset() addSDs [(exp_ln_iff RS iffD2 RS sym)],simpset()));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1148
qed "ln_inj_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1149
Addsimps [ln_inj_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1150
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1151
Goal "ln 1 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1152
by (rtac (exp_inj_iff RS iffD1) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1153
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1154
qed "ln_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1155
Addsimps [ln_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1156
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1157
Goal "0 < x ==> ln(inverse x) = - ln x";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1158
by (res_inst_tac [("a1","ln x")] (add_left_cancel RS iffD1) 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1159
by (auto_tac (claset(),simpset() addsimps [positive_imp_inverse_positive,ln_mult RS sym]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1160
qed "ln_inverse";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1161
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1162
Goalw [real_divide_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1163
    "[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1164
by (auto_tac (claset(),simpset() addsimps [positive_imp_inverse_positive,
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1165
    ln_mult,ln_inverse]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1166
qed "ln_div";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1167
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1168
Goal "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1169
by (REPEAT(dtac (exp_ln_iff RS iffD2) 1));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1170
by (REPEAT(dtac subst 1 THEN assume_tac 2));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1171
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1172
qed "ln_less_cancel_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1173
Addsimps [ln_less_cancel_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1174
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1175
Goalw [real_le_def] "[| 0 < x; 0 < y|] ==> (ln x <= ln y) = (x <= y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1176
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1177
qed "ln_le_cancel_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1178
Addsimps [ln_le_cancel_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1179
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1180
Goal "0 < x ==> ln(x ^ n) = real n * ln(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1181
by (auto_tac (claset() addSDs [exp_total],simpset() 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1182
    addsimps [exp_real_of_nat_mult RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1183
qed "ln_realpow";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1184
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1185
Goal "0 <= x ==> ln(1 + x) <= x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1186
by (rtac (ln_exp RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1187
by (rtac (ln_le_cancel_iff RS iffD2) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1188
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1189
qed "ln_add_one_self_le_self";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1190
Addsimps [ln_add_one_self_le_self];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1191
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1192
Goal "0 < x ==> ln x < x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1193
by (rtac order_less_le_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1194
by (rtac ln_add_one_self_le_self 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1195
by (rtac (ln_less_cancel_iff RS iffD2) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1196
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1197
qed "ln_less_self";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1198
Addsimps [ln_less_self];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1199
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1200
Goal "1 <= x ==> 0 <= ln x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1201
by (subgoal_tac "0 < x" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1202
by (rtac order_less_le_trans 2 THEN assume_tac 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1203
by (rtac (exp_le_cancel_iff RS iffD1) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1204
by (auto_tac (claset(),simpset() addsimps 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1205
    [exp_ln_iff RS sym] delsimps [exp_ln_iff]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1206
qed "ln_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1207
Addsimps [ln_ge_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1208
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1209
Goal "1 < x ==> 0 < ln x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1210
by (rtac (exp_less_cancel_iff RS iffD1) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1211
by (rtac (exp_ln_iff RS iffD2 RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1212
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1213
qed "ln_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1214
Addsimps [ln_gt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1215
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1216
Goal "[| 0 < x; x ~= 1 |] ==> ln x ~= 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1217
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1218
by (dtac (exp_inj_iff RS iffD2) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1219
by (dtac (exp_ln_iff RS iffD2) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1220
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1221
qed "ln_not_eq_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1222
Addsimps [ln_not_eq_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1223
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1224
Goal "[| 0 < x; x < 1 |] ==> ln x < 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1225
by (rtac (exp_less_cancel_iff RS iffD1) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1226
by (auto_tac (claset(),simpset() addsimps [exp_ln_iff RS sym]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1227
    delsimps [exp_ln_iff]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1228
qed "ln_less_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1229
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1230
Goal "exp u = x ==> ln x = u";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1231
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1232
qed "exp_ln_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1233
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1234
Addsimps [hypreal_less_not_refl];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1235
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1236
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1237
(* Basic properties of the trig functions                                   *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1238
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1239
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1240
Goalw [sin_def] "sin 0 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1241
by (auto_tac (claset() addSIs [(sums_unique RS sym),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1242
    LIMSEQ_const],simpset() addsimps [sums_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1243
qed "sin_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1244
Addsimps [sin_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1245
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1246
Goal "(ALL m. n <= m --> f m = 0) --> f sums sumr 0 n f";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1247
by (auto_tac (claset() addIs [series_zero],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1248
qed "lemma_series_zero2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1249
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1250
Goalw [cos_def] "cos 0 = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1251
by (rtac (sums_unique RS sym) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1252
by (cut_inst_tac [("n","1"),("f","(%n. (if even n then (- 1) ^ (n div 2)/ \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1253
\   (real (fact n)) else 0) * 0 ^ n)")] lemma_series_zero2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1254
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1255
qed "cos_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1256
Addsimps [cos_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1257
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1258
Goal "DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1259
by (rtac DERIV_mult 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1260
qed "DERIV_sin_sin_mult";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1261
Addsimps [DERIV_sin_sin_mult];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1262
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1263
Goal "DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1264
by (cut_inst_tac [("x","x")] DERIV_sin_sin_mult 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1265
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1266
qed "DERIV_sin_sin_mult2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1267
Addsimps [DERIV_sin_sin_mult2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1268
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1269
Goal "DERIV (%x. sin(x) ^ 2) x :> cos(x) * sin(x) + cos(x) * sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1270
by (auto_tac (claset(),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1271
              simpset() addsimps [numeral_2_eq_2, real_mult_assoc RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1272
qed "DERIV_sin_realpow2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1273
Addsimps [DERIV_sin_realpow2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1274
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1275
Goal "DERIV (%x. sin(x) ^ 2) x :> 2 * cos(x) * sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1276
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1277
qed "DERIV_sin_realpow2a";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1278
Addsimps [ DERIV_sin_realpow2a];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1279
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1280
Goal "DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1281
by (rtac DERIV_mult 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1282
qed "DERIV_cos_cos_mult";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1283
Addsimps [DERIV_cos_cos_mult];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1284
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1285
Goal "DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1286
by (cut_inst_tac [("x","x")] DERIV_cos_cos_mult 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1287
by (auto_tac (claset(),simpset() addsimps mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1288
qed "DERIV_cos_cos_mult2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1289
Addsimps [DERIV_cos_cos_mult2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1290
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1291
Goal "DERIV (%x. cos(x) ^ 2) x :> -sin(x) * cos(x) + -sin(x) * cos(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1292
by (auto_tac (claset(),
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1293
              simpset() addsimps [numeral_2_eq_2, real_mult_assoc RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1294
qed "DERIV_cos_realpow2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1295
Addsimps [DERIV_cos_realpow2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1296
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1297
Goal "DERIV (%x. cos(x) ^ 2) x :> -2 * cos(x) * sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1298
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1299
qed "DERIV_cos_realpow2a";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1300
Addsimps [DERIV_cos_realpow2a];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1301
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1302
Goal "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1303
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1304
val lemma_DERIV_subst = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1305
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1306
Goal "DERIV (%x. cos(x) ^ 2) x :> -(2 * cos(x) * sin(x))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1307
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1308
by (rtac DERIV_cos_realpow2a 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1309
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1310
qed "DERIV_cos_realpow2b";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1311
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1312
(* most useful *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1313
Goal "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1314
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1315
by (rtac DERIV_cos_cos_mult2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1316
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1317
qed "DERIV_cos_cos_mult3";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1318
Addsimps [DERIV_cos_cos_mult3];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1319
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1320
Goalw [real_diff_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1321
     "ALL x. DERIV (%x. sin(x) ^ 2 + cos(x) ^ 2) x :> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1322
\            (2*cos(x)*sin(x) - 2*cos(x)*sin(x))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1323
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1324
by (rtac DERIV_add 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1325
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1326
qed "DERIV_sin_circle_all";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1327
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1328
Goal "ALL x. DERIV (%x. sin(x) ^ 2 + cos(x) ^ 2) x :> 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1329
by (cut_facts_tac [DERIV_sin_circle_all] 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1330
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1331
qed "DERIV_sin_circle_all_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1332
Addsimps [DERIV_sin_circle_all_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1333
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1334
Goal "(sin(x) ^ 2) + (cos(x) ^ 2) = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1335
by (cut_inst_tac [("x","x"),("y","0")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1336
    (DERIV_sin_circle_all_zero RS DERIV_isconst_all) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1337
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1338
qed "sin_cos_squared_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1339
Addsimps [sin_cos_squared_add];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1340
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1341
Goal "(cos(x) ^ 2) + (sin(x) ^ 2) = 1";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1342
by (stac real_add_commute 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1343
by (simp_tac (simpset() delsimps [realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1344
qed "sin_cos_squared_add2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1345
Addsimps [sin_cos_squared_add2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1346
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1347
Goal "cos x * cos x + sin x * sin x = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1348
by (cut_inst_tac [("x","x")] sin_cos_squared_add2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1349
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1350
qed "sin_cos_squared_add3";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1351
Addsimps [sin_cos_squared_add3];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1352
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1353
Goal "(sin(x) ^ 2) = 1 - (cos(x) ^ 2)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1354
by (res_inst_tac [("a1","(cos(x) ^ 2)")] (add_right_cancel RS iffD1) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1355
by (simp_tac (simpset() delsimps [realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1356
qed "sin_squared_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1357
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1358
Goal "(cos(x) ^ 2) = 1 - (sin(x) ^ 2)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1359
by (res_inst_tac [("a1","(sin(x) ^ 2)")] (add_right_cancel RS iffD1) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1360
by (simp_tac (simpset() delsimps [realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1361
qed "cos_squared_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1362
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1363
Goal "[| 1 < x; 0 <= y |] ==> 1 < x + (y::real)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1364
by (arith_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1365
qed "real_gt_one_ge_zero_add_less";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1366
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1367
Goalw [real_le_def] "abs(sin x) <= 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1368
by (rtac notI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1369
by (dtac realpow_two_gt_one 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1370
by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1371
by (dres_inst_tac [("r1","cos x")] (realpow_two_le RSN 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1372
    (2, real_gt_one_ge_zero_add_less)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1373
by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1374
				 delsimps [realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1375
qed "abs_sin_le_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1376
Addsimps [abs_sin_le_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1377
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1378
Goal "- 1 <= sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1379
by (full_simp_tac (simpset() addsimps [simplify (simpset()) (abs_sin_le_one RS 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1380
    (abs_le_interval_iff RS iffD1))]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1381
qed "sin_ge_minus_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1382
Addsimps [sin_ge_minus_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1383
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1384
Goal "-1 <= sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1385
by (rtac (simplify (simpset()) sin_ge_minus_one) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1386
qed "sin_ge_minus_one2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1387
Addsimps [sin_ge_minus_one2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1388
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1389
Goal "sin x <= 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1390
by (full_simp_tac (simpset() addsimps [abs_sin_le_one RS 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1391
    (abs_le_interval_iff RS iffD1)]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1392
qed "sin_le_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1393
Addsimps [sin_le_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1394
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1395
Goalw [real_le_def] "abs(cos x) <= 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1396
by (rtac notI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1397
by (dtac realpow_two_gt_one 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1398
by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1399
by (dres_inst_tac [("r1","sin x")] (realpow_two_le RSN 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1400
    (2, real_gt_one_ge_zero_add_less)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1401
by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1402
				 delsimps [realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1403
qed "abs_cos_le_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1404
Addsimps [abs_cos_le_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1405
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1406
Goal "- 1 <= cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1407
by (full_simp_tac (simpset() addsimps [simplify (simpset())(abs_cos_le_one RS 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1408
    (abs_le_interval_iff RS iffD1))]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1409
qed "cos_ge_minus_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1410
Addsimps [cos_ge_minus_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1411
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1412
Goal "-1 <= cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1413
by (rtac (simplify (simpset()) cos_ge_minus_one) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1414
qed "cos_ge_minus_one2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1415
Addsimps [cos_ge_minus_one2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1416
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1417
Goal "cos x <= 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1418
by (full_simp_tac (simpset() addsimps [abs_cos_le_one RS 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1419
    (abs_le_interval_iff RS iffD1)]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1420
qed "cos_le_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1421
Addsimps [cos_le_one];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1422
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1423
Goal "DERIV g x :> m ==> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1424
\     DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1425
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1426
by (res_inst_tac [("f","(%x. x ^ n)")] DERIV_chain2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1427
by (rtac DERIV_pow 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1428
qed "DERIV_fun_pow";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1429
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1430
Goal "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1431
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1432
by (res_inst_tac [("f","exp")] DERIV_chain2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1433
by (rtac DERIV_exp 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1434
qed "DERIV_fun_exp";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1435
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1436
Goal "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1437
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1438
by (res_inst_tac [("f","sin")] DERIV_chain2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1439
by (rtac DERIV_sin 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1440
qed "DERIV_fun_sin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1441
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1442
Goal "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1443
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1444
by (res_inst_tac [("f","cos")] DERIV_chain2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1445
by (rtac DERIV_cos 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1446
qed "DERIV_fun_cos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1447
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1448
(* FIXME: remove this quick, crude tactic *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1449
exception DERIV_name;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1450
fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1451
|   get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1452
|   get_fun_name _ = raise DERIV_name;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1453
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1454
val deriv_rulesI = [DERIV_Id,DERIV_const,DERIV_cos,DERIV_cmult,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1455
                    DERIV_sin, DERIV_exp, DERIV_inverse,DERIV_pow,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1456
                    DERIV_add, DERIV_diff, DERIV_mult, DERIV_minus,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1457
                    DERIV_inverse_fun,DERIV_quotient,DERIV_fun_pow,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1458
                    DERIV_fun_exp,DERIV_fun_sin,DERIV_fun_cos,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1459
                    DERIV_Id,DERIV_const,DERIV_cos];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1460
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1461
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1462
fun deriv_tac i = (resolve_tac deriv_rulesI i) ORELSE 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1463
                   ((rtac (read_instantiate [("f",get_fun_name (getgoal i))] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1464
                     DERIV_chain2) i) handle DERIV_name => no_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1465
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1466
val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1467
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1468
(* lemma *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1469
Goal "ALL x. \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1470
\        DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1471
\              (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1472
by (Step_tac 1 THEN rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1473
by DERIV_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1474
by (auto_tac (claset(),simpset() addsimps [real_diff_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1475
    left_distrib,right_distrib] @ 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1476
    mult_ac @ add_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1477
val lemma_DERIV_sin_cos_add = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1478
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1479
Goal "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1480
\     (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1481
by (cut_inst_tac [("y","0"),("x","x"),("y7","y")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1482
    (lemma_DERIV_sin_cos_add RS DERIV_isconst_all) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1483
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1484
qed "sin_cos_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1485
Addsimps [sin_cos_add];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1486
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1487
Goal "sin (x + y) = sin x * cos y + cos x * sin y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1488
by (cut_inst_tac [("x","x"),("y","y")] sin_cos_add 1); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1489
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1490
         simpset()  addsimps [numeral_2_eq_2] delsimps [sin_cos_add]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1491
qed "sin_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1492
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1493
Goal "cos (x + y) = cos x * cos y - sin x * sin y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1494
by (cut_inst_tac [("x","x"),("y","y")] sin_cos_add 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1495
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_add]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1496
qed "cos_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1497
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1498
Goal "ALL x. \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1499
\         DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1500
by (Step_tac 1 THEN rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1501
by DERIV_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1502
by (auto_tac (claset(),simpset() addsimps [real_diff_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1503
    left_distrib,right_distrib]
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1504
    @ mult_ac @ add_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1505
val lemma_DERIV_sin_cos_minus = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1506
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1507
Goal "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1508
by (cut_inst_tac [("y","0"),("x","x")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1509
    (lemma_DERIV_sin_cos_minus RS DERIV_isconst_all) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1510
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1511
qed "sin_cos_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1512
Addsimps [sin_cos_minus];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1513
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1514
Goal "sin (-x) = -sin(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1515
by (cut_inst_tac [("x","x")] sin_cos_minus 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1516
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1517
              simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_minus]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1518
qed "sin_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1519
Addsimps [sin_minus];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1520
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1521
Goal "cos (-x) = cos(x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1522
by (cut_inst_tac [("x","x")] sin_cos_minus 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1523
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1524
              simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_minus]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1525
qed "cos_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1526
Addsimps [cos_minus];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1527
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1528
Goalw [real_diff_def] "sin (x - y) = sin x * cos y - cos x * sin y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1529
by (simp_tac (simpset() addsimps [sin_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1530
qed "sin_diff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1531
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1532
Goal "sin (x - y) = cos y * sin x - sin y * cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1533
by (simp_tac (simpset() addsimps [sin_diff,real_mult_commute]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1534
qed "sin_diff2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1535
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1536
Goalw [real_diff_def] "cos (x - y) = cos x * cos y + sin x * sin y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1537
by (simp_tac (simpset() addsimps [cos_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1538
qed "cos_diff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1539
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1540
Goal "cos (x - y) = cos y * cos x + sin y * sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1541
by (simp_tac (simpset() addsimps [cos_diff,real_mult_commute]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1542
qed "cos_diff2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1543
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1544
Goal "sin(2 * x) = 2* sin x * cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1545
by (cut_inst_tac [("x","x"),("y","x")] sin_add 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1546
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1547
qed "sin_double";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1548
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1549
Addsimps [sin_double];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1550
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1551
Goal "cos(2* x) = (cos(x) ^ 2) - (sin(x) ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1552
by (cut_inst_tac [("x","x"),("y","x")] cos_add 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1553
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1554
qed "cos_double";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1555
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1556
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1557
(* Show that there's a least positive x with cos(x) = 0; hence define pi    *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1558
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1559
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1560
Goal "(%n. (- 1) ^ n /(real (fact (2 * n + 1))) * \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1561
\      x ^ (2 * n + 1)) sums  sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1562
by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((sin_converges 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1563
    RS sums_summable) RS sums_group)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1564
by (auto_tac (claset(),simpset() addsimps mult_ac@[sin_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1565
qed "sin_paired";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1566
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1567
Goal "real (Suc (Suc (Suc (Suc 2)))) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1568
\     real (2::nat) * real (Suc 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1569
by (simp_tac (simpset() addsimps [numeral_2_eq_2, real_of_nat_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1570
val lemma_real_of_nat_six_mult = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1571
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1572
Goal "[|0 < x; x < 2 |] ==> 0 < sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1573
by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((sin_paired
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1574
    RS sums_summable) RS sums_group)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1575
by (rotate_tac 2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1576
by (dtac ((sin_paired RS sums_unique) RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1577
by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1578
by (ftac sums_unique 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1579
by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1580
by (res_inst_tac [("n1","0")] (series_pos_less RSN (2,order_le_less_trans)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1581
by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1582
by (etac sums_summable 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1583
by (case_tac "m=0" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1584
by (Asm_simp_tac 1);
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1585
by (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x" 1); 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1586
by (asm_full_simp_tac (HOL_ss addsimps [mult_less_cancel_left]) 1); 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1587
by (asm_full_simp_tac (simpset() addsimps []) 1); 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1588
by (asm_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym, real_mult_assoc RS sym]) 1);
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1589
by (stac (CLAIM "6 = 2 * (3::real)") 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1590
by (rtac real_mult_less_mono 1); (*mult_strict_mono would be stronger*)  
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1591
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc]));
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1592
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1593
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1594
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1595
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1596
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1597
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1598
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1599
by (stac real_of_nat_mult 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1600
by (simp_tac (simpset() addsimps [real_divide_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1601
    inverse_mult_distrib] delsimps [fact_Suc]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1602
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1603
    delsimps [fact_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1604
by (multr_by_tac "real (Suc (Suc (4*m)))" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1605
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1606
    delsimps [fact_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1607
by (multr_by_tac "real (Suc (Suc (Suc (4*m))))" 1);
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1608
by (auto_tac (claset(),simpset() addsimps [mult_assoc,mult_less_cancel_left] 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1609
    delsimps [fact_Suc]));
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1610
by (auto_tac (claset(),simpset() addsimps [
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1611
        CLAIM "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * (x::real))"] 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1612
    delsimps [fact_Suc]));
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1613
by (subgoal_tac "0 < x ^ (4 * m)" 1);
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1614
by (asm_simp_tac (simpset() addsimps [realpow_gt_zero]) 2); 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1615
by (asm_simp_tac (simpset() addsimps [mult_less_cancel_left]) 1); 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1616
by (rtac real_mult_less_mono 1); (*mult_strict_mono would be stronger*)
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1617
by (ALLGOALS(Asm_simp_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1618
by (TRYALL(rtac real_less_trans));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1619
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1620
by (res_inst_tac [("y","0")] order_less_le_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1621
by (ALLGOALS(Asm_simp_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1622
qed "sin_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1623
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1624
Goal "[|0 < x; x < 2 |] ==> 0 < sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1625
by (auto_tac (claset() addIs [sin_gt_zero],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1626
qed "sin_gt_zero1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1627
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1628
Goal "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1629
by (auto_tac (claset(),
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1630
    simpset() addsimps [cos_squared_eq,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1631
    minus_add_distrib RS sym, neg_less_0_iff_less,sin_gt_zero1,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1632
    real_add_order,realpow_gt_zero,cos_double] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1633
    delsimps [realpow_Suc,minus_add_distrib]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1634
qed "cos_double_less_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1635
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1636
Goal  "(%n. (- 1) ^ n /(real (fact (2 * n))) * x ^ (2 * n)) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1637
\      sums cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1638
by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((cos_converges 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1639
    RS sums_summable) RS sums_group)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1640
by (auto_tac (claset(),simpset() addsimps  mult_ac@[cos_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1641
qed "cos_paired";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1642
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1643
Addsimps [realpow_gt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1644
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1645
Goal "cos (2) < 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1646
by (cut_inst_tac [("x","2")] cos_paired 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1647
by (dtac sums_minus 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1648
by (rtac (CLAIM "- x < -y ==> (y::real) < x") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1649
by (ftac sums_unique 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1650
by (res_inst_tac [("R2.0",
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1651
    "sumr 0 (Suc (Suc (Suc 0))) (%n. -((- 1) ^ n /(real (fact(2 * n))) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1652
\               * 2 ^ (2 * n)))")] real_less_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1653
by (simp_tac (simpset() addsimps [fact_num_eq_if,realpow_num_eq_if] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1654
    delsimps [fact_Suc,realpow_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1655
by (simp_tac (simpset() addsimps [real_mult_assoc] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1656
    delsimps [sumr_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1657
by (rtac sumr_pos_lt_pair 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1658
by (etac sums_summable 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1659
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1660
by (simp_tac (simpset() addsimps [real_divide_def,real_mult_assoc RS sym] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1661
    delsimps [fact_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1662
by (rtac real_mult_inverse_cancel2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1663
by (TRYALL(rtac (real_of_nat_fact_gt_zero)));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1664
by (simp_tac (simpset() addsimps [real_mult_assoc RS sym] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1665
    delsimps [fact_Suc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1666
by (rtac ((CLAIM "real(n::nat) * 4 = real(4 * n)") RS ssubst) 1);
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1667
by (stac fact_Suc 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1668
by (stac real_of_nat_mult 1);
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1669
by (stac real_of_nat_mult 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1670
by (rtac real_mult_less_mono 1); (*mult_strict_mono would be stronger*)
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1671
by (Force_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1672
by (Force_tac 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1673
by (rtac real_of_nat_fact_gt_zero 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1674
by (rtac (real_of_nat_less_iff RS iffD2) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1675
by (rtac fact_less_mono 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1676
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1677
qed "cos_two_less_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1678
Addsimps [cos_two_less_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1679
Addsimps [cos_two_less_zero RS real_not_refl2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1680
Addsimps [cos_two_less_zero RS order_less_imp_le];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1681
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1682
Goal "EX! x. 0 <= x & x <= 2 & cos x = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1683
by (subgoal_tac "EX x.  0 <= x & x <= 2 & cos x = 0" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1684
by (rtac IVT2  2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1685
by (auto_tac (claset() addIs [DERIV_isCont,DERIV_cos],simpset ()));
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1686
by (cut_inst_tac [("x","xa"),("y","y")] linorder_less_linear 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1687
by (rtac ccontr 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1688
by (subgoal_tac "(ALL x. cos differentiable x) & \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1689
\   (ALL x. isCont cos x)" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1690
by (auto_tac (claset() addIs [DERIV_cos,DERIV_isCont],simpset() 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1691
    addsimps [differentiable_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1692
by (dres_inst_tac [("f","cos")] Rolle 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1693
by (dres_inst_tac [("f","cos")] Rolle 5);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1694
by (auto_tac (claset() addSDs [DERIV_cos RS DERIV_unique],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1695
    simpset() addsimps [differentiable_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1696
by (dres_inst_tac [("y1","xa")] (order_le_less_trans RS sin_gt_zero) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1697
by (assume_tac 1 THEN rtac order_less_le_trans 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1698
by (dres_inst_tac [("y1","y")] (order_le_less_trans RS sin_gt_zero) 4);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1699
by (assume_tac 4 THEN rtac order_less_le_trans 4);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1700
by (assume_tac 1 THEN assume_tac 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1701
by (ALLGOALS (Asm_full_simp_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1702
qed "cos_is_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1703
    
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1704
Goalw [pi_def] "pi/2 = (@x. 0 <= x & x <= 2 & cos x = 0)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1705
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1706
qed "pi_half";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1707
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1708
Goal "cos (pi / 2) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1709
by (rtac (cos_is_zero RS ex1E) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1710
by (auto_tac (claset() addSIs [someI2],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1711
    simpset() addsimps [pi_half]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1712
qed "cos_pi_half";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1713
Addsimps [cos_pi_half];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1714
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1715
Goal "0 < pi / 2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1716
by (rtac (cos_is_zero RS ex1E) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1717
by (auto_tac (claset(),simpset() addsimps [pi_half]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1718
by (rtac someI2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1719
by (Blast_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1720
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1721
by (dres_inst_tac [("y","xa")] real_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1722
by (Step_tac 1 THEN Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1723
qed "pi_half_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1724
Addsimps [pi_half_gt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1725
Addsimps [(pi_half_gt_zero RS real_not_refl2) RS not_sym];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1726
Addsimps [pi_half_gt_zero RS order_less_imp_le];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1727
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1728
Goal "pi / 2 < 2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1729
by (rtac (cos_is_zero RS ex1E) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1730
by (auto_tac (claset(),simpset() addsimps [pi_half]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1731
by (rtac someI2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1732
by (Blast_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1733
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1734
by (dres_inst_tac [("x","xa")] order_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1735
by (Step_tac 1 THEN Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1736
qed "pi_half_less_two";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1737
Addsimps [pi_half_less_two];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1738
Addsimps [pi_half_less_two RS real_not_refl2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1739
Addsimps [pi_half_less_two RS order_less_imp_le];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1740
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1741
Goal "0 < pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1742
by (multr_by_tac "inverse 2" 1);
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1743
by (Simp_tac 1);
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1744
by (cut_facts_tac [pi_half_gt_zero] 1);
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1745
by (full_simp_tac (HOL_ss addsimps [thm"mult_left_zero", real_divide_def]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1746
qed "pi_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1747
Addsimps [pi_gt_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1748
Addsimps [(pi_gt_zero RS real_not_refl2) RS not_sym];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1749
Addsimps [pi_gt_zero RS CLAIM "(x::real) < y ==> ~ y < x"];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1750
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1751
Goal "0 <= pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1752
by (auto_tac (claset() addIs [order_less_imp_le],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1753
qed "pi_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1754
Addsimps [pi_ge_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1755
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1756
Goal "-(pi/2) < 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1757
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1758
qed "minus_pi_half_less_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1759
Addsimps [minus_pi_half_less_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1760
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1761
Goal "sin(pi/2) = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1762
by (cut_inst_tac [("x","pi/2")] sin_cos_squared_add2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1763
by (cut_facts_tac [[pi_half_gt_zero,pi_half_less_two] MRS sin_gt_zero] 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1764
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2])); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1765
qed "sin_pi_half";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1766
Addsimps [sin_pi_half];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1767
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1768
Goal "cos pi = - 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1769
by (cut_inst_tac [("x","pi/2"),("y","pi/2")] cos_add 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1770
by (Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1771
qed "cos_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1772
Addsimps [cos_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1773
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1774
Goal "sin pi = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1775
by (cut_inst_tac [("x","pi/2"),("y","pi/2")] sin_add 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1776
by (Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1777
qed "sin_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1778
Addsimps [sin_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1779
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1780
Goalw [real_diff_def] "sin x = cos (pi/2 - x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1781
by (simp_tac (simpset() addsimps [cos_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1782
qed "sin_cos_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1783
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1784
Goal "-sin x = cos (x + pi/2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1785
by (simp_tac (simpset() addsimps [cos_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1786
qed "minus_sin_cos_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1787
Addsimps [minus_sin_cos_eq RS sym];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1788
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1789
Goalw [real_diff_def] "cos x = sin (pi/2 - x)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1790
by (simp_tac (simpset() addsimps [sin_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1791
qed "cos_sin_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1792
Addsimps [sin_cos_eq RS sym, cos_sin_eq RS sym];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1793
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1794
Goal "sin (x + pi) = - sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1795
by (simp_tac (simpset() addsimps [sin_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1796
qed "sin_periodic_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1797
Addsimps [sin_periodic_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1798
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1799
Goal "sin (pi + x) = - sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1800
by (simp_tac (simpset() addsimps [sin_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1801
qed "sin_periodic_pi2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1802
Addsimps [sin_periodic_pi2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1803
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1804
Goal "cos (x + pi) = - cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1805
by (simp_tac (simpset() addsimps [cos_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1806
qed "cos_periodic_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1807
Addsimps [cos_periodic_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1808
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1809
Goal "sin (x + 2*pi) = sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1810
by (simp_tac (simpset() addsimps [sin_add,cos_double,numeral_2_eq_2]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1811
  (*FIXME: just needs x^n for literals!*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1812
qed "sin_periodic";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1813
Addsimps [sin_periodic];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1814
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1815
Goal "cos (x + 2*pi) = cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1816
by (simp_tac (simpset() addsimps [cos_add,cos_double,numeral_2_eq_2]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1817
  (*FIXME: just needs x^n for literals!*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1818
qed "cos_periodic";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1819
Addsimps [cos_periodic];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1820
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1821
Goal "cos (real n * pi) = (-(1::real)) ^ n";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1822
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1823
by (auto_tac (claset(),simpset() addsimps 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1824
    [real_of_nat_Suc,left_distrib]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1825
qed "cos_npi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1826
Addsimps [cos_npi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1827
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1828
Goal "sin (real (n::nat) * pi) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1829
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1830
by (auto_tac (claset(),simpset() addsimps 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1831
    [real_of_nat_Suc,left_distrib]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1832
qed "sin_npi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1833
Addsimps [sin_npi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1834
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1835
Goal "sin (pi * real (n::nat)) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1836
by (cut_inst_tac [("n","n")] sin_npi 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1837
by (auto_tac (claset(),simpset() addsimps [real_mult_commute] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1838
    delsimps [sin_npi]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1839
qed "sin_npi2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1840
Addsimps [sin_npi2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1841
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1842
Goal "cos (2 * pi) = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1843
by (simp_tac (simpset() addsimps [cos_double,numeral_2_eq_2]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1844
  (*FIXME: just needs x^n for literals!*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1845
qed "cos_two_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1846
Addsimps [cos_two_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1847
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1848
Goal "sin (2 * pi) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1849
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1850
qed "sin_two_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1851
Addsimps [sin_two_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1852
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1853
Goal "[| 0 < x; x < pi/2 |] ==> 0 < sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1854
by (rtac sin_gt_zero 1);
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1855
by (assume_tac 1); 
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1856
by (rtac real_less_trans 1 THEN assume_tac 1);
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1857
by (rtac pi_half_less_two 1); 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1858
qed "sin_gt_zero2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1859
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1860
Goal "[| - pi/2 < x; x < 0 |] ==> sin x < 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1861
by (rtac (CLAIM "(0::real) < - x ==> x < 0") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1862
by (rtac (sin_minus RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1863
by (rtac sin_gt_zero2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1864
by (rtac (CLAIM "-y < x ==> -x < (y::real)") 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1865
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1866
qed "sin_less_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1867
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1868
Goal "pi < 4";
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1869
by (cut_facts_tac [pi_half_less_two] 1);
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1870
by Auto_tac; 
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1871
qed "pi_less_4";
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1872
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1873
Goal "[| 0 < x; x < pi/2 |] ==> 0 < cos x";
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1874
by (cut_facts_tac [pi_less_4] 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1875
by (cut_inst_tac [("f","cos"),("a","0"),("b","x"),("y","0")] IVT2_objl 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1876
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1877
by (cut_facts_tac [cos_is_zero] 5);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1878
by (Step_tac 5);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1879
by (dres_inst_tac [("x","xa")] spec 5);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1880
by (dres_inst_tac [("x","pi/2")] spec 5);
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1881
by (force_tac (claset(), simpset() addsimps []) 1); 
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1882
by (force_tac (claset(), simpset() addsimps []) 1); 
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1883
by (force_tac (claset(), simpset() addsimps []) 1); 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1884
by (auto_tac (claset() addSDs [ pi_half_less_two RS order_less_trans, 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1885
    CLAIM "~ m <= n ==> n < (m::real)"]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1886
    addIs [DERIV_isCont,DERIV_cos],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1887
qed "cos_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1888
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1889
Goal "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1890
by (res_inst_tac [("R1.0","x"),("R2.0","0")] real_linear_less2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1891
by (rtac (cos_minus RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1892
by (rtac cos_gt_zero 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1893
by (rtac (CLAIM "-y < x ==> -x < (y::real)") 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1894
by (auto_tac (claset() addIs [cos_gt_zero],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1895
qed "cos_gt_zero_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1896
 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1897
Goal "[| -(pi/2) <= x; x <= pi/2 |] ==> 0 <= cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1898
by (auto_tac (claset(),HOL_ss addsimps [real_le_less,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1899
    cos_gt_zero_pi]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1900
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1901
qed "cos_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1902
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1903
Goal "[| 0 < x; x < pi  |] ==> 0 < sin x";
12486
0ed8bdd883e0 isatool expandshort;
wenzelm
parents: 12481
diff changeset
  1904
by (stac sin_cos_eq 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1905
by (rotate_tac 1 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1906
by (dtac (real_sum_of_halves RS ssubst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1907
by (auto_tac (claset() addSIs [cos_gt_zero_pi],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1908
    simpset() delsimps [sin_cos_eq RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1909
qed "sin_gt_zero_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1910
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1911
Goal "[| 0 <= x; x <= pi |] ==> 0 <= sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1912
by (auto_tac (claset(),simpset() addsimps [real_le_less,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1913
    sin_gt_zero_pi]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1914
qed "sin_ge_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1915
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1916
Goal "[| - 1 <= y; y <= 1 |] ==> EX! x. 0 <= x & x <= pi & (cos x = y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1917
by (subgoal_tac "EX x.  0 <= x & x <= pi & cos x = y" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1918
by (rtac IVT2  2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1919
by (auto_tac (claset() addIs [order_less_imp_le,DERIV_isCont,DERIV_cos],
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1920
    simpset ()));
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  1921
by (cut_inst_tac [("x","xa"),("y","y")] linorder_less_linear 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1922
by (rtac ccontr 1 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1923
by (dres_inst_tac [("f","cos")] Rolle 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1924
by (dres_inst_tac [("f","cos")] Rolle 5);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1925
by (auto_tac (claset() addIs [order_less_imp_le,DERIV_isCont,DERIV_cos]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1926
    addSDs [DERIV_cos RS DERIV_unique],simpset() addsimps [differentiable_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1927
by (auto_tac (claset() addDs [[order_le_less_trans,order_less_le_trans] MRS 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1928
    sin_gt_zero_pi],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1929
qed "cos_total";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1930
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1931
Goal "[| - 1 <= y; y <= 1 |] ==> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1932
\     EX! x. -(pi/2) <= x & x <= pi/2 & (sin x = y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1933
by (rtac ccontr 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1934
by (subgoal_tac "ALL x. (-(pi/2) <= x & x <= pi/2 & (sin x = y)) \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1935
\   = (0 <= (x + pi/2) & (x + pi/2) <= pi & \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1936
\     (cos(x + pi/2) = -y))" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1937
by (etac swap 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1938
by (asm_full_simp_tac (simpset() delsimps [minus_sin_cos_eq RS sym]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1939
by (dtac (CLAIM "(x::real) <= y ==> -y <= -x") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1940
by (dtac (CLAIM "(x::real) <= y ==> -y <= -x") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1941
by (dtac cos_total 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1942
by (Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1943
by (etac ex1E 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1944
by (res_inst_tac [("a","x - (pi/2)")] ex1I 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1945
by (simp_tac (simpset() addsimps [real_add_assoc]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1946
by (rotate_tac 3 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1947
by (dres_inst_tac [("x","xa + pi/2")] spec 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1948
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1949
by (TRYALL(Asm_full_simp_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1950
by (auto_tac (claset(),simpset() addsimps [CLAIM "(-x <= y) = (-y <= (x::real))"]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1951
qed "sin_total";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1952
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1953
Goal "(EX n. P (n::nat)) = (EX n. P n & (ALL m. m < n --> ~ P m))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1954
by (rtac iffI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1955
by (rtac contrapos_pp 1 THEN assume_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1956
by (EVERY1[Simp_tac, rtac allI, rtac nat_less_induct]);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1957
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1958
qed "less_induct_ex_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1959
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1960
Goal "[| 0 < y; 0 <= x |] ==> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1961
\     EX n. real n * y <= x & x < real (Suc n) * y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1962
by (auto_tac (claset() addSDs [reals_Archimedean3],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1963
by (dres_inst_tac [("x","x")] spec 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1964
by (dtac (less_induct_ex_iff RS iffD1) 1 THEN Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1965
by (case_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1966
by (res_inst_tac [("x","nat")] exI 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1967
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1968
qed "reals_Archimedean4";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1969
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1970
(* Pre Isabelle99-2 proof was simpler- numerals arithmetic 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1971
   now causes some unwanted re-arrangements of literals!   *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1972
Goal "[| 0 <= x; cos x = 0 |] ==> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1973
\     EX n. ~even n & x = real n * (pi/2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1974
by (dtac (pi_gt_zero RS reals_Archimedean4) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1975
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1976
by (subgoal_tac 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1977
    "0 <= x - real n * pi & (x - real n * pi) <= pi & \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1978
\    (cos(x - real n * pi) = 0)" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1979
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1980
by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1981
    left_distrib]) 2);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1982
by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1983
by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1984
by (subgoal_tac "EX! x. 0 <= x & x <= pi & cos x = 0" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1985
by (rtac cos_total 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1986
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1987
by (dres_inst_tac [("x","x - real n * pi")] spec 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1988
by (dres_inst_tac [("x","pi/2")] spec 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1989
by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1990
by (res_inst_tac [("x","Suc (2 * n)")] exI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1991
by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  1992
    left_distrib]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1993
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1994
qed "cos_zero_lemma";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1995
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1996
Goal "[| 0 <= x; sin x = 0 |] ==> \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1997
\     EX n. even n & x = real n * (pi/2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1998
by (subgoal_tac 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  1999
    "EX n. ~ even n & x + pi/2  = real n * (pi/2)" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2000
by (rtac cos_zero_lemma 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2001
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2002
by (res_inst_tac [("x","n - 1")] exI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2003
by (rtac (CLAIM "-y <= x ==> -x <= (y::real)") 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2004
by (rtac real_le_trans 2 THEN assume_tac 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2005
by (auto_tac (claset(),simpset() addsimps [odd_not_even RS sym,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2006
    odd_Suc_mult_two_ex,real_of_nat_Suc,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2007
    left_distrib,real_mult_assoc RS sym]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2008
qed "sin_zero_lemma";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2009
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2010
(* also spoilt by numeral arithmetic *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2011
Goal "(cos x = 0) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2012
\     ((EX n. ~even n & (x = real n * (pi/2))) |   \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2013
\      (EX n. ~even n & (x = -(real n * (pi/2)))))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2014
by (rtac iffI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2015
by (cut_inst_tac [("x","x")] (CLAIM "0 <= (x::real) | x <= 0") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2016
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2017
by (dtac cos_zero_lemma 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2018
by (dtac (CLAIM "(x::real) <= 0 ==> 0 <= -x") 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2019
by (dtac cos_zero_lemma 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2020
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2021
by (dtac (CLAIM "-x = y ==> x = -(y::real)") 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2022
by (auto_tac (claset(),HOL_ss addsimps [odd_not_even RS sym,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2023
    odd_Suc_mult_two_ex,real_of_nat_Suc,left_distrib]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2024
by (auto_tac (claset(),simpset() addsimps [cos_add]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2025
qed "cos_zero_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2026
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2027
(* ditto: but to a lesser extent *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2028
Goal "(sin x = 0) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2029
\     ((EX n. even n & (x = real n * (pi/2))) |   \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2030
\      (EX n. even n & (x = -(real n * (pi/2)))))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2031
by (rtac iffI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2032
by (cut_inst_tac [("x","x")] (CLAIM "0 <= (x::real) | x <= 0") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2033
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2034
by (dtac sin_zero_lemma 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2035
by (dtac (CLAIM "(x::real) <= 0 ==> 0 <= -x") 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2036
by (dtac sin_zero_lemma 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2037
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2038
by (dtac (CLAIM "-x = y ==> x = -(y::real)") 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2039
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2040
qed "sin_zero_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2041
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2042
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2043
(* Tangent                                                                  *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2044
(* ------------------------------------------------------------------------ *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2045
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2046
Goalw [tan_def] "tan 0 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2047
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2048
qed "tan_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2049
Addsimps [tan_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2050
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2051
Goalw [tan_def] "tan pi = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2052
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2053
qed "tan_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2054
Addsimps [tan_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2055
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2056
Goalw [tan_def] "tan (real (n::nat) * pi) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2057
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2058
qed "tan_npi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2059
Addsimps [tan_npi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2060
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2061
Goalw [tan_def] "tan (-x) = - tan x";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2062
by (simp_tac (simpset() addsimps [minus_mult_left]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2063
qed "tan_minus";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2064
Addsimps [tan_minus];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2065
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2066
Goalw [tan_def] "tan (x + 2*pi) = tan x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2067
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2068
qed "tan_periodic";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2069
Addsimps [tan_periodic];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2070
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2071
Goalw [tan_def,real_divide_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2072
      "[| cos x ~= 0; cos y ~= 0 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2073
\       ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)";
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  2074
by (auto_tac (claset(),
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  2075
   simpset() delsimps [thm"Ring_and_Field.inverse_mult_distrib"]
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2076
	     addsimps [inverse_mult_distrib RS sym] @ mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2077
by (res_inst_tac [("c1","cos x * cos y")] (real_mult_right_cancel RS subst) 1);
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  2078
by (auto_tac (claset(), 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  2079
    simpset() delsimps [thm"Ring_and_Field.inverse_mult_distrib"]
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2080
	      addsimps [mult_assoc, left_diff_distrib,cos_add]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2081
val lemma_tan_add1 = result();
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2082
Addsimps [lemma_tan_add1];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2083
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2084
Goalw [tan_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2085
      "[| cos x ~= 0; cos y ~= 0 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2086
\      ==> tan x + tan y = sin(x + y)/(cos x * cos y)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2087
by (res_inst_tac [("c1","cos x * cos y")] (real_mult_right_cancel RS subst) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2088
by (auto_tac (claset(), simpset() addsimps [mult_assoc, left_distrib]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2089
by (simp_tac (simpset() addsimps [sin_add]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2090
qed "add_tan_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2091
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2092
Goal "[| cos x ~= 0; cos y ~= 0; cos (x + y) ~= 0 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2093
\     ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2094
by (asm_simp_tac (simpset() addsimps [add_tan_eq]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2095
by (simp_tac (simpset() addsimps [tan_def]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2096
qed "tan_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2097
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2098
Goal "[| cos x ~= 0; cos (2 * x) ~= 0 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2099
\     ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2100
by (auto_tac (claset(),simpset() addsimps [asm_full_simplify 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2101
    (simpset() addsimps [real_mult_2 RS sym] delsimps [lemma_tan_add1]) 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2102
    (read_instantiate [("x","x"),("y","x")] tan_add),numeral_2_eq_2]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2103
    delsimps [lemma_tan_add1]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2104
qed "tan_double";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2105
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2106
Goalw [tan_def,real_divide_def] "[| 0 < x; x < pi/2 |] ==> 0 < tan x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2107
by (auto_tac (claset() addSIs [sin_gt_zero2,cos_gt_zero_pi]
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2108
    addSIs [real_mult_order, positive_imp_inverse_positive],simpset()));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2109
qed "tan_gt_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2110
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2111
Goal "[| - pi/2 < x; x < 0 |] ==> tan x < 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2112
by (rtac (CLAIM "(0::real) < - x ==> x < 0") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2113
by (rtac (tan_minus RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2114
by (rtac tan_gt_zero 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2115
by (rtac (CLAIM "-x < y ==> -y < (x::real)") 2 THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2116
qed "tan_less_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2117
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2118
Goal "cos x ~= 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse(cos x ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2119
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2120
by DERIV_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2121
by (auto_tac (claset(),simpset() addsimps [real_divide_def,numeral_2_eq_2]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2122
qed "lemma_DERIV_tan";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2123
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2124
Goal "cos x ~= 0 ==> DERIV tan x :> inverse(cos(x) ^ 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2125
by (auto_tac (claset() addDs [lemma_DERIV_tan],simpset()
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2126
    addsimps [(tan_def RS meta_eq_to_obj_eq) RS sym]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2127
qed "DERIV_tan";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2128
Addsimps [DERIV_tan];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2129
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2130
Goalw [real_divide_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2131
      "(%x. cos(x)/sin(x)) -- pi/2 --> 0";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2132
by (res_inst_tac [("a1","1")] ((mult_left_zero) RS subst) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2133
by (rtac LIM_mult2 1); 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2134
by (rtac (inverse_1 RS subst) 2);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2135
by (rtac LIM_inverse 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2136
by (fold_tac [real_divide_def]);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2137
by (auto_tac (claset() addSIs [DERIV_isCont],simpset() 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2138
    addsimps [(isCont_def RS meta_eq_to_obj_eq)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2139
    RS sym, cos_pi_half RS sym, sin_pi_half RS sym] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2140
    delsimps [cos_pi_half,sin_pi_half]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2141
by (DERIV_tac THEN Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2142
qed "LIM_cos_div_sin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2143
Addsimps [LIM_cos_div_sin];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2144
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2145
Goal "0 < y ==> EX x. 0 < x & x < pi/2 & y < tan x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2146
by (cut_facts_tac [LIM_cos_div_sin] 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2147
by (asm_full_simp_tac (HOL_ss addsimps [LIM_def]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2148
by (dres_inst_tac [("x","inverse y")] spec 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2149
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2150
by (Force_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2151
by (dres_inst_tac [("d1.0","s")] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2152
    (pi_half_gt_zero RSN (2,real_lbound_gt_zero)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2153
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2154
by (res_inst_tac [("x","(pi/2) - e")] exI 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2155
by (Asm_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2156
by (dres_inst_tac [("x","(pi/2) - e")] spec 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2157
by (auto_tac (claset(),simpset() addsimps [abs_eqI2,tan_def]));
14309
f508492af9b4 moving HyperArith0.ML to other theories
paulson
parents: 14305
diff changeset
  2158
by (rtac (inverse_less_iff_less RS iffD1) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2159
by (auto_tac (claset(),simpset() addsimps [real_divide_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2160
by (rtac (real_mult_order) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2161
by (subgoal_tac "0 < sin e" 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2162
by (subgoal_tac "0 < cos e" 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2163
by (auto_tac (claset() addIs [cos_gt_zero,sin_gt_zero2],simpset()
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2164
    addsimps [inverse_mult_distrib,abs_mult]));
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2165
by (dres_inst_tac [("a","cos e")] (positive_imp_inverse_positive) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2166
by (dres_inst_tac [("x","inverse (cos e)")] abs_eqI2 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2167
by (auto_tac (claset() addSDs [abs_eqI2],simpset() addsimps mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2168
qed "lemma_tan_total";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2169
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2170
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2171
Goal "0 <= y ==> EX x. 0 <= x & x < pi/2 & tan x = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2172
by (ftac real_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2173
by (Step_tac 1 THEN Force_tac 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2174
by (dtac lemma_tan_total 1 THEN Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2175
by (cut_inst_tac [("f","tan"),("a","0"),("b","x"),("y","y")] IVT_objl 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2176
by (auto_tac (claset() addSIs [DERIV_tan RS DERIV_isCont],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2177
by (dres_inst_tac [("y","xa")] order_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2178
by (auto_tac (claset() addDs [cos_gt_zero],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2179
qed "tan_total_pos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2180
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2181
Goal "EX x. -(pi/2) < x & x < (pi/2) & tan x = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2182
by (cut_inst_tac [("y","y")] (CLAIM "0 <= (y::real) | 0 <= -y") 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2183
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2184
by (dtac tan_total_pos 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2185
by (dtac tan_total_pos 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2186
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2187
by (res_inst_tac [("x","-x")] exI 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2188
by (auto_tac (claset() addSIs [exI],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2189
qed "lemma_tan_total1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2190
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2191
Goal "EX! x. -(pi/2) < x & x < (pi/2) & tan x = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2192
by (cut_inst_tac [("y","y")] lemma_tan_total1 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2193
by (Auto_tac);
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2194
by (cut_inst_tac [("x","xa"),("y","y")] linorder_less_linear 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2195
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2196
by (subgoal_tac "EX z. xa < z & z < y & DERIV tan z :> 0" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2197
by (subgoal_tac "EX z. y < z & z < xa & DERIV tan z :> 0" 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2198
by (rtac Rolle 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2199
by (rtac Rolle 7);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2200
by (auto_tac (claset() addSIs [DERIV_tan,DERIV_isCont,exI],simpset()
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2201
    addsimps [differentiable_def]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2202
by (TRYALL(rtac DERIV_tan));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2203
by (TRYALL(dtac (DERIV_tan RSN (2,DERIV_unique))));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2204
by (TRYALL(rtac (real_not_refl2 RS not_sym)));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2205
by (auto_tac (claset() addSIs [cos_gt_zero_pi],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2206
by (ALLGOALS(subgoal_tac "0 < cos z")); 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2207
by (Force_tac 1 THEN Force_tac 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2208
by (ALLGOALS(thin_tac "cos z = 0"));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2209
by (auto_tac (claset() addSIs [cos_gt_zero_pi],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2210
qed "tan_total";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2211
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2212
Goal "[| - 1 <= y; y <= 1 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2213
\     ==> -(pi/2) <= arcsin y & arcsin y <= pi & sin(arcsin y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2214
by (dtac sin_total 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2215
by (etac ex1E 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2216
by (rewtac arcsin_def);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2217
by (rtac someI2 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2218
by (EVERY1[assume_tac, Blast_tac, Step_tac]);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2219
by (rtac real_le_trans 1 THEN assume_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2220
by (Force_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2221
qed "arcsin_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2222
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2223
Goal "[| - 1 <= y; y <= 1 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2224
\     ==> -(pi/2) <= arcsin y & \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2225
\          arcsin y <= pi/2 & sin(arcsin y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2226
by (dtac sin_total 1 THEN assume_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2227
by (etac ex1E 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2228
by (rewtac arcsin_def);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2229
by (rtac someI2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2230
by (ALLGOALS(Blast_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2231
qed "arcsin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2232
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2233
Goal "[| - 1 <= y; y <= 1 |] ==> sin(arcsin y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2234
by (blast_tac (claset() addDs [arcsin]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2235
qed "sin_arcsin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2236
Addsimps [sin_arcsin];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2237
      
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2238
Goal "[| -1 <= y; y <= 1 |] ==> sin(arcsin y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2239
by (auto_tac (claset() addIs [sin_arcsin],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2240
qed "sin_arcsin2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2241
Addsimps [sin_arcsin2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2242
      
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2243
Goal "[| - 1 <= y; y <= 1 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2244
\     ==> -(pi/2) <= arcsin y & arcsin y <= pi/2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2245
by (blast_tac (claset() addDs [arcsin]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2246
qed "arcsin_bounded";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2247
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2248
Goal "[| - 1 <= y; y <= 1 |] ==> -(pi/2) <= arcsin y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2249
by (blast_tac (claset() addDs [arcsin]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2250
qed "arcsin_lbound";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2251
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2252
Goal "[| - 1 <= y; y <= 1 |] ==> arcsin y <= pi/2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2253
by (blast_tac (claset() addDs [arcsin]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2254
qed "arcsin_ubound";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2255
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2256
Goal "[| - 1 < y; y < 1 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2257
\     ==> -(pi/2) < arcsin y & arcsin y < pi/2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2258
by (ftac order_less_imp_le 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2259
by (forw_inst_tac [("y","y")] order_less_imp_le 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2260
by (ftac arcsin_bounded 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2261
by (Step_tac 1 THEN Asm_full_simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2262
by (dres_inst_tac [("y","arcsin y")] order_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2263
by (dres_inst_tac [("y","pi/2")] order_le_imp_less_or_eq 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2264
by (Step_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2265
by (ALLGOALS(dres_inst_tac [("f","sin")] arg_cong));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2266
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2267
qed "arcsin_lt_bounded";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2268
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2269
Goalw [arcsin_def] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2270
  "[|-(pi/2) <= x; x <= pi/2 |] ==> arcsin(sin x) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2271
by (rtac some1_equality 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2272
by (rtac sin_total 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2273
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2274
qed "arcsin_sin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2275
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2276
Goal "[| - 1 <= y; y <= 1 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2277
\     ==> 0 <= arcos y & arcos y <= pi & cos(arcos y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2278
by (dtac cos_total 1 THEN assume_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2279
by (etac ex1E 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2280
by (rewtac arcos_def);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2281
by (rtac someI2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2282
by (ALLGOALS(Blast_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2283
qed "arcos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2284
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2285
Goal "[| - 1 <= y; y <= 1 |] ==> cos(arcos y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2286
by (blast_tac (claset() addDs [arcos]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2287
qed "cos_arcos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2288
Addsimps [cos_arcos];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2289
      
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2290
Goal "[| -1 <= y; y <= 1 |] ==> cos(arcos y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2291
by (auto_tac (claset() addIs [cos_arcos],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2292
qed "cos_arcos2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2293
Addsimps [cos_arcos2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2294
      
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2295
Goal "[| - 1 <= y; y <= 1 |] ==> 0 <= arcos y & arcos y <= pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2296
by (blast_tac (claset() addDs [arcos]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2297
qed "arcos_bounded";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2298
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2299
Goal "[| - 1 <= y; y <= 1 |] ==> 0 <= arcos y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2300
by (blast_tac (claset() addDs [arcos]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2301
qed "arcos_lbound";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2302
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2303
Goal "[| - 1 <= y; y <= 1 |] ==> arcos y <= pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2304
by (blast_tac (claset() addDs [arcos]) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2305
qed "arcos_ubound";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2306
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2307
Goal "[| - 1 < y; y < 1 |] \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2308
\     ==> 0 < arcos y & arcos y < pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2309
by (ftac order_less_imp_le 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2310
by (forw_inst_tac [("y","y")] order_less_imp_le 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2311
by (ftac arcos_bounded 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2312
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2313
by (dres_inst_tac [("y","arcos y")] order_le_imp_less_or_eq 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2314
by (dres_inst_tac [("y","pi")] order_le_imp_less_or_eq 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2315
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2316
by (ALLGOALS(dres_inst_tac [("f","cos")] arg_cong));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2317
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2318
qed "arcos_lt_bounded";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2319
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2320
Goalw [arcos_def] "[|0 <= x; x <= pi |] ==> arcos(cos x) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2321
by (auto_tac (claset() addSIs [some1_equality,cos_total],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2322
qed "arcos_cos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2323
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2324
Goalw [arcos_def] "[|x <= 0; -pi <= x |] ==> arcos(cos x) = -x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2325
by (auto_tac (claset() addSIs [some1_equality,cos_total],simpset()));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2326
qed "arcos_cos2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2327
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2328
Goal "- (pi/2) < arctan y  & \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2329
\     arctan y < pi/2 & tan (arctan y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2330
by (cut_inst_tac [("y","y")] tan_total 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2331
by (etac ex1E 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2332
by (rewtac arctan_def);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2333
by (rtac someI2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2334
by (ALLGOALS(Blast_tac));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2335
qed "arctan";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2336
Addsimps [arctan];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2337
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2338
Goal "tan(arctan y) = y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2339
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2340
qed "tan_arctan";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2341
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2342
Goal "- (pi/2) < arctan y  & arctan y < pi/2";
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  2343
by (asm_full_simp_tac (HOL_ss addsimps [arctan]) 1); 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2344
qed "arctan_bounded";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2345
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2346
Goal "- (pi/2) < arctan y";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2347
by (Auto_tac);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2348
qed "arctan_lbound";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2349
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2350
Goal "arctan y < pi/2";
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  2351
by (asm_full_simp_tac (HOL_ss addsimps [arctan]) 1); 
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2352
qed "arctan_ubound";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2353
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2354
Goalw [arctan_def]
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2355
      "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2356
by (rtac some1_equality 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2357
by (rtac tan_total 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2358
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2359
qed "arctan_tan";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2360
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2361
Goal "arctan 0 = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2362
by (rtac (asm_full_simplify (simpset()) 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2363
     (read_instantiate [("x","0")] arctan_tan)) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2364
qed "arctan_zero_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2365
Addsimps [arctan_zero_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2366
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2367
(* ------------------------------------------------------------------------- *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2368
(* Differentiation of arctan.                                                *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2369
(* ------------------------------------------------------------------------- *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2370
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2371
Goal "cos(arctan x) ~= 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2372
by (auto_tac (claset(),simpset() addsimps [cos_zero_iff]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2373
by (case_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2374
by (case_tac "n" 3);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2375
by (cut_inst_tac [("y","x")] arctan_ubound 2);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2376
by (cut_inst_tac [("y","x")] arctan_lbound 4);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2377
by (auto_tac (claset(),
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2378
     simpset() addsimps [real_of_nat_Suc, left_distrib,real_le_def, mult_less_0_iff] 
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2379
     delsimps [arctan]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2380
qed "cos_arctan_not_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2381
Addsimps [cos_arctan_not_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2382
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2383
Goal "cos x ~= 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2384
by (rtac (realpow_inverse RS subst) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2385
by (res_inst_tac [("c1","cos(x) ^ 2")] (real_mult_right_cancel RS iffD1) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2386
by (auto_tac (claset() addDs [realpow_not_zero], simpset() addsimps
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2387
    [realpow_mult,left_distrib,realpow_divide,
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2388
     tan_def,real_mult_assoc,realpow_inverse RS sym] 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2389
     delsimps [realpow_Suc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2390
qed "tan_sec";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2391
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2392
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2393
(*--------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2394
(* Some more theorems- developed while at ICASE (07/2001)                   *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2395
(*--------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2396
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2397
Goal "sin (xa + 1 / 2 * real (Suc m) * pi) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2398
\     cos (xa + 1 / 2 * real  (m) * pi)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2399
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2400
    real_of_nat_Suc,left_distrib,right_distrib]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2401
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2402
qed "lemma_sin_cos_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2403
Addsimps [lemma_sin_cos_eq];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2404
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2405
Goal "sin (xa + real (Suc m) * pi / 2) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2406
\     cos (xa + real (m) * pi / 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2407
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2408
    real_of_nat_Suc,left_distrib,right_distrib]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2409
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2410
qed "lemma_sin_cos_eq2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2411
Addsimps [lemma_sin_cos_eq2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2412
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2413
Goal "DERIV (%x. sin (x + k)) xa :> cos (xa + k)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2414
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2415
by (res_inst_tac [("f","sin"),("g","%x. x + k")] DERIV_chain2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2416
by DERIV_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2417
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2418
qed "DERIV_sin_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2419
Addsimps [DERIV_sin_add];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2420
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2421
(* which further simplifies to (- 1 ^ m) !! *)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2422
Goal "sin ((real m + 1/2) * pi) = cos (real m * pi)";
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2423
by (auto_tac (claset(),simpset() addsimps [right_distrib,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2424
    sin_add,left_distrib] @ mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2425
qed "sin_cos_npi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2426
Addsimps [sin_cos_npi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2427
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2428
Goal "sin (real (Suc (2 * n)) * pi / 2) = (- 1) ^ n";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2429
by (cut_inst_tac [("m","n")] sin_cos_npi 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2430
by (auto_tac (claset(),HOL_ss addsimps [real_of_nat_Suc,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2431
    left_distrib,real_divide_def]));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2432
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2433
qed "sin_cos_npi2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2434
Addsimps [ sin_cos_npi2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2435
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2436
Goal "cos (2 * real (n::nat) * pi) = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2437
by (auto_tac (claset(),simpset() addsimps [cos_double,
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2438
    real_mult_assoc,realpow_add RS sym,numeral_2_eq_2]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2439
  (*FIXME: just needs x^n for literals!*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2440
qed "cos_2npi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2441
Addsimps [cos_2npi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2442
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2443
Goal "cos (3 / 2 * pi) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2444
by (rtac (CLAIM "(1::real) + 1/2 = 3/2" RS subst) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2445
by (stac left_distrib 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2446
by (auto_tac (claset(),simpset() addsimps [cos_add] @ mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2447
qed "cos_3over2_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2448
Addsimps [cos_3over2_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2449
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2450
Goal "sin (2 * real (n::nat) * pi) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2451
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2452
qed "sin_2npi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2453
Addsimps [sin_2npi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2454
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2455
Goal "sin (3 / 2 * pi) = - 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2456
by (rtac (CLAIM "(1::real) + 1/2 = 3/2" RS subst) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2457
by (stac left_distrib 1);
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2458
by (auto_tac (claset(),simpset() addsimps [sin_add] @mult_ac));
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2459
qed "sin_3over2_pi";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2460
Addsimps [sin_3over2_pi];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2461
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2462
Goal "cos(xa + 1 / 2 * real (Suc m) * pi) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2463
\     -sin  (xa + 1 / 2 * real (m) * pi)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2464
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2465
    real_of_nat_Suc,right_distrib,left_distrib,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2466
    minus_mult_right]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2467
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2468
qed "lemma_cos_sin_eq";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2469
Addsimps [lemma_cos_sin_eq];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2470
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2471
Goal "cos (xa + real (Suc m) * pi / 2) = \
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2472
\     -sin (xa + real (m) * pi / 2)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2473
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2474
    real_of_nat_Suc,left_distrib,right_distrib]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2475
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2476
qed "lemma_cos_sin_eq2";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2477
Addsimps [lemma_cos_sin_eq2];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2478
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2479
Goal "cos (pi * real (Suc (2 * m)) / 2) = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2480
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def,
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2481
    real_of_nat_Suc,left_distrib,right_distrib]) 1);
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2482
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2483
qed "cos_pi_eq_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2484
Addsimps [cos_pi_eq_zero];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2485
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2486
Goal "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2487
by (rtac lemma_DERIV_subst 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2488
by (res_inst_tac [("f","cos"),("g","%x. x + k")] DERIV_chain2 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2489
by DERIV_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2490
by (Simp_tac 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2491
qed "DERIV_cos_add";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2492
Addsimps [DERIV_cos_add];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2493
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2494
Goal "isCont cos x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2495
by (rtac (DERIV_cos RS DERIV_isCont) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2496
qed "isCont_cos";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2497
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2498
Goal "isCont sin x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2499
by (rtac (DERIV_sin RS DERIV_isCont) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2500
qed "isCont_sin";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2501
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2502
Goal "isCont exp x";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2503
by (rtac (DERIV_exp RS DERIV_isCont) 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2504
qed "isCont_exp";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2505
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2506
val isCont_simp = [isCont_exp,isCont_sin,isCont_cos];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2507
Addsimps isCont_simp;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2508
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2509
(** more theorems: e.g. used in complex geometry  **)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2510
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2511
Goal "sin x = 0 ==> abs(cos x) = 1";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2512
by (auto_tac (claset(),simpset() addsimps [sin_zero_iff,even_mult_two_ex]));
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2513
qed "sin_zero_abs_cos_one";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2514
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2515
Goal "(exp x = 1) = (x = 0)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2516
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2517
by (dres_inst_tac [("f","ln")] arg_cong 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2518
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2519
qed "exp_eq_one_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2520
Addsimps [exp_eq_one_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2521
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2522
Goal "cos x = 1 ==> sin x = 0";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2523
by (cut_inst_tac [("x","x")] sin_cos_squared_add3 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2524
by Auto_tac;
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2525
qed "cos_one_sin_zero";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2526
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2527
(*-------------------------------------------------------------------------------*)
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2528
(* A few extra theorems                                                          *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2529
(*-------------------------------------------------------------------------------*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2530
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2531
Goal "[| 0 <= x; x < y |] ==> root(Suc n) x < root(Suc n) y";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2532
by (ftac order_le_less_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2533
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2534
by (forw_inst_tac [("n1","n")] (real_root_pow_pos2 RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2535
by (rotate_tac 1 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2536
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2537
by (forw_inst_tac [("n1","n")] (real_root_pow_pos RS ssubst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2538
by (rotate_tac 3 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2539
by (dres_inst_tac [("y","root (Suc n) y ^ Suc n")] order_less_imp_le 1 );
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2540
by (forw_inst_tac [("n","n")] real_root_pos_pos_le 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2541
by (forw_inst_tac [("n","n")] real_root_pos_pos 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2542
by (dres_inst_tac [("x","root (Suc n) x"),
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2543
    ("y","root (Suc n) y")] realpow_increasing 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2544
by (assume_tac 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2545
by (dres_inst_tac [("x","root (Suc n) x")] order_le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2546
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2547
by (dres_inst_tac [("f","%x. x ^ (Suc n)")] arg_cong 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2548
by (auto_tac (claset(),simpset() addsimps [real_root_pow_pos2]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2549
    delsimps [realpow_Suc]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2550
qed "real_root_less_mono";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2551
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2552
Goal "[| 0 <= x; x <= y |] ==> root(Suc n) x <= root(Suc n) y";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2553
by (dres_inst_tac [("y","y")] order_le_imp_less_or_eq 1 );
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2554
by (auto_tac (claset() addDs [real_root_less_mono]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2555
    addIs [order_less_imp_le],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2556
qed "real_root_le_mono";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2557
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2558
Goal "[| 0 <= x; 0 <= y |] ==> (root(Suc n) x < root(Suc n) y) = (x < y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2559
by (auto_tac (claset() addIs [real_root_less_mono],simpset()));
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2560
by (rtac ccontr 1 THEN dtac (linorder_not_less RS iffD1) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2561
by (dres_inst_tac [("x","y"),("n","n")] real_root_le_mono 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2562
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2563
qed "real_root_less_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2564
Addsimps [real_root_less_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2565
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2566
Goal "[| 0 <= x; 0 <= y |] ==> (root(Suc n) x <= root(Suc n) y) = (x <= y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2567
by (auto_tac (claset() addIs [real_root_le_mono],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2568
by (simp_tac (simpset() addsimps [real_le_def]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2569
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2570
by (dres_inst_tac [("x","y"),("n","n")] real_root_less_mono 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2571
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2572
qed "real_root_le_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2573
Addsimps [real_root_le_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2574
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2575
Goal "[| 0 <= x; 0 <= y |] ==> (root(Suc n) x = root(Suc n) y) = (x = y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2576
by (auto_tac (claset() addSIs [real_le_anti_sym],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2577
by (res_inst_tac [("n1","n")] (real_root_le_iff RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2578
by (res_inst_tac [("n1","n")] (real_root_le_iff RS iffD1) 4);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2579
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2580
qed "real_root_eq_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2581
Addsimps [real_root_eq_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2582
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2583
Goal "[| 0 <= x; 0 <= y; y ^ (Suc n) = x |] ==> root (Suc n) x = y";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2584
by (auto_tac (claset() addDs [real_root_pos2],
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2585
    simpset() delsimps [realpow_Suc]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2586
qed "real_root_pos_unique";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2587
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2588
Goal "[| 0 <= x; 0 <= y |]\
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2589
\     ==> root(Suc n) (x * y) = root(Suc n) x * root(Suc n) y";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2590
by (rtac real_root_pos_unique 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2591
by (auto_tac (claset() addSIs [real_root_pos_pos_le],simpset() 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2592
    addsimps [realpow_mult,zero_le_mult_iff,
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2593
    real_root_pow_pos2] delsimps [realpow_Suc]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2594
qed "real_root_mult";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2595
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2596
Goal "0 <= x ==> (root(Suc n) (inverse x) = inverse(root(Suc n) x))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2597
by (rtac real_root_pos_unique 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2598
by (auto_tac (claset() addIs [real_root_pos_pos_le],simpset() 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2599
    addsimps [realpow_inverse RS sym,real_root_pow_pos2] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2600
    delsimps [realpow_Suc]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2601
qed "real_root_inverse";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2602
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2603
Goalw [real_divide_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2604
     "[| 0 <= x; 0 <= y |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2605
\     ==> (root(Suc n) (x / y) = root(Suc n) x / root(Suc n) y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2606
by (auto_tac (claset(),simpset() addsimps [real_root_mult,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2607
    real_root_inverse]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2608
qed "real_root_divide";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2609
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2610
Goalw [sqrt_def] "[| 0 <= x; x < y |] ==> sqrt(x) < sqrt(y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2611
by (auto_tac (claset() addIs [real_root_less_mono],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2612
qed "real_sqrt_less_mono";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2613
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2614
Goalw [sqrt_def] "[| 0 <= x; x <= y |] ==> sqrt(x) <= sqrt(y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2615
by (auto_tac (claset() addIs [real_root_le_mono],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2616
qed "real_sqrt_le_mono";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2617
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2618
Goalw [sqrt_def] "[| 0 <= x; 0 <= y |] ==> (sqrt(x) < sqrt(y)) = (x < y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2619
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2620
qed "real_sqrt_less_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2621
Addsimps [real_sqrt_less_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2622
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2623
Goalw [sqrt_def] "[| 0 <= x; 0 <= y |] ==> (sqrt(x) <= sqrt(y)) = (x <= y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2624
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2625
qed "real_sqrt_le_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2626
Addsimps [real_sqrt_le_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2627
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2628
Goalw [sqrt_def] "[| 0 <= x; 0 <= y |] ==> (sqrt(x) = sqrt(y)) = (x = y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2629
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2630
qed "real_sqrt_eq_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2631
Addsimps [real_sqrt_eq_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2632
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2633
Goal "(sqrt(x ^ 2 + y ^ 2) < 1) = (x ^ 2 + y ^ 2 < 1)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2634
by (rtac (real_sqrt_one RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2635
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2636
by (rtac real_sqrt_less_mono 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2637
by (dtac (rotate_prems 2 (real_sqrt_less_iff RS iffD1)) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2638
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2639
qed "real_sqrt_sos_less_one_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2640
Addsimps [real_sqrt_sos_less_one_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2641
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2642
Goal "(sqrt(x ^ 2 + y ^ 2) = 1) = (x ^ 2 + y ^ 2 = 1)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2643
by (rtac (real_sqrt_one RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2644
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2645
by (dtac (rotate_prems 2 (real_sqrt_eq_iff RS iffD1)) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2646
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2647
qed "real_sqrt_sos_eq_one_iff";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2648
Addsimps [real_sqrt_sos_eq_one_iff];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2649
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2650
Goalw [real_divide_def] "(((r::real) * a) / (r * r)) = a / r";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2651
by (case_tac "r=0" 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2652
by (auto_tac (claset(),simpset() addsimps [inverse_mult_distrib] @ mult_ac));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2653
qed "real_divide_square_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2654
Addsimps [real_divide_square_eq];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2655
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2656
(*-------------------------------------------------------------------------------*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2657
(* More theorems about sqrt, transcendental functions etc. needed in Complex.ML  *)
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2658
(*-------------------------------------------------------------------------------*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
  2659
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2660
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2661
Goalw [real_divide_def]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2662
    "0 < x ==> 0 <= x/(sqrt (x * x + y * y))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2663
by (ftac ((real_sqrt_sum_squares_ge1 RSN (2,order_less_le_trans)) 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2664
    RS (CLAIM "0 < x ==> 0 < inverse (x::real)")) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2665
by (rtac (real_mult_order RS order_less_imp_le) 1);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2666
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2667
qed "lemma_real_divide_sqrt";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2668
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2669
Goal "0 < x ==> -(1::real) <= x/(sqrt (x * x + y * y))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2670
by (rtac real_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2671
by (rtac lemma_real_divide_sqrt 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2672
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2673
qed "lemma_real_divide_sqrt_ge_minus_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2674
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2675
Goal "x < 0 ==> 0 < sqrt (x * x + y * y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2676
by (rtac real_sqrt_gt_zero 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2677
by (rtac (ARITH_PROVE "[| 0 < x; 0 <= y |] ==> (0::real) < x + y") 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2678
by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2679
qed "real_sqrt_sum_squares_gt_zero1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2680
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2681
Goal "0 < x ==> 0 < sqrt (x * x + y * y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2682
by (rtac real_sqrt_gt_zero 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2683
by (rtac (ARITH_PROVE "[| 0 < x; 0 <= y |] ==> (0::real) < x + y") 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2684
by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2685
qed "real_sqrt_sum_squares_gt_zero2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2686
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2687
Goal "x ~= 0 ==> 0 < sqrt(x ^ 2 + y ^ 2)";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2688
by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2689
by (auto_tac (claset() addIs [real_sqrt_sum_squares_gt_zero2,
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2690
    real_sqrt_sum_squares_gt_zero1],simpset() addsimps [numeral_2_eq_2]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2691
qed "real_sqrt_sum_squares_gt_zero3";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2692
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2693
Goal "y ~= 0 ==> 0 < sqrt(x ^ 2 + y ^ 2)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2694
by (dres_inst_tac [("y","x")] real_sqrt_sum_squares_gt_zero3 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2695
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2696
qed "real_sqrt_sum_squares_gt_zero3a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2697
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2698
Goal "sqrt(x ^ 2 + y ^ 2) = x ==> y = 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2699
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2700
by (forw_inst_tac [("x","x")] real_sum_squares_not_zero2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2701
by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2702
by (forw_inst_tac [("x","x"),("y","y")] real_sum_square_gt_zero2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2703
by (dtac real_sqrt_gt_zero_pow2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2704
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2705
qed "real_sqrt_sum_squares_eq_cancel";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2706
Addsimps [real_sqrt_sum_squares_eq_cancel];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2707
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2708
Goal "sqrt(x ^ 2 + y ^ 2) = y ==> x = 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2709
by (res_inst_tac [("x","y")] real_sqrt_sum_squares_eq_cancel 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2710
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2711
qed "real_sqrt_sum_squares_eq_cancel2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2712
Addsimps [real_sqrt_sum_squares_eq_cancel2];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2713
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2714
Goal "x < 0 ==> x/(sqrt (x * x + y * y)) <= 1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2715
by (dtac (ARITH_PROVE "x < 0 ==> (0::real) < -x") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2716
by (dres_inst_tac [("y","y")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2717
    lemma_real_divide_sqrt_ge_minus_one 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2718
by (dtac (ARITH_PROVE "x <= y ==> -y <= -(x::real)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2719
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2720
qed "lemma_real_divide_sqrt_le_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2721
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2722
Goal "x < 0 ==> -(1::real) <= x/(sqrt (x * x + y * y))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2723
by (case_tac "y = 0" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2724
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2725
by (ftac abs_minus_eqI2 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2726
by (auto_tac (claset(),simpset() addsimps [inverse_minus_eq]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2727
by (rtac order_less_imp_le 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2728
by (res_inst_tac [("z1","sqrt(x * x + y * y)")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2729
     (real_mult_less_iff1 RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2730
by (forw_inst_tac [("y2","y")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2731
    (real_sqrt_sum_squares_gt_zero1 RS real_not_refl2 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2732
    RS not_sym) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2733
by (auto_tac (claset() addIs [real_sqrt_sum_squares_gt_zero1],
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2734
    simpset() addsimps mult_ac));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2735
by (rtac (ARITH_PROVE "-x < y ==> -y < (x::real)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2736
by (cut_inst_tac [("x","-x"),("y","y")] real_sqrt_sum_squares_ge1 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2737
by (dtac real_le_imp_less_or_eq 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2738
by (Step_tac 1);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2739
by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2]) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2740
by (dtac (sym RS real_sqrt_sum_squares_eq_cancel) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2741
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2742
qed "lemma_real_divide_sqrt_ge_minus_one2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2743
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2744
Goal "0 < x ==> x/(sqrt (x * x + y * y)) <= 1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2745
by (dtac (ARITH_PROVE "0 < x ==> -x < (0::real)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2746
by (dres_inst_tac [("y","y")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2747
    lemma_real_divide_sqrt_ge_minus_one2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2748
by (dtac (ARITH_PROVE "x <= y ==> -y <= -(x::real)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2749
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2750
qed "lemma_real_divide_sqrt_le_one2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2751
(* was qed "lemma_real_mult_self_rinv_sqrt_squared5" *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2752
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2753
Goal "-(1::real)<= x / sqrt (x * x + y * y)";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2754
by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2755
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2756
by (rtac lemma_real_divide_sqrt_ge_minus_one2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2757
by (rtac lemma_real_divide_sqrt_ge_minus_one 3);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2758
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2759
qed "cos_x_y_ge_minus_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2760
Addsimps [cos_x_y_ge_minus_one];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2761
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2762
Goal "-(1::real)<= y / sqrt (x * x + y * y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2763
by (cut_inst_tac [("x","y"),("y","x")] cos_x_y_ge_minus_one 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2764
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2765
qed "cos_x_y_ge_minus_one1a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2766
Addsimps [cos_x_y_ge_minus_one1a,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2767
          simplify (simpset()) cos_x_y_ge_minus_one1a];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2768
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2769
Goal "x / sqrt (x * x + y * y) <= 1";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2770
by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2771
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2772
by (rtac lemma_real_divide_sqrt_le_one 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2773
by (rtac lemma_real_divide_sqrt_le_one2 3);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2774
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2775
qed "cos_x_y_le_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2776
Addsimps [cos_x_y_le_one];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2777
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2778
Goal "y / sqrt (x * x + y * y) <= 1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2779
by (cut_inst_tac [("x","y"),("y","x")] cos_x_y_le_one 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2780
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2781
qed "cos_x_y_le_one2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2782
Addsimps [cos_x_y_le_one2];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2783
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2784
Addsimps [[cos_x_y_ge_minus_one,cos_x_y_le_one] MRS cos_arcos];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2785
Addsimps [[cos_x_y_ge_minus_one,cos_x_y_le_one] MRS arcos_bounded];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2786
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2787
Addsimps [[cos_x_y_ge_minus_one1a,cos_x_y_le_one2] MRS cos_arcos];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2788
Addsimps [[cos_x_y_ge_minus_one1a,cos_x_y_le_one2] MRS arcos_bounded];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2789
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2790
Goal "-(1::real) <= abs(x) / sqrt (x * x + y * y)";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2791
by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2792
by (auto_tac (claset(),simpset() addsimps [abs_minus_eqI2,abs_eqI2]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2793
by (dtac lemma_real_divide_sqrt_ge_minus_one 1 THEN Force_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2794
qed "cos_rabs_x_y_ge_minus_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2795
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2796
Addsimps [cos_rabs_x_y_ge_minus_one,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2797
          simplify (simpset()) cos_rabs_x_y_ge_minus_one];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2798
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2799
Goal "abs(x) / sqrt (x * x + y * y) <= 1";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2800
by (cut_inst_tac [("x","x"),("y","0")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2801
by (auto_tac (claset(),simpset() addsimps [abs_minus_eqI2,abs_eqI2]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2802
by (dtac lemma_real_divide_sqrt_ge_minus_one2 1 THEN Force_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2803
qed "cos_rabs_x_y_le_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2804
Addsimps [cos_rabs_x_y_le_one];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2805
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2806
Addsimps [[cos_rabs_x_y_ge_minus_one,cos_rabs_x_y_le_one] MRS cos_arcos];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2807
Addsimps [[cos_rabs_x_y_ge_minus_one,cos_rabs_x_y_le_one] MRS arcos_bounded];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2808
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2809
Goal "-pi < 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2810
by (Simp_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2811
qed "minus_pi_less_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2812
Addsimps [minus_pi_less_zero];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2813
Addsimps [minus_pi_less_zero RS order_less_imp_le];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2814
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2815
Goal "[| -(1::real) <= y; y <= 1 |] ==> -pi <= arcos y";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2816
by (rtac real_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2817
by (rtac arcos_lbound 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2818
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2819
qed "arcos_ge_minus_pi";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2820
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2821
Addsimps [[cos_x_y_ge_minus_one,cos_x_y_le_one] MRS arcos_ge_minus_pi];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2822
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2823
(* How tedious! *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2824
Goal "[| x + (y::real) ~= 0; 1 - z = x/(x + y) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2825
\     |] ==> z = y/(x + y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2826
by (res_inst_tac [("c1","x + y")] (real_mult_right_cancel RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2827
by (forw_inst_tac [("c1","x + y")] (real_mult_right_cancel RS iffD2) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2828
by (assume_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2829
by (rotate_tac 2 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2830
by (dtac (real_mult_assoc RS subst) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2831
by (rotate_tac 2 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2832
by (ftac (real_mult_inv_left RS subst) 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2833
by (assume_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2834
by (thin_tac "(1 - z) * (x + y) = x /(x + y) * (x + y)" 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2835
by (thin_tac "1 - z = x /(x + y)" 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2836
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]));
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2837
by (auto_tac (claset(),simpset() addsimps [right_distrib,
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2838
    left_diff_distrib]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2839
qed "lemma_divide_rearrange";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2840
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2841
Goal "[| 0 < x * x + y * y; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2842
\       1 - sin xa ^ 2 = (x / sqrt (x * x + y * y)) ^ 2 \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2843
\     |] ==> sin xa ^ 2 = (y / sqrt (x * x + y * y)) ^ 2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2844
by (auto_tac (claset() addIs [lemma_divide_rearrange],simpset() 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2845
    addsimps [realpow_divide,real_sqrt_gt_zero_pow2,
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2846
    real_power_two RS sym]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2847
qed "lemma_cos_sin_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2848
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2849
Goal "[| 0 < x * x + y * y; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2850
\       1 - cos xa ^ 2 = (y / sqrt (x * x + y * y)) ^ 2 \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2851
\     |] ==> cos xa ^ 2 = (x / sqrt (x * x + y * y)) ^ 2";
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  2852
by (auto_tac (claset(),
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  2853
    simpset() addsimps [realpow_divide,
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2854
          real_sqrt_gt_zero_pow2,real_power_two RS sym]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2855
by (rtac (real_add_commute RS subst) 1);
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  2856
by (rtac lemma_divide_rearrange 1); 
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  2857
by (asm_full_simp_tac (simpset() addsimps []) 1);
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  2858
by (asm_full_simp_tac (simpset() addsimps [add_commute]) 1);  
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2859
qed "lemma_sin_cos_eq";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2860
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2861
Goal "[| x ~= 0; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2862
\        cos xa = x / sqrt (x * x + y * y) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2863
\     |] ==>  sin xa = y / sqrt (x * x + y * y) | \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2864
\             sin xa = - y / sqrt (x * x + y * y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2865
by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2866
by (forw_inst_tac [("y","y")] real_sum_square_gt_zero 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2867
by (asm_full_simp_tac (simpset() addsimps [cos_squared_eq]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2868
by (subgoal_tac "sin xa ^ 2 =  (y / sqrt (x * x + y * y)) ^ 2" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2869
by (rtac lemma_cos_sin_eq 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2870
by (Force_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2871
by (Asm_full_simp_tac 2);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2872
by (auto_tac (claset(),simpset() addsimps [realpow_two_disj,numeral_2_eq_2] 
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2873
    delsimps [realpow_Suc]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2874
qed "sin_x_y_disj";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2875
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2876
Goal "y ~= 0 ==> x / sqrt (x * x + y * y) ~= -(1::real)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2877
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2878
by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2879
by (auto_tac (claset(),simpset() addsimps [realpow_divide,
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2880
    real_power_two RS sym]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2881
by (forw_inst_tac [("x","x")] real_sum_squares_not_zero2 1);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2882
by (asm_full_simp_tac (HOL_ss addsimps [real_power_two RS sym]) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2883
by (forw_inst_tac [("c1","x ^ 2 + y ^ 2")]
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2884
    (real_mult_right_cancel RS iffD2) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2885
by (assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2886
by (dres_inst_tac [("y","x ^ 2 + y ^ 2"),("x","x ^ 2")] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2887
    (CLAIM "y ~= (0::real) ==> (x/y) * y= x") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2888
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2889
qed "cos_not_eq_minus_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2890
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2891
Goalw [arcos_def] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2892
  "arcos (x / sqrt (x * x + y * y)) = pi ==> y = 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2893
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2894
by (rtac swap 1 THEN assume_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2895
by (rtac (([cos_x_y_ge_minus_one,cos_x_y_le_one] MRS cos_total) RS 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2896
     ((CLAIM "EX! x. P x ==> EX x. P x") RS someI2_ex)) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2897
by (auto_tac (claset() addDs [cos_not_eq_minus_one],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2898
qed "arcos_eq_pi_cancel";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2899
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2900
Goalw [real_divide_def] "x ~= 0 ==> x / sqrt (x * x + y * y) ~= 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2901
by (forw_inst_tac [("y3","y")] (real_sqrt_sum_squares_gt_zero3 
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2902
    RS real_not_refl2 RS not_sym RS nonzero_imp_inverse_nonzero) 1);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2903
by (auto_tac (claset(),simpset() addsimps [real_power_two]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2904
qed "lemma_cos_not_eq_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2905
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2906
Goal "[| x ~= 0; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2907
\        sin xa = y / sqrt (x * x + y * y) \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2908
\     |] ==>  cos xa = x / sqrt (x * x + y * y) | \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2909
\             cos xa = - x / sqrt (x * x + y * y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2910
by (dres_inst_tac [("f","%x. x ^ 2")] arg_cong 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2911
by (forw_inst_tac [("y","y")] real_sum_square_gt_zero 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2912
by (asm_full_simp_tac (simpset() addsimps [sin_squared_eq] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2913
    delsimps [realpow_Suc]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2914
by (subgoal_tac "cos xa ^ 2 =  (x / sqrt (x * x + y * y)) ^ 2" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2915
by (rtac lemma_sin_cos_eq 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2916
by (Force_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2917
by (Asm_full_simp_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2918
by (auto_tac (claset(),simpset() addsimps [realpow_two_disj,
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2919
    numeral_2_eq_2] delsimps [realpow_Suc]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2920
qed "cos_x_y_disj";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2921
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2922
Goal "0 < y ==> - y / sqrt (x * x + y * y) < 0";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2923
by (case_tac "x = 0" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2924
by (auto_tac (claset(),simpset() addsimps [abs_eqI2]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2925
by (dres_inst_tac [("y","y")] real_sqrt_sum_squares_gt_zero3 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  2926
by (auto_tac (claset(),simpset() addsimps [zero_less_mult_iff,
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2927
    real_divide_def,real_power_two])); 
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2928
qed "real_sqrt_divide_less_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2929
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2930
Goal "[| x ~= 0; 0 < y |] ==> EX r a. x = r * cos a & y = r * sin a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2931
by (res_inst_tac [("x","sqrt(x ^ 2 + y ^ 2)")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2932
by (res_inst_tac [("x","arcos(x / sqrt (x * x + y * y))")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2933
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2934
by (dres_inst_tac [("y2","y")] (real_sqrt_sum_squares_gt_zero3 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2935
    RS real_not_refl2 RS not_sym) 1);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2936
by (auto_tac (claset(),simpset() addsimps [real_power_two]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2937
by (rewtac arcos_def);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2938
by (cut_inst_tac [("x1","x"),("y1","y")] ([cos_x_y_ge_minus_one,cos_x_y_le_one] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2939
    MRS cos_total) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2940
by (rtac someI2_ex 1 THEN Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2941
by (thin_tac 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2942
    "EX! xa. 0 <= xa & xa <= pi & cos xa = x / sqrt (x * x + y * y)" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2943
by (ftac sin_x_y_disj 1 THEN Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2944
by (dres_inst_tac [("y2","y")] (real_sqrt_sum_squares_gt_zero3 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2945
    RS real_not_refl2 RS not_sym) 1);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2946
by (auto_tac (claset(),simpset() addsimps [real_power_two]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2947
by (dtac sin_ge_zero 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2948
by (dres_inst_tac [("x","x")] real_sqrt_divide_less_zero 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2949
qed "polar_ex1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2950
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2951
Goal "x * x = -(y * y) ==> y = (0::real)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2952
by (auto_tac (claset() addIs [real_sum_squares_cancel],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2953
qed "real_sum_squares_cancel2a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2954
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2955
Goal "[| x ~= 0; y < 0 |] ==> EX r a. x = r * cos a & y = r * sin a";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2956
by (cut_inst_tac [("x","0"),("y","x")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2957
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2958
by (res_inst_tac [("x","sqrt(x ^ 2 + y ^ 2)")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2959
by (res_inst_tac [("x","arcsin(y / sqrt (x * x + y * y))")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2960
by (auto_tac (claset() addDs [real_sum_squares_cancel2a],
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2961
    simpset() addsimps [real_power_two]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2962
by (rewtac arcsin_def);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2963
by (cut_inst_tac [("x1","x"),("y1","y")] ([cos_x_y_ge_minus_one1a,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2964
    cos_x_y_le_one2] MRS sin_total) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2965
by (rtac someI2_ex 1 THEN Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2966
by (thin_tac "EX! xa. - (pi/2) <= xa & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2967
\                xa <= pi/2 & sin xa = y / sqrt (x * x + y * y)" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2968
by (ftac ((CLAIM "0 < x ==> (x::real) ~= 0") RS cos_x_y_disj) 1 THEN Blast_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2969
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2970
by (dtac cos_ge_zero 1 THEN Force_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2971
by (dres_inst_tac [("x","y")] real_sqrt_divide_less_zero 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2972
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2973
by (dtac (ARITH_PROVE "(y::real) < 0 ==> 0 < - y") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2974
by (dtac (CLAIM "x < (0::real) ==> x ~= 0" RS polar_ex1) 1 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2975
by (REPEAT(etac exE 1));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2976
by (res_inst_tac [("x","r")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2977
by (res_inst_tac [("x","-a")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2978
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2979
qed "polar_ex2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2980
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2981
Goal "EX r a. x = r * cos a & y = r * sin a";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2982
by (case_tac "x = 0" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2983
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2984
by (res_inst_tac [("x","y")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2985
by (res_inst_tac [("x","pi/2")] exI 1 THEN Auto_tac);
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  2986
by (cut_inst_tac [("x","0"),("y","y")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2987
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2988
by (res_inst_tac [("x","x")] exI 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2989
by (res_inst_tac [("x","0")] exI 2 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2990
by (ALLGOALS(blast_tac (claset() addIs [polar_ex1,polar_ex2])));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2991
qed "polar_Ex";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2992
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2993
Goal "abs x <= sqrt (x ^ 2 + y ^ 2)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2994
by (res_inst_tac [("n","1")] realpow_increasing 1);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2995
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym]));
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  2996
by (simp_tac (simpset() addsimps [numeral_2_eq_2]) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2997
qed "real_sqrt_ge_abs1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2998
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  2999
Goal "abs y <= sqrt (x ^ 2 + y ^ 2)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3000
by (rtac (real_add_commute RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3001
by (rtac real_sqrt_ge_abs1 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3002
qed "real_sqrt_ge_abs2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3003
Addsimps [real_sqrt_ge_abs1,real_sqrt_ge_abs2];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3004
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3005
Goal "0 < sqrt 2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3006
by (auto_tac (claset() addIs [real_sqrt_gt_zero],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3007
qed "real_sqrt_two_gt_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3008
Addsimps [real_sqrt_two_gt_zero];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3009
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3010
Goal "0 <= sqrt 2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3011
by (auto_tac (claset() addIs [real_sqrt_ge_zero],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3012
qed "real_sqrt_two_ge_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3013
Addsimps [real_sqrt_two_ge_zero];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3014
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3015
Goal "1 < sqrt 2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3016
by (res_inst_tac [("y","7/5")] order_less_le_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3017
by (res_inst_tac [("n","1")] realpow_increasing 2);
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  3018
by (auto_tac (claset(),simpset() addsimps [real_sqrt_gt_zero_pow2,numeral_2_eq_2 RS sym] 
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3019
    delsimps [realpow_Suc]));
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  3020
by (simp_tac (simpset() addsimps [numeral_2_eq_2]) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3021
qed "real_sqrt_two_gt_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3022
Addsimps [real_sqrt_two_gt_one];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3023
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3024
Goal "0 < u ==> u / sqrt 2 < u";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3025
by (res_inst_tac [("z1","inverse u")] (real_mult_less_iff1 RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3026
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3027
by (res_inst_tac [("z1","sqrt 2")] (real_mult_less_iff1 RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3028
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3029
qed "lemma_real_divide_sqrt_less";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3030
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3031
(* needed for infinitely close relation over the nonstandard complex numbers *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3032
Goal "[| 0 < u; x < u/2; y < u/2; 0 <= x; 0 <= y |] ==> sqrt (x ^ 2 + y ^ 2) < u";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3033
by (res_inst_tac [("y","u/sqrt 2")] order_le_less_trans 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3034
by (etac lemma_real_divide_sqrt_less 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3035
by (res_inst_tac [("n","1")] realpow_increasing 1);
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14270
diff changeset
  3036
by (auto_tac (claset(),
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14270
diff changeset
  3037
    simpset() addsimps [real_0_le_divide_iff,realpow_divide,
14322
fa78e7eb1dac deleting redundant theorems
paulson
parents: 14309
diff changeset
  3038
    real_sqrt_gt_zero_pow2,numeral_2_eq_2 RS sym] delsimps [realpow_Suc]));
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3039
by (res_inst_tac [("t","u ^ 2")] (real_sum_of_halves RS subst) 1);
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14331
diff changeset
  3040
by (rtac add_mono 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3041
by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3042
by (ALLGOALS(rtac ((CLAIM "(2::real) ^ 2 = 4") RS subst)));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3043
by (ALLGOALS(rtac (realpow_mult RS subst)));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3044
by (ALLGOALS(rtac realpow_le));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3045
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3046
qed "lemma_sqrt_hcomplex_capprox";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3047
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3048
Addsimps [real_sqrt_sum_squares_ge_zero RS abs_eqI1];
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3049
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3050
(* A few theorems involving ln and derivatives, etc *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3051
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3052
Goal "DERIV ln z :> l ==> DERIV (%x. exp (ln x)) z :> exp (ln z) * l";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3053
by (etac DERIV_fun_exp 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3054
qed "lemma_DERIV_ln";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3055
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3056
Goal "0 < z ==> ( *f* (%x. exp (ln x))) z = z";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3057
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3058
by (auto_tac (claset(),simpset() addsimps [starfun,
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3059
    hypreal_zero_def,hypreal_less]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3060
qed "STAR_exp_ln";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3061
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3062
Goal "[|e : Infinitesimal; 0 < x |] ==> 0 < hypreal_of_real x + e";
14331
8dbbb7cf3637 re-organized numeric lemmas
paulson
parents: 14322
diff changeset
  3063
by (res_inst_tac [("c1","-e")] (add_less_cancel_right RS iffD1) 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3064
by (auto_tac (claset() addIs [Infinitesimal_less_SReal],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3065
qed "hypreal_add_Infinitesimal_gt_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3066
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3067
Goalw [nsderiv_def,NSLIM_def] "0 < z ==> NSDERIV (%x. exp (ln x)) z :> 1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3068
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3069
by (rtac ccontr 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3070
by (subgoal_tac "0 < hypreal_of_real z + h" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3071
by (dtac STAR_exp_ln 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3072
by (rtac hypreal_add_Infinitesimal_gt_zero 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3073
by (dtac (CLAIM "h ~= 0 ==> h/h = (1::hypreal)") 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3074
by (auto_tac (claset(),simpset() addsimps [exp_ln_iff RS sym] 
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3075
    delsimps [exp_ln_iff]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3076
qed "NSDERIV_exp_ln_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3077
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3078
Goal "0 < z ==> DERIV (%x. exp (ln x)) z :> 1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3079
by (auto_tac (claset() addIs [NSDERIV_exp_ln_one],
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3080
    simpset() addsimps [NSDERIV_DERIV_iff RS sym]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3081
qed "DERIV_exp_ln_one";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3082
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3083
Goal "[| 0 < z; DERIV ln z :> l |] ==>  exp (ln z) * l = 1";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3084
by (rtac DERIV_unique 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3085
by (rtac lemma_DERIV_ln 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3086
by (rtac DERIV_exp_ln_one 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3087
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3088
qed "lemma_DERIV_ln2";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3089
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3090
Goal "[| 0 < z; DERIV ln z :> l |] ==>  l = 1/(exp (ln z))";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3091
by (res_inst_tac [("c1","exp(ln z)")] (real_mult_left_cancel RS iffD1) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3092
by (auto_tac (claset() addIs [lemma_DERIV_ln2],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3093
qed "lemma_DERIV_ln3";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3094
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3095
Goal "[| 0 < z; DERIV ln z :> l |] ==>  l = 1/z";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3096
by (res_inst_tac [("t","z")] (exp_ln_iff RS iffD2 RS subst) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3097
by (auto_tac (claset() addIs [lemma_DERIV_ln3],simpset()));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3098
qed "lemma_DERIV_ln4";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3099
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3100
(* need to rename second isCont_inverse *)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3101
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3102
Goal "[| 0 < d; ALL z. abs(z - x) <= d --> g(f(z)) = z; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3103
\        ALL z. abs(z - x) <= d --> isCont f z |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3104
\     ==> isCont g (f x)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3105
by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3106
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3107
by (dres_inst_tac [("d1.0","r")] real_lbound_gt_zero 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3108
by (assume_tac 1 THEN Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3109
by (subgoal_tac "ALL z. abs(z - x) <= e --> (g(f z) = z)" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3110
by (Force_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3111
by (subgoal_tac "ALL z. abs(z - x) <= e --> isCont f z" 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3112
by (Force_tac 2);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3113
by (dres_inst_tac [("d","e")] isCont_inj_range 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3114
by (assume_tac 2 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3115
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3116
by (res_inst_tac [("x","ea")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3117
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3118
by (rotate_tac 4 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3119
by (dres_inst_tac [("x","f(x) + xa")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3120
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3121
by (dtac sym 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3122
by (arith_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3123
qed "isCont_inv_fun";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3124
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3125
(*
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3126
Goalw [isCont_def]  
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3127
      "[| isCont f x; f x ~= 0 |] ==> isCont (%x. inverse (f x)) x";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3128
by (blast_tac (claset() addIs [LIM_inverse]) 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3129
qed "isCont_inverse";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3130
*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3131
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3132
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3133
Goal "[| 0 < d; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3134
\        ALL z. abs(z - x) <= d --> g(f(z)) = z; \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3135
\        ALL z. abs(z - x) <= d --> isCont f z |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3136
\      ==> EX e. 0 < (e::real) & \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3137
\            (ALL y. 0 < abs(y - f(x::real)) & abs(y - f(x)) < e --> f(g(y)) = y)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3138
by (dtac isCont_inj_range 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3139
by (assume_tac 2 THEN assume_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3140
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3141
by (res_inst_tac [("x","e")] exI 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3142
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3143
by (dres_inst_tac [("x","y")] spec 1 THEN Auto_tac);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3144
qed "isCont_inv_fun_inv";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3145
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3146
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3147
(* Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3148
Goal "[| f -- c --> l; 0 < l |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3149
\        ==> EX r. 0 < r & (ALL x. x ~= c & abs (c - x) < r --> 0 < f x)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3150
by (auto_tac (claset(),simpset() addsimps [LIM_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3151
by (dres_inst_tac [("x","l/2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3152
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3153
by (Force_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3154
by (res_inst_tac [("x","s")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3155
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3156
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3157
by (dres_inst_tac [("x","x")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3158
by (auto_tac (claset(),HOL_ss addsimps [abs_interval_iff]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3159
by (auto_tac (claset(),simpset() addsimps [CLAIM "(l::real) + -(l/2) = l/2",
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3160
    CLAIM "(a < f + - l) = (l + a < (f::real))"]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3161
qed "LIM_fun_gt_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3162
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3163
Goal "[| f -- c --> l; l < 0 |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3164
\        ==> EX r. 0 < r & (ALL x. x ~= c & abs (c - x) < r --> f x < 0)";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3165
by (auto_tac (claset(),simpset() addsimps [LIM_def]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3166
by (dres_inst_tac [("x","-l/2")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3167
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3168
by (Force_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3169
by (res_inst_tac [("x","s")] exI 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3170
by (Step_tac 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3171
by (rotate_tac 2 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3172
by (dres_inst_tac [("x","x")] spec 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3173
by (auto_tac (claset(),HOL_ss addsimps [abs_interval_iff]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3174
by (auto_tac (claset(),simpset() addsimps [CLAIM "(l::real) + -(l/2) = l/2",
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3175
    CLAIM "(f + - l < a) = ((f::real) < l + a)"]));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3176
qed "LIM_fun_less_zero";
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3177
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3178
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3179
Goal "[| f -- c --> l; l ~= 0 |] \
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3180
\        ==> EX r. 0 < r & (ALL x. x ~= c & abs (c - x) < r --> f x ~= 0)";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
  3181
by (cut_inst_tac [("x","l"),("y","0")] linorder_less_linear 1);
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3182
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3183
by (dtac LIM_fun_less_zero 1);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3184
by (dtac LIM_fun_gt_zero 3);
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3185
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3186
by (ALLGOALS(res_inst_tac [("x","r")] exI));
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3187
by Auto_tac;
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 13601
diff changeset
  3188
qed "LIM_fun_not_zero";