src/HOL/Transcendental.thy
author haftmann
Tue, 28 Apr 2009 15:50:30 +0200
changeset 31017 2c227493ea56
parent 30273 ecd6f0ca62ea
child 31148 7ba7c1f8bc22
permissions -rw-r--r--
stripped class recpower further
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     1
(*  Title       : Transcendental.thy
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
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents: 12196
diff changeset
     4
                  1999,2001 University of Edinburgh
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
     5
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     6
*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     7
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
     8
header{*Power Series, Transcendental Functions etc.*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
     9
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15086
diff changeset
    10
theory Transcendental
25600
73431bd8c4c4 joined EvenOdd theory with Parity
haftmann
parents: 25153
diff changeset
    11
imports Fact Series Deriv NthRoot
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15086
diff changeset
    12
begin
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    13
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
    14
subsection {* Properties of Power Series *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    15
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    16
lemma lemma_realpow_diff:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
    17
  fixes y :: "'a::monoid_mult"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    18
  shows "p \<le> n \<Longrightarrow> y ^ (Suc n - p) = (y ^ (n - p)) * y"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    19
proof -
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    20
  assume "p \<le> n"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    21
  hence "Suc n - p = Suc (n - p)" by (rule Suc_diff_le)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
    22
  thus ?thesis by (simp add: power_commutes)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    23
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    24
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    25
lemma lemma_realpow_diff_sumr:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
    26
  fixes y :: "'a::{comm_semiring_0,monoid_mult}" shows
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    27
     "(\<Sum>p=0..<Suc n. (x ^ p) * y ^ (Suc n - p)) =  
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    28
      y * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))"
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
    29
by (simp add: setsum_right_distrib lemma_realpow_diff mult_ac
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
    30
         del: setsum_op_ivl_Suc cong: strong_setsum_cong)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    31
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
    32
lemma lemma_realpow_diff_sumr2:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
    33
  fixes y :: "'a::{comm_ring,monoid_mult}" shows
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
    34
     "x ^ (Suc n) - y ^ (Suc n) =  
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    35
      (x - y) * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
    36
apply (induct n, simp)
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
    37
apply (simp del: setsum_op_ivl_Suc)
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15546
diff changeset
    38
apply (subst setsum_op_ivl_Suc)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    39
apply (subst lemma_realpow_diff_sumr)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    40
apply (simp add: right_distrib del: setsum_op_ivl_Suc)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    41
apply (subst mult_left_commute [where a="x - y"])
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    42
apply (erule subst)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
    43
apply (simp add: algebra_simps)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    44
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    45
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
    46
lemma lemma_realpow_rev_sumr:
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
    47
     "(\<Sum>p=0..<Suc n. (x ^ p) * (y ^ (n - p))) =  
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    48
      (\<Sum>p=0..<Suc n. (x ^ (n - p)) * (y ^ p))"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    49
apply (rule setsum_reindex_cong [where f="\<lambda>i. n - i"])
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    50
apply (rule inj_onI, simp)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    51
apply auto
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    52
apply (rule_tac x="n - x" in image_eqI, simp, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    53
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    54
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    55
text{*Power series has a `circle` of convergence, i.e. if it sums for @{term
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    56
x}, then it sums absolutely for @{term z} with @{term "\<bar>z\<bar> < \<bar>x\<bar>"}.*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    57
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
    58
lemma powser_insidea:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
    59
  fixes x z :: "'a::{real_normed_field,banach}"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    60
  assumes 1: "summable (\<lambda>n. f n * x ^ n)"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    61
  assumes 2: "norm z < norm x"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    62
  shows "summable (\<lambda>n. norm (f n * z ^ n))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    63
proof -
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    64
  from 2 have x_neq_0: "x \<noteq> 0" by clarsimp
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    65
  from 1 have "(\<lambda>n. f n * x ^ n) ----> 0"
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    66
    by (rule summable_LIMSEQ_zero)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    67
  hence "convergent (\<lambda>n. f n * x ^ n)"
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    68
    by (rule convergentI)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    69
  hence "Cauchy (\<lambda>n. f n * x ^ n)"
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    70
    by (simp add: Cauchy_convergent_iff)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    71
  hence "Bseq (\<lambda>n. f n * x ^ n)"
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    72
    by (rule Cauchy_Bseq)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    73
  then obtain K where 3: "0 < K" and 4: "\<forall>n. norm (f n * x ^ n) \<le> K"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    74
    by (simp add: Bseq_def, safe)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    75
  have "\<exists>N. \<forall>n\<ge>N. norm (norm (f n * z ^ n)) \<le>
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    76
                   K * norm (z ^ n) * inverse (norm (x ^ n))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    77
  proof (intro exI allI impI)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    78
    fix n::nat assume "0 \<le> n"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    79
    have "norm (norm (f n * z ^ n)) * norm (x ^ n) =
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    80
          norm (f n * x ^ n) * norm (z ^ n)"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    81
      by (simp add: norm_mult abs_mult)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    82
    also have "\<dots> \<le> K * norm (z ^ n)"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    83
      by (simp only: mult_right_mono 4 norm_ge_zero)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    84
    also have "\<dots> = K * norm (z ^ n) * (inverse (norm (x ^ n)) * norm (x ^ n))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    85
      by (simp add: x_neq_0)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    86
    also have "\<dots> = K * norm (z ^ n) * inverse (norm (x ^ n)) * norm (x ^ n)"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    87
      by (simp only: mult_assoc)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    88
    finally show "norm (norm (f n * z ^ n)) \<le>
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    89
                  K * norm (z ^ n) * inverse (norm (x ^ n))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    90
      by (simp add: mult_le_cancel_right x_neq_0)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    91
  qed
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    92
  moreover have "summable (\<lambda>n. K * norm (z ^ n) * inverse (norm (x ^ n)))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    93
  proof -
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    94
    from 2 have "norm (norm (z * inverse x)) < 1"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    95
      using x_neq_0
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    96
      by (simp add: nonzero_norm_divide divide_inverse [symmetric])
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    97
    hence "summable (\<lambda>n. norm (z * inverse x) ^ n)"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
    98
      by (rule summable_geometric)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
    99
    hence "summable (\<lambda>n. K * norm (z * inverse x) ^ n)"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   100
      by (rule summable_mult)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   101
    thus "summable (\<lambda>n. K * norm (z ^ n) * inverse (norm (x ^ n)))"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   102
      using x_neq_0
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   103
      by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   104
                    power_inverse norm_power mult_assoc)
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   105
  qed
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   106
  ultimately show "summable (\<lambda>n. norm (f n * z ^ n))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   107
    by (rule summable_comparison_test)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   108
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   109
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   110
lemma powser_inside:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   111
  fixes f :: "nat \<Rightarrow> 'a::{real_normed_field,banach}" shows
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   112
     "[| summable (%n. f(n) * (x ^ n)); norm z < norm x |]  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   113
      ==> summable (%n. f(n) * (z ^ n))"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   114
by (rule powser_insidea [THEN summable_norm_cancel])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   115
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   116
lemma sum_split_even_odd: fixes f :: "nat \<Rightarrow> real" shows
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   117
  "(\<Sum> i = 0 ..< 2 * n. if even i then f i else g i) = 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   118
   (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   119
proof (induct n)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   120
  case (Suc n)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   121
  have "(\<Sum> i = 0 ..< 2 * Suc n. if even i then f i else g i) = 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   122
        (\<Sum> i = 0 ..< n. f (2 * i)) + (\<Sum> i = 0 ..< n. g (2 * i + 1)) + (f (2 * n) + g (2 * n + 1))"
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   123
    using Suc.hyps unfolding One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   124
  also have "\<dots> = (\<Sum> i = 0 ..< Suc n. f (2 * i)) + (\<Sum> i = 0 ..< Suc n. g (2 * i + 1))" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   125
  finally show ?case .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   126
qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   127
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   128
lemma sums_if': fixes g :: "nat \<Rightarrow> real" assumes "g sums x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   129
  shows "(\<lambda> n. if even n then 0 else g ((n - 1) div 2)) sums x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   130
  unfolding sums_def
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   131
proof (rule LIMSEQ_I)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   132
  fix r :: real assume "0 < r"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   133
  from `g sums x`[unfolded sums_def, THEN LIMSEQ_D, OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   134
  obtain no where no_eq: "\<And> n. n \<ge> no \<Longrightarrow> (norm (setsum g { 0..<n } - x) < r)" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   135
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   136
  let ?SUM = "\<lambda> m. \<Sum> i = 0 ..< m. if even i then 0 else g ((i - 1) div 2)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   137
  { fix m assume "m \<ge> 2 * no" hence "m div 2 \<ge> no" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   138
    have sum_eq: "?SUM (2 * (m div 2)) = setsum g { 0 ..< m div 2 }" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   139
      using sum_split_even_odd by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   140
    hence "(norm (?SUM (2 * (m div 2)) - x) < r)" using no_eq unfolding sum_eq using `m div 2 \<ge> no` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   141
    moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   142
    have "?SUM (2 * (m div 2)) = ?SUM m"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   143
    proof (cases "even m")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   144
      case True show ?thesis unfolding even_nat_div_two_times_two[OF True, unfolded numeral_2_eq_2[symmetric]] ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   145
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   146
      case False hence "even (Suc m)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   147
      from even_nat_div_two_times_two[OF this, unfolded numeral_2_eq_2[symmetric]] odd_nat_plus_one_div_two[OF False, unfolded numeral_2_eq_2[symmetric]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   148
      have eq: "Suc (2 * (m div 2)) = m" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   149
      hence "even (2 * (m div 2))" using `odd m` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   150
      have "?SUM m = ?SUM (Suc (2 * (m div 2)))" unfolding eq ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   151
      also have "\<dots> = ?SUM (2 * (m div 2))" using `even (2 * (m div 2))` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   152
      finally show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   153
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   154
    ultimately have "(norm (?SUM m - x) < r)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   155
  }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   156
  thus "\<exists> no. \<forall> m \<ge> no. norm (?SUM m - x) < r" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   157
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   158
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   159
lemma sums_if: fixes g :: "nat \<Rightarrow> real" assumes "g sums x" and "f sums y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   160
  shows "(\<lambda> n. if even n then f (n div 2) else g ((n - 1) div 2)) sums (x + y)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   161
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   162
  let ?s = "\<lambda> n. if even n then 0 else f ((n - 1) div 2)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   163
  { fix B T E have "(if B then (0 :: real) else E) + (if B then T else 0) = (if B then T else E)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   164
      by (cases B) auto } note if_sum = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   165
  have g_sums: "(\<lambda> n. if even n then 0 else g ((n - 1) div 2)) sums x" using sums_if'[OF `g sums x`] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   166
  { 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   167
    have "?s 0 = 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   168
    have Suc_m1: "\<And> n. Suc n - 1 = n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   169
    { fix B T E have "(if \<not> B then T else E) = (if B then E else T)" by auto } note if_eq = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   170
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   171
    have "?s sums y" using sums_if'[OF `f sums y`] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   172
    from this[unfolded sums_def, THEN LIMSEQ_Suc] 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   173
    have "(\<lambda> n. if even n then f (n div 2) else 0) sums y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   174
      unfolding sums_def setsum_shift_lb_Suc0_0_upt[where f="?s", OF `?s 0 = 0`, symmetric]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   175
                image_Suc_atLeastLessThan[symmetric] setsum_reindex[OF inj_Suc, unfolded comp_def]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   176
                even_nat_Suc Suc_m1 if_eq .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   177
  } from sums_add[OF g_sums this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   178
  show ?thesis unfolding if_sum .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   179
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   180
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   181
subsection {* Alternating series test / Leibniz formula *}
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   182
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   183
lemma sums_alternating_upper_lower:
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   184
  fixes a :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   185
  assumes mono: "\<And>n. a (Suc n) \<le> a n" and a_pos: "\<And>n. 0 \<le> a n" and "a ----> 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   186
  shows "\<exists>l. ((\<forall>n. (\<Sum>i=0..<2*n. -1^i*a i) \<le> l) \<and> (\<lambda> n. \<Sum>i=0..<2*n. -1^i*a i) ----> l) \<and> 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   187
             ((\<forall>n. l \<le> (\<Sum>i=0..<2*n + 1. -1^i*a i)) \<and> (\<lambda> n. \<Sum>i=0..<2*n + 1. -1^i*a i) ----> l)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   188
  (is "\<exists>l. ((\<forall>n. ?f n \<le> l) \<and> _) \<and> ((\<forall>n. l \<le> ?g n) \<and> _)")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   189
proof -
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   190
  have fg_diff: "\<And>n. ?f n - ?g n = - a (2 * n)" unfolding One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   191
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   192
  have "\<forall> n. ?f n \<le> ?f (Suc n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   193
  proof fix n show "?f n \<le> ?f (Suc n)" using mono[of "2*n"] by auto qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   194
  moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   195
  have "\<forall> n. ?g (Suc n) \<le> ?g n"
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   196
  proof fix n show "?g (Suc n) \<le> ?g n" using mono[of "Suc (2*n)"]
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   197
    unfolding One_nat_def by auto qed
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   198
  moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   199
  have "\<forall> n. ?f n \<le> ?g n" 
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   200
  proof fix n show "?f n \<le> ?g n" using fg_diff a_pos
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   201
    unfolding One_nat_def by auto qed
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   202
  moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   203
  have "(\<lambda> n. ?f n - ?g n) ----> 0" unfolding fg_diff
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   204
  proof (rule LIMSEQ_I)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   205
    fix r :: real assume "0 < r"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   206
    with `a ----> 0`[THEN LIMSEQ_D] 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   207
    obtain N where "\<And> n. n \<ge> N \<Longrightarrow> norm (a n - 0) < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   208
    hence "\<forall> n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   209
    thus "\<exists> N. \<forall> n \<ge> N. norm (- a (2 * n) - 0) < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   210
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   211
  ultimately
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   212
  show ?thesis by (rule lemma_nest_unique)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   213
qed 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   214
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   215
lemma summable_Leibniz': fixes a :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   216
  assumes a_zero: "a ----> 0" and a_pos: "\<And> n. 0 \<le> a n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   217
  and a_monotone: "\<And> n. a (Suc n) \<le> a n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   218
  shows summable: "summable (\<lambda> n. (-1)^n * a n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   219
  and "\<And>n. (\<Sum>i=0..<2*n. (-1)^i*a i) \<le> (\<Sum>i. (-1)^i*a i)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   220
  and "(\<lambda>n. \<Sum>i=0..<2*n. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   221
  and "\<And>n. (\<Sum>i. (-1)^i*a i) \<le> (\<Sum>i=0..<2*n+1. (-1)^i*a i)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   222
  and "(\<lambda>n. \<Sum>i=0..<2*n+1. (-1)^i*a i) ----> (\<Sum>i. (-1)^i*a i)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   223
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   224
  let "?S n" = "(-1)^n * a n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   225
  let "?P n" = "\<Sum>i=0..<n. ?S i"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   226
  let "?f n" = "?P (2 * n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   227
  let "?g n" = "?P (2 * n + 1)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   228
  obtain l :: real where below_l: "\<forall> n. ?f n \<le> l" and "?f ----> l" and above_l: "\<forall> n. l \<le> ?g n" and "?g ----> l"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   229
    using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   230
  
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   231
  let ?Sa = "\<lambda> m. \<Sum> n = 0..<m. ?S n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   232
  have "?Sa ----> l"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   233
  proof (rule LIMSEQ_I)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   234
    fix r :: real assume "0 < r"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   235
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   236
    with `?f ----> l`[THEN LIMSEQ_D] 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   237
    obtain f_no where f: "\<And> n. n \<ge> f_no \<Longrightarrow> norm (?f n - l) < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   238
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   239
    from `0 < r` `?g ----> l`[THEN LIMSEQ_D] 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   240
    obtain g_no where g: "\<And> n. n \<ge> g_no \<Longrightarrow> norm (?g n - l) < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   241
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   242
    { fix n :: nat
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   243
      assume "n \<ge> (max (2 * f_no) (2 * g_no))" hence "n \<ge> 2 * f_no" and "n \<ge> 2 * g_no" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   244
      have "norm (?Sa n - l) < r"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   245
      proof (cases "even n")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   246
	case True from even_nat_div_two_times_two[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   247
	have n_eq: "2 * (n div 2) = n" unfolding numeral_2_eq_2[symmetric] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   248
	with `n \<ge> 2 * f_no` have "n div 2 \<ge> f_no" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   249
	from f[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   250
	show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   251
      next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   252
	case False hence "even (n - 1)" using even_num_iff odd_pos by auto 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   253
	from even_nat_div_two_times_two[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   254
	have n_eq: "2 * ((n - 1) div 2) = n - 1" unfolding numeral_2_eq_2[symmetric] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   255
	hence range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   256
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   257
	from n_eq `n \<ge> 2 * g_no` have "(n - 1) div 2 \<ge> g_no" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   258
	from g[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   259
	show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost range_eq .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   260
      qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   261
    }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   262
    thus "\<exists> no. \<forall> n \<ge> no. norm (?Sa n - l) < r" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   263
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   264
  hence sums_l: "(\<lambda>i. (-1)^i * a i) sums l" unfolding sums_def atLeastLessThanSuc_atLeastAtMost[symmetric] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   265
  thus "summable ?S" using summable_def by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   266
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   267
  have "l = suminf ?S" using sums_unique[OF sums_l] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   268
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   269
  { fix n show "suminf ?S \<le> ?g n" unfolding sums_unique[OF sums_l, symmetric] using above_l by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   270
  { fix n show "?f n \<le> suminf ?S" unfolding sums_unique[OF sums_l, symmetric] using below_l by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   271
  show "?g ----> suminf ?S" using `?g ----> l` `l = suminf ?S` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   272
  show "?f ----> suminf ?S" using `?f ----> l` `l = suminf ?S` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   273
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   274
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   275
theorem summable_Leibniz: fixes a :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   276
  assumes a_zero: "a ----> 0" and "monoseq a"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   277
  shows "summable (\<lambda> n. (-1)^n * a n)" (is "?summable")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   278
  and "0 < a 0 \<longrightarrow> (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i=0..<2*n. -1^i * a i .. \<Sum>i=0..<2*n+1. -1^i * a i})" (is "?pos")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   279
  and "a 0 < 0 \<longrightarrow> (\<forall>n. (\<Sum>i. -1^i*a i) \<in> { \<Sum>i=0..<2*n+1. -1^i * a i .. \<Sum>i=0..<2*n. -1^i * a i})" (is "?neg")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   280
  and "(\<lambda>n. \<Sum>i=0..<2*n. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?f")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   281
  and "(\<lambda>n. \<Sum>i=0..<2*n+1. -1^i*a i) ----> (\<Sum>i. -1^i*a i)" (is "?g")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   282
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   283
  have "?summable \<and> ?pos \<and> ?neg \<and> ?f \<and> ?g"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   284
  proof (cases "(\<forall> n. 0 \<le> a n) \<and> (\<forall>m. \<forall>n\<ge>m. a n \<le> a m)")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   285
    case True
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   286
    hence ord: "\<And>n m. m \<le> n \<Longrightarrow> a n \<le> a m" and ge0: "\<And> n. 0 \<le> a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   287
    { fix n have "a (Suc n) \<le> a n" using ord[where n="Suc n" and m=n] by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   288
    note leibniz = summable_Leibniz'[OF `a ----> 0` ge0] and mono = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   289
    from leibniz[OF mono]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   290
    show ?thesis using `0 \<le> a 0` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   291
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   292
    let ?a = "\<lambda> n. - a n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   293
    case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   294
    with monoseq_le[OF `monoseq a` `a ----> 0`]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   295
    have "(\<forall> n. a n \<le> 0) \<and> (\<forall>m. \<forall>n\<ge>m. a m \<le> a n)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   296
    hence ord: "\<And>n m. m \<le> n \<Longrightarrow> ?a n \<le> ?a m" and ge0: "\<And> n. 0 \<le> ?a n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   297
    { fix n have "?a (Suc n) \<le> ?a n" using ord[where n="Suc n" and m=n] by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   298
    note monotone = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   299
    note leibniz = summable_Leibniz'[OF _ ge0, of "\<lambda>x. x", OF LIMSEQ_minus[OF `a ----> 0`, unfolded minus_zero] monotone]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   300
    have "summable (\<lambda> n. (-1)^n * ?a n)" using leibniz(1) by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   301
    then obtain l where "(\<lambda> n. (-1)^n * ?a n) sums l" unfolding summable_def by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   302
    from this[THEN sums_minus]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   303
    have "(\<lambda> n. (-1)^n * a n) sums -l" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   304
    hence ?summable unfolding summable_def by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   305
    moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   306
    have "\<And> a b :: real. \<bar> - a - - b \<bar> = \<bar>a - b\<bar>" unfolding minus_diff_minus by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   307
    
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   308
    from suminf_minus[OF leibniz(1), unfolded mult_minus_right minus_minus]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   309
    have move_minus: "(\<Sum>n. - (-1 ^ n * a n)) = - (\<Sum>n. -1 ^ n * a n)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   310
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   311
    have ?pos using `0 \<le> ?a 0` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   312
    moreover have ?neg using leibniz(2,4) unfolding mult_minus_right setsum_negf move_minus neg_le_iff_le by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   313
    moreover have ?f and ?g using leibniz(3,5)[unfolded mult_minus_right setsum_negf move_minus, THEN LIMSEQ_minus_cancel] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   314
    ultimately show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   315
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   316
  from this[THEN conjunct1] this[THEN conjunct2, THEN conjunct1] this[THEN conjunct2, THEN conjunct2, THEN conjunct1] this[THEN conjunct2, THEN conjunct2, THEN conjunct2, THEN conjunct1]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   317
       this[THEN conjunct2, THEN conjunct2, THEN conjunct2, THEN conjunct2]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   318
  show ?summable and ?pos and ?neg and ?f and ?g .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   319
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   320
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
   321
subsection {* Term-by-Term Differentiability of Power Series *}
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   322
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   323
definition
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   324
  diffs :: "(nat => 'a::ring_1) => nat => 'a" where
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   325
  "diffs c = (%n. of_nat (Suc n) * c(Suc n))"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   326
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   327
text{*Lemma about distributing negation over it*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   328
lemma diffs_minus: "diffs (%n. - c n) = (%n. - diffs c n)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   329
by (simp add: diffs_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   330
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   331
lemma sums_Suc_imp:
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   332
  assumes f: "f 0 = 0"
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   333
  shows "(\<lambda>n. f (Suc n)) sums s \<Longrightarrow> (\<lambda>n. f n) sums s"
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   334
unfolding sums_def
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   335
apply (rule LIMSEQ_imp_Suc)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   336
apply (subst setsum_shift_lb_Suc0_0_upt [where f=f, OF f, symmetric])
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   337
apply (simp only: setsum_shift_bounds_Suc_ivl)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   338
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   339
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   340
lemma diffs_equiv:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   341
     "summable (%n. (diffs c)(n) * (x ^ n)) ==>  
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   342
      (%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums  
15546
5188ce7316b7 suminf -> \<Sum>
nipkow
parents: 15544
diff changeset
   343
         (\<Sum>n. (diffs c)(n) * (x ^ n))"
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   344
unfolding diffs_def
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   345
apply (drule summable_sums)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   346
apply (rule sums_Suc_imp, simp_all)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   347
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   348
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   349
lemma lemma_termdiff1:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   350
  fixes z :: "'a :: {monoid_mult,comm_ring}" shows
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   351
  "(\<Sum>p=0..<m. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) =  
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   352
   (\<Sum>p=0..<m. (z ^ p) * (((z + h) ^ (m - p)) - (z ^ (m - p))))"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
   353
by(auto simp add: algebra_simps power_add [symmetric] cong: strong_setsum_cong)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   354
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   355
lemma sumr_diff_mult_const2:
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   356
  "setsum f {0..<n} - of_nat n * (r::'a::ring_1) = (\<Sum>i = 0..<n. f i - r)"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   357
by (simp add: setsum_subtractf)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   358
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   359
lemma lemma_termdiff2:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   360
  fixes h :: "'a :: {field}"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   361
  assumes h: "h \<noteq> 0" shows
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   362
  "((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) =
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   363
   h * (\<Sum>p=0..< n - Suc 0. \<Sum>q=0..< n - Suc 0 - p.
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   364
        (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs")
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   365
apply (subgoal_tac "h * ?lhs = h * ?rhs", simp add: h)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   366
apply (simp add: right_diff_distrib diff_divide_distrib h)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   367
apply (simp add: mult_assoc [symmetric])
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   368
apply (cases "n", simp)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   369
apply (simp add: lemma_realpow_diff_sumr2 h
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   370
                 right_diff_distrib [symmetric] mult_assoc
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
   371
            del: power_Suc setsum_op_ivl_Suc of_nat_Suc)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   372
apply (subst lemma_realpow_rev_sumr)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   373
apply (subst sumr_diff_mult_const2)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   374
apply simp
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   375
apply (simp only: lemma_termdiff1 setsum_right_distrib)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   376
apply (rule setsum_cong [OF refl])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   377
apply (simp add: diff_minus [symmetric] less_iff_Suc_add)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   378
apply (clarify)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   379
apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
   380
            del: setsum_op_ivl_Suc power_Suc)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   381
apply (subst mult_assoc [symmetric], subst power_add [symmetric])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   382
apply (simp add: mult_ac)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   383
done
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   384
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   385
lemma real_setsum_nat_ivl_bounded2:
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   386
  fixes K :: "'a::ordered_semidom"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   387
  assumes f: "\<And>p::nat. p < n \<Longrightarrow> f p \<le> K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   388
  assumes K: "0 \<le> K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   389
  shows "setsum f {0..<n-k} \<le> of_nat n * K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   390
apply (rule order_trans [OF setsum_mono])
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   391
apply (rule f, simp)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   392
apply (simp add: mult_right_mono K)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   393
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   394
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   395
lemma lemma_termdiff3:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   396
  fixes h z :: "'a::{real_normed_field}"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   397
  assumes 1: "h \<noteq> 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   398
  assumes 2: "norm z \<le> K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   399
  assumes 3: "norm (z + h) \<le> K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   400
  shows "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0))
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   401
          \<le> of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   402
proof -
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   403
  have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) =
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   404
        norm (\<Sum>p = 0..<n - Suc 0. \<Sum>q = 0..<n - Suc 0 - p.
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   405
          (z + h) ^ q * z ^ (n - 2 - q)) * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   406
    apply (subst lemma_termdiff2 [OF 1])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   407
    apply (subst norm_mult)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   408
    apply (rule mult_commute)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   409
    done
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   410
  also have "\<dots> \<le> of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   411
  proof (rule mult_right_mono [OF _ norm_ge_zero])
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   412
    from norm_ge_zero 2 have K: "0 \<le> K" by (rule order_trans)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   413
    have le_Kn: "\<And>i j n. i + j = n \<Longrightarrow> norm ((z + h) ^ i * z ^ j) \<le> K ^ n"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   414
      apply (erule subst)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   415
      apply (simp only: norm_mult norm_power power_add)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   416
      apply (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   417
      done
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   418
    show "norm (\<Sum>p = 0..<n - Suc 0. \<Sum>q = 0..<n - Suc 0 - p.
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   419
              (z + h) ^ q * z ^ (n - 2 - q))
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   420
          \<le> of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   421
      apply (intro
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   422
         order_trans [OF norm_setsum]
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   423
         real_setsum_nat_ivl_bounded2
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   424
         mult_nonneg_nonneg
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   425
         zero_le_imp_of_nat
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   426
         zero_le_power K)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   427
      apply (rule le_Kn, simp)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   428
      done
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   429
  qed
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   430
  also have "\<dots> = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   431
    by (simp only: mult_assoc)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   432
  finally show ?thesis .
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   433
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   434
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   435
lemma lemma_termdiff4:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   436
  fixes f :: "'a::{real_normed_field} \<Rightarrow>
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   437
              'b::real_normed_vector"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   438
  assumes k: "0 < (k::real)"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   439
  assumes le: "\<And>h. \<lbrakk>h \<noteq> 0; norm h < k\<rbrakk> \<Longrightarrow> norm (f h) \<le> K * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   440
  shows "f -- 0 --> 0"
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   441
unfolding LIM_def diff_0_right
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   442
proof (safe)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   443
  let ?h = "of_real (k / 2)::'a"
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   444
  have "?h \<noteq> 0" and "norm ?h < k" using k by simp_all
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   445
  hence "norm (f ?h) \<le> K * norm ?h" by (rule le)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   446
  hence "0 \<le> K * norm ?h" by (rule order_trans [OF norm_ge_zero])
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   447
  hence zero_le_K: "0 \<le> K" using k by (simp add: zero_le_mult_iff)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   448
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   449
  fix r::real assume r: "0 < r"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   450
  show "\<exists>s. 0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < s \<longrightarrow> norm (f x) < r)"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   451
  proof (cases)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   452
    assume "K = 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   453
    with k r le have "0 < k \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < k \<longrightarrow> norm (f x) < r)"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   454
      by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   455
    thus "\<exists>s. 0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < s \<longrightarrow> norm (f x) < r)" ..
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   456
  next
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   457
    assume K_neq_zero: "K \<noteq> 0"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   458
    with zero_le_K have K: "0 < K" by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   459
    show "\<exists>s. 0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < s \<longrightarrow> norm (f x) < r)"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   460
    proof (rule exI, safe)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   461
      from k r K show "0 < min k (r * inverse K / 2)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   462
        by (simp add: mult_pos_pos positive_imp_inverse_positive)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   463
    next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   464
      fix x::'a
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   465
      assume x1: "x \<noteq> 0" and x2: "norm x < min k (r * inverse K / 2)"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   466
      from x2 have x3: "norm x < k" and x4: "norm x < r * inverse K / 2"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   467
        by simp_all
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   468
      from x1 x3 le have "norm (f x) \<le> K * norm x" by simp
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   469
      also from x4 K have "K * norm x < K * (r * inverse K / 2)"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   470
        by (rule mult_strict_left_mono)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   471
      also have "\<dots> = r / 2"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   472
        using K_neq_zero by simp
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   473
      also have "r / 2 < r"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   474
        using r by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   475
      finally show "norm (f x) < r" .
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   476
    qed
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   477
  qed
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   478
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   479
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   480
lemma lemma_termdiff5:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   481
  fixes g :: "'a::{real_normed_field} \<Rightarrow>
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   482
              nat \<Rightarrow> 'b::banach"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   483
  assumes k: "0 < (k::real)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   484
  assumes f: "summable f"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   485
  assumes le: "\<And>h n. \<lbrakk>h \<noteq> 0; norm h < k\<rbrakk> \<Longrightarrow> norm (g h n) \<le> f n * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   486
  shows "(\<lambda>h. suminf (g h)) -- 0 --> 0"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   487
proof (rule lemma_termdiff4 [OF k])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   488
  fix h::'a assume "h \<noteq> 0" and "norm h < k"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   489
  hence A: "\<forall>n. norm (g h n) \<le> f n * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   490
    by (simp add: le)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   491
  hence "\<exists>N. \<forall>n\<ge>N. norm (norm (g h n)) \<le> f n * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   492
    by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   493
  moreover from f have B: "summable (\<lambda>n. f n * norm h)"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   494
    by (rule summable_mult2)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   495
  ultimately have C: "summable (\<lambda>n. norm (g h n))"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   496
    by (rule summable_comparison_test)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   497
  hence "norm (suminf (g h)) \<le> (\<Sum>n. norm (g h n))"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   498
    by (rule summable_norm)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   499
  also from A C B have "(\<Sum>n. norm (g h n)) \<le> (\<Sum>n. f n * norm h)"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   500
    by (rule summable_le)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   501
  also from f have "(\<Sum>n. f n * norm h) = suminf f * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   502
    by (rule suminf_mult2 [symmetric])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   503
  finally show "norm (suminf (g h)) \<le> suminf f * norm h" .
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   504
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   505
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   506
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   507
text{* FIXME: Long proofs*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   508
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   509
lemma termdiffs_aux:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   510
  fixes x :: "'a::{real_normed_field,banach}"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   511
  assumes 1: "summable (\<lambda>n. diffs (diffs c) n * K ^ n)"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   512
  assumes 2: "norm x < norm K"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   513
  shows "(\<lambda>h. \<Sum>n. c n * (((x + h) ^ n - x ^ n) / h
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   514
             - of_nat n * x ^ (n - Suc 0))) -- 0 --> 0"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   515
proof -
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   516
  from dense [OF 2]
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   517
  obtain r where r1: "norm x < r" and r2: "r < norm K" by fast
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   518
  from norm_ge_zero r1 have r: "0 < r"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   519
    by (rule order_le_less_trans)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   520
  hence r_neq_0: "r \<noteq> 0" by simp
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   521
  show ?thesis
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   522
  proof (rule lemma_termdiff5)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   523
    show "0 < r - norm x" using r1 by simp
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   524
  next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   525
    from r r2 have "norm (of_real r::'a) < norm K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   526
      by simp
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   527
    with 1 have "summable (\<lambda>n. norm (diffs (diffs c) n * (of_real r ^ n)))"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   528
      by (rule powser_insidea)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   529
    hence "summable (\<lambda>n. diffs (diffs (\<lambda>n. norm (c n))) n * r ^ n)"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   530
      using r
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   531
      by (simp add: diffs_def norm_mult norm_power del: of_nat_Suc)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   532
    hence "summable (\<lambda>n. of_nat n * diffs (\<lambda>n. norm (c n)) n * r ^ (n - Suc 0))"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   533
      by (rule diffs_equiv [THEN sums_summable])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   534
    also have "(\<lambda>n. of_nat n * diffs (\<lambda>n. norm (c n)) n * r ^ (n - Suc 0))
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   535
      = (\<lambda>n. diffs (%m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   536
      apply (rule ext)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   537
      apply (simp add: diffs_def)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   538
      apply (case_tac n, simp_all add: r_neq_0)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   539
      done
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   540
    finally have "summable 
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   541
      (\<lambda>n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   542
      by (rule diffs_equiv [THEN sums_summable])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   543
    also have
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   544
      "(\<lambda>n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) *
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   545
           r ^ (n - Suc 0)) =
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   546
       (\<lambda>n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   547
      apply (rule ext)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   548
      apply (case_tac "n", simp)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   549
      apply (case_tac "nat", simp)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   550
      apply (simp add: r_neq_0)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   551
      done
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   552
    finally show
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   553
      "summable (\<lambda>n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" .
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   554
  next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   555
    fix h::'a and n::nat
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   556
    assume h: "h \<noteq> 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   557
    assume "norm h < r - norm x"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   558
    hence "norm x + norm h < r" by simp
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   559
    with norm_triangle_ineq have xh: "norm (x + h) < r"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   560
      by (rule order_le_less_trans)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   561
    show "norm (c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)))
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   562
          \<le> norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   563
      apply (simp only: norm_mult mult_assoc)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   564
      apply (rule mult_left_mono [OF _ norm_ge_zero])
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   565
      apply (simp (no_asm) add: mult_assoc [symmetric])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   566
      apply (rule lemma_termdiff3)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   567
      apply (rule h)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   568
      apply (rule r1 [THEN order_less_imp_le])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   569
      apply (rule xh [THEN order_less_imp_le])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   570
      done
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   571
  qed
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   572
qed
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
   573
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   574
lemma termdiffs:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   575
  fixes K x :: "'a::{real_normed_field,banach}"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   576
  assumes 1: "summable (\<lambda>n. c n * K ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   577
  assumes 2: "summable (\<lambda>n. (diffs c) n * K ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   578
  assumes 3: "summable (\<lambda>n. (diffs (diffs c)) n * K ^ n)"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   579
  assumes 4: "norm x < norm K"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   580
  shows "DERIV (\<lambda>x. \<Sum>n. c n * x ^ n) x :> (\<Sum>n. (diffs c) n * x ^ n)"
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   581
unfolding deriv_def
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   582
proof (rule LIM_zero_cancel)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   583
  show "(\<lambda>h. (suminf (\<lambda>n. c n * (x + h) ^ n) - suminf (\<lambda>n. c n * x ^ n)) / h
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   584
            - suminf (\<lambda>n. diffs c n * x ^ n)) -- 0 --> 0"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   585
  proof (rule LIM_equal2)
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   586
    show "0 < norm K - norm x" using 4 by (simp add: less_diff_eq)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   587
  next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   588
    fix h :: 'a
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   589
    assume "h \<noteq> 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   590
    assume "norm (h - 0) < norm K - norm x"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   591
    hence "norm x + norm h < norm K" by simp
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   592
    hence 5: "norm (x + h) < norm K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   593
      by (rule norm_triangle_ineq [THEN order_le_less_trans])
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   594
    have A: "summable (\<lambda>n. c n * x ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   595
      by (rule powser_inside [OF 1 4])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   596
    have B: "summable (\<lambda>n. c n * (x + h) ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   597
      by (rule powser_inside [OF 1 5])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   598
    have C: "summable (\<lambda>n. diffs c n * x ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   599
      by (rule powser_inside [OF 2 4])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   600
    show "((\<Sum>n. c n * (x + h) ^ n) - (\<Sum>n. c n * x ^ n)) / h
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   601
             - (\<Sum>n. diffs c n * x ^ n) = 
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   602
          (\<Sum>n. c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)))"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   603
      apply (subst sums_unique [OF diffs_equiv [OF C]])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   604
      apply (subst suminf_diff [OF B A])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   605
      apply (subst suminf_divide [symmetric])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   606
      apply (rule summable_diff [OF B A])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   607
      apply (subst suminf_diff)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   608
      apply (rule summable_divide)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   609
      apply (rule summable_diff [OF B A])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   610
      apply (rule sums_summable [OF diffs_equiv [OF C]])
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   611
      apply (rule arg_cong [where f="suminf"], rule ext)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
   612
      apply (simp add: algebra_simps)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   613
      done
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   614
  next
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   615
    show "(\<lambda>h. \<Sum>n. c n * (((x + h) ^ n - x ^ n) / h -
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   616
               of_nat n * x ^ (n - Suc 0))) -- 0 --> 0"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   617
        by (rule termdiffs_aux [OF 3 4])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   618
  qed
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   619
qed
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   620
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   621
29695
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   622
subsection{* Some properties of factorials *}
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   623
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   624
lemma real_of_nat_fact_not_zero [simp]: "real (fact n) \<noteq> 0"
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   625
by auto
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   626
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   627
lemma real_of_nat_fact_gt_zero [simp]: "0 < real(fact n)"
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   628
by auto
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   629
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   630
lemma real_of_nat_fact_ge_zero [simp]: "0 \<le> real(fact n)"
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   631
by simp
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   632
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   633
lemma inv_real_of_nat_fact_gt_zero [simp]: "0 < inverse (real (fact n))"
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   634
by (auto simp add: positive_imp_inverse_positive)
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   635
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   636
lemma inv_real_of_nat_fact_ge_zero [simp]: "0 \<le> inverse (real (fact n))"
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   637
by (auto intro: order_less_imp_le)
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   638
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   639
subsection {* Derivability of power series *}
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   640
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   641
lemma DERIV_series': fixes f :: "real \<Rightarrow> nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   642
  assumes DERIV_f: "\<And> n. DERIV (\<lambda> x. f x n) x0 :> (f' x0 n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   643
  and allf_summable: "\<And> x. x \<in> {a <..< b} \<Longrightarrow> summable (f x)" and x0_in_I: "x0 \<in> {a <..< b}"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   644
  and "summable (f' x0)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   645
  and "summable L" and L_def: "\<And> n x y. \<lbrakk> x \<in> { a <..< b} ; y \<in> { a <..< b} \<rbrakk> \<Longrightarrow> \<bar> f x n - f y n \<bar> \<le> L n * \<bar> x - y \<bar>"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   646
  shows "DERIV (\<lambda> x. suminf (f x)) x0 :> (suminf (f' x0))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   647
  unfolding deriv_def
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   648
proof (rule LIM_I)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   649
  fix r :: real assume "0 < r" hence "0 < r/3" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   650
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   651
  obtain N_L where N_L: "\<And> n. N_L \<le> n \<Longrightarrow> \<bar> \<Sum> i. L (i + n) \<bar> < r/3" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   652
    using suminf_exist_split[OF `0 < r/3` `summable L`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   653
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   654
  obtain N_f' where N_f': "\<And> n. N_f' \<le> n \<Longrightarrow> \<bar> \<Sum> i. f' x0 (i + n) \<bar> < r/3" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   655
    using suminf_exist_split[OF `0 < r/3` `summable (f' x0)`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   656
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   657
  let ?N = "Suc (max N_L N_f')"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   658
  have "\<bar> \<Sum> i. f' x0 (i + ?N) \<bar> < r/3" (is "?f'_part < r/3") and
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   659
    L_estimate: "\<bar> \<Sum> i. L (i + ?N) \<bar> < r/3" using N_L[of "?N"] and N_f' [of "?N"] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   660
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   661
  let "?diff i x" = "(f (x0 + x) i - f x0 i) / x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   662
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   663
  let ?r = "r / (3 * real ?N)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   664
  have "0 < 3 * real ?N" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   665
  from divide_pos_pos[OF `0 < r` this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   666
  have "0 < ?r" .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   667
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   668
  let "?s n" = "SOME s. 0 < s \<and> (\<forall> x. x \<noteq> 0 \<and> \<bar> x \<bar> < s \<longrightarrow> \<bar> ?diff n x - f' x0 n \<bar> < ?r)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   669
  def S' \<equiv> "Min (?s ` { 0 ..< ?N })"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   670
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   671
  have "0 < S'" unfolding S'_def
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   672
  proof (rule iffD2[OF Min_gr_iff])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   673
    show "\<forall> x \<in> (?s ` { 0 ..< ?N }). 0 < x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   674
    proof (rule ballI)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   675
      fix x assume "x \<in> ?s ` {0..<?N}"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   676
      then obtain n where "x = ?s n" and "n \<in> {0..<?N}" using image_iff[THEN iffD1] by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   677
      from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def] 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   678
      obtain s where s_bound: "0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < s \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   679
      have "0 < ?s n" by (rule someI2[where a=s], auto simp add: s_bound)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   680
      thus "0 < x" unfolding `x = ?s n` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   681
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   682
  qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   683
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   684
  def S \<equiv> "min (min (x0 - a) (b - x0)) S'"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   685
  hence "0 < S" and S_a: "S \<le> x0 - a" and S_b: "S \<le> b - x0" and "S \<le> S'" using x0_in_I and `0 < S'`
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   686
    by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   687
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   688
  { fix x assume "x \<noteq> 0" and "\<bar> x \<bar> < S"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   689
    hence x_in_I: "x0 + x \<in> { a <..< b }" using S_a S_b by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   690
    
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   691
    note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   692
    note div_smbl = summable_divide[OF diff_smbl]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   693
    note all_smbl = summable_diff[OF div_smbl `summable (f' x0)`]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   694
    note ign = summable_ignore_initial_segment[where k="?N"]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   695
    note diff_shft_smbl = summable_diff[OF ign[OF allf_summable[OF x_in_I]] ign[OF allf_summable[OF x0_in_I]]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   696
    note div_shft_smbl = summable_divide[OF diff_shft_smbl]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   697
    note all_shft_smbl = summable_diff[OF div_smbl ign[OF `summable (f' x0)`]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   698
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   699
    { fix n
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   700
      have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   701
	using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] unfolding abs_divide .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   702
      hence "\<bar> ( \<bar> ?diff (n + ?N) x \<bar>) \<bar> \<le> L (n + ?N)" using `x \<noteq> 0` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   703
    } note L_ge = summable_le2[OF allI[OF this] ign[OF `summable L`]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   704
    from order_trans[OF summable_rabs[OF conjunct1[OF L_ge]] L_ge[THEN conjunct2]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   705
    have "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> (\<Sum> i. L (i + ?N))" .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   706
    hence "\<bar> \<Sum> i. ?diff (i + ?N) x \<bar> \<le> r / 3" (is "?L_part \<le> r/3") using L_estimate by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   707
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   708
    have "\<bar>\<Sum>n \<in> { 0 ..< ?N}. ?diff n x - f' x0 n \<bar> \<le> (\<Sum>n \<in> { 0 ..< ?N}. \<bar>?diff n x - f' x0 n \<bar>)" ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   709
    also have "\<dots> < (\<Sum>n \<in> { 0 ..< ?N}. ?r)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   710
    proof (rule setsum_strict_mono)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   711
      fix n assume "n \<in> { 0 ..< ?N}"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   712
      have "\<bar> x \<bar> < S" using `\<bar> x \<bar> < S` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   713
      also have "S \<le> S'" using `S \<le> S'` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   714
      also have "S' \<le> ?s n" unfolding S'_def 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   715
      proof (rule Min_le_iff[THEN iffD2])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   716
	have "?s n \<in> (?s ` {0..<?N}) \<and> ?s n \<le> ?s n" using `n \<in> { 0 ..< ?N}` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   717
	thus "\<exists> a \<in> (?s ` {0..<?N}). a \<le> ?s n" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   718
      qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   719
      finally have "\<bar> x \<bar> < ?s n" .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   720
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   721
      from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   722
      have "\<forall>x. x \<noteq> 0 \<and> \<bar>x\<bar> < ?s n \<longrightarrow> \<bar>?diff n x - f' x0 n\<bar> < ?r" .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   723
      with `x \<noteq> 0` and `\<bar>x\<bar> < ?s n`
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   724
      show "\<bar>?diff n x - f' x0 n\<bar> < ?r" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   725
    qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   726
    also have "\<dots> = of_nat (card {0 ..< ?N}) * ?r" by (rule setsum_constant)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   727
    also have "\<dots> = real ?N * ?r" unfolding real_eq_of_nat by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   728
    also have "\<dots> = r/3" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   729
    finally have "\<bar>\<Sum>n \<in> { 0 ..< ?N}. ?diff n x - f' x0 n \<bar> < r / 3" (is "?diff_part < r / 3") .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   730
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   731
    from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   732
    have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> = 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   733
                    \<bar> \<Sum>n. ?diff n x - f' x0 n \<bar>" unfolding suminf_diff[OF div_smbl `summable (f' x0)`, symmetric] using suminf_divide[OF diff_smbl, symmetric] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   734
    also have "\<dots> \<le> ?diff_part + \<bar> (\<Sum>n. ?diff (n + ?N) x) - (\<Sum> n. f' x0 (n + ?N)) \<bar>" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF `summable (f' x0)`]] by (rule abs_triangle_ineq)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   735
    also have "\<dots> \<le> ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   736
    also have "\<dots> < r /3 + r/3 + r/3" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   737
      using `?diff_part < r/3` `?L_part \<le> r/3` and `?f'_part < r/3` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   738
    finally have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> < r"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   739
      by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   740
  } thus "\<exists> s > 0. \<forall> x. x \<noteq> 0 \<and> norm (x - 0) < s \<longrightarrow> 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   741
      norm (((\<Sum>n. f (x0 + x) n) - (\<Sum>n. f x0 n)) / x - (\<Sum>n. f' x0 n)) < r" using `0 < S`
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   742
    unfolding real_norm_def diff_0_right by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   743
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   744
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   745
lemma DERIV_power_series': fixes f :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   746
  assumes converges: "\<And> x. x \<in> {-R <..< R} \<Longrightarrow> summable (\<lambda> n. f n * real (Suc n) * x^n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   747
  and x0_in_I: "x0 \<in> {-R <..< R}" and "0 < R"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   748
  shows "DERIV (\<lambda> x. (\<Sum> n. f n * x^(Suc n))) x0 :> (\<Sum> n. f n * real (Suc n) * x0^n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   749
  (is "DERIV (\<lambda> x. (suminf (?f x))) x0 :> (suminf (?f' x0))")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   750
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   751
  { fix R' assume "0 < R'" and "R' < R" and "-R' < x0" and "x0 < R'"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   752
    hence "x0 \<in> {-R' <..< R'}" and "R' \<in> {-R <..< R}" and "x0 \<in> {-R <..< R}" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   753
    have "DERIV (\<lambda> x. (suminf (?f x))) x0 :> (suminf (?f' x0))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   754
    proof (rule DERIV_series')
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   755
      show "summable (\<lambda> n. \<bar>f n * real (Suc n) * R'^n\<bar>)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   756
      proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   757
	have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using `0 < R'` `0 < R` `R' < R` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   758
	hence in_Rball: "(R' + R) / 2 \<in> {-R <..< R}" using `R' < R` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   759
	have "norm R' < norm ((R' + R) / 2)" using `0 < R'` `0 < R` `R' < R` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   760
	from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   761
      qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   762
      { fix n x y assume "x \<in> {-R' <..< R'}" and "y \<in> {-R' <..< R'}"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   763
	show "\<bar>?f x n - ?f y n\<bar> \<le> \<bar>f n * real (Suc n) * R'^n\<bar> * \<bar>x-y\<bar>"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   764
	proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   765
	  have "\<bar>f n * x ^ (Suc n) - f n * y ^ (Suc n)\<bar> = (\<bar>f n\<bar> * \<bar>x-y\<bar>) * \<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar>" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   766
	    unfolding right_diff_distrib[symmetric] lemma_realpow_diff_sumr2 abs_mult by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   767
	  also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   768
	  proof (rule mult_left_mono)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   769
	    have "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> (\<Sum>p = 0..<Suc n. \<bar>x ^ p * y ^ (n - p)\<bar>)" by (rule setsum_abs)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   770
	    also have "\<dots> \<le> (\<Sum>p = 0..<Suc n. R' ^ n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   771
	    proof (rule setsum_mono)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   772
	      fix p assume "p \<in> {0..<Suc n}" hence "p \<le> n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   773
	      { fix n fix x :: real assume "x \<in> {-R'<..<R'}"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   774
		hence "\<bar>x\<bar> \<le> R'"  by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   775
		hence "\<bar>x^n\<bar> \<le> R'^n" unfolding power_abs by (rule power_mono, auto)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   776
	      } from mult_mono[OF this[OF `x \<in> {-R'<..<R'}`, of p] this[OF `y \<in> {-R'<..<R'}`, of "n-p"]] `0 < R'`
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   777
	      have "\<bar>x^p * y^(n-p)\<bar> \<le> R'^p * R'^(n-p)" unfolding abs_mult by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   778
	      thus "\<bar>x^p * y^(n-p)\<bar> \<le> R'^n" unfolding power_add[symmetric] using `p \<le> n` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   779
	    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   780
	    also have "\<dots> = real (Suc n) * R' ^ n" unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   781
	    finally show "\<bar>\<Sum>p = 0..<Suc n. x ^ p * y ^ (n - p)\<bar> \<le> \<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>" unfolding abs_real_of_nat_cancel abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF `0 < R'`]]] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   782
	    show "0 \<le> \<bar>f n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult[symmetric] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   783
	  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   784
	  also have "\<dots> = \<bar>f n * real (Suc n) * R' ^ n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult real_mult_assoc[symmetric] by algebra
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   785
	  finally show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   786
	qed }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   787
      { fix n
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   788
	from DERIV_pow[of "Suc n" x0, THEN DERIV_cmult[where c="f n"]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   789
	show "DERIV (\<lambda> x. ?f x n) x0 :> (?f' x0 n)" unfolding real_mult_assoc by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   790
      { fix x assume "x \<in> {-R' <..< R'}" hence "R' \<in> {-R <..< R}" and "norm x < norm R'" using assms `R' < R` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   791
	have "summable (\<lambda> n. f n * x^n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   792
	proof (rule summable_le2[THEN conjunct1, OF _ powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`]], rule allI)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   793
	  fix n
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   794
	  have le: "\<bar>f n\<bar> * 1 \<le> \<bar>f n\<bar> * real (Suc n)" by (rule mult_left_mono, auto)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   795
	  show "\<bar>f n * x ^ n\<bar> \<le> norm (f n * real (Suc n) * x ^ n)" unfolding real_norm_def abs_mult
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   796
	    by (rule mult_right_mono, auto simp add: le[unfolded mult_1_right])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   797
	qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   798
	from this[THEN summable_mult2[where c=x], unfolded real_mult_assoc, unfolded real_mult_commute]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   799
	show "summable (?f x)" by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   800
      show "summable (?f' x0)" using converges[OF `x0 \<in> {-R <..< R}`] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   801
      show "x0 \<in> {-R' <..< R'}" using `x0 \<in> {-R' <..< R'}` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   802
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   803
  } note for_subinterval = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   804
  let ?R = "(R + \<bar>x0\<bar>) / 2"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   805
  have "\<bar>x0\<bar> < ?R" using assms by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   806
  hence "- ?R < x0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   807
  proof (cases "x0 < 0")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   808
    case True
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   809
    hence "- x0 < ?R" using `\<bar>x0\<bar> < ?R` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   810
    thus ?thesis unfolding neg_less_iff_less[symmetric, of "- x0"] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   811
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   812
    case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   813
    have "- ?R < 0" using assms by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   814
    also have "\<dots> \<le> x0" using False by auto 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   815
    finally show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   816
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   817
  hence "0 < ?R" "?R < R" "- ?R < x0" and "x0 < ?R" using assms by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   818
  from for_subinterval[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   819
  show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   820
qed
29695
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   821
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
   822
subsection {* Exponential Function *}
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   823
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   824
definition
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   825
  exp :: "'a \<Rightarrow> 'a::{real_normed_field,banach}" where
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   826
  "exp x = (\<Sum>n. x ^ n /\<^sub>R real (fact n))"
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   827
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   828
lemma summable_exp_generic:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   829
  fixes x :: "'a::{real_normed_algebra_1,banach}"
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   830
  defines S_def: "S \<equiv> \<lambda>n. x ^ n /\<^sub>R real (fact n)"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   831
  shows "summable S"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   832
proof -
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   833
  have S_Suc: "\<And>n. S (Suc n) = (x * S n) /\<^sub>R real (Suc n)"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
   834
    unfolding S_def by (simp del: mult_Suc)
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   835
  obtain r :: real where r0: "0 < r" and r1: "r < 1"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   836
    using dense [OF zero_less_one] by fast
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   837
  obtain N :: nat where N: "norm x < real N * r"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   838
    using reals_Archimedean3 [OF r0] by fast
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   839
  from r1 show ?thesis
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   840
  proof (rule ratio_test [rule_format])
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   841
    fix n :: nat
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   842
    assume n: "N \<le> n"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   843
    have "norm x \<le> real N * r"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   844
      using N by (rule order_less_imp_le)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   845
    also have "real N * r \<le> real (Suc n) * r"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   846
      using r0 n by (simp add: mult_right_mono)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   847
    finally have "norm x * norm (S n) \<le> real (Suc n) * r * norm (S n)"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   848
      using norm_ge_zero by (rule mult_right_mono)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   849
    hence "norm (x * S n) \<le> real (Suc n) * r * norm (S n)"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   850
      by (rule order_trans [OF norm_mult_ineq])
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   851
    hence "norm (x * S n) / real (Suc n) \<le> r * norm (S n)"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   852
      by (simp add: pos_divide_le_eq mult_ac)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   853
    thus "norm (S (Suc n)) \<le> r * norm (S n)"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   854
      by (simp add: S_Suc norm_scaleR inverse_eq_divide)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   855
  qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   856
qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   857
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   858
lemma summable_norm_exp:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   859
  fixes x :: "'a::{real_normed_algebra_1,banach}"
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   860
  shows "summable (\<lambda>n. norm (x ^ n /\<^sub>R real (fact n)))"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   861
proof (rule summable_norm_comparison_test [OF exI, rule_format])
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   862
  show "summable (\<lambda>n. norm x ^ n /\<^sub>R real (fact n))"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   863
    by (rule summable_exp_generic)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   864
next
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   865
  fix n show "norm (x ^ n /\<^sub>R real (fact n)) \<le> norm x ^ n /\<^sub>R real (fact n)"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   866
    by (simp add: norm_scaleR norm_power_ineq)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   867
qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   868
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   869
lemma summable_exp: "summable (%n. inverse (real (fact n)) * x ^ n)"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   870
by (insert summable_exp_generic [where x=x], simp)
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   871
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   872
lemma exp_converges: "(\<lambda>n. x ^ n /\<^sub>R real (fact n)) sums exp x"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   873
unfolding exp_def by (rule summable_exp_generic [THEN summable_sums])
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   874
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   875
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   876
lemma exp_fdiffs: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   877
      "diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))"
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23413
diff changeset
   878
by (simp add: diffs_def mult_assoc [symmetric] real_of_nat_def of_nat_mult
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   879
         del: mult_Suc of_nat_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   880
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   881
lemma diffs_of_real: "diffs (\<lambda>n. of_real (f n)) = (\<lambda>n. of_real (diffs f n))"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   882
by (simp add: diffs_def)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   883
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   884
lemma lemma_exp_ext: "exp = (\<lambda>x. \<Sum>n. x ^ n /\<^sub>R real (fact n))"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   885
by (auto intro!: ext simp add: exp_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   886
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   887
lemma DERIV_exp [simp]: "DERIV exp x :> exp(x)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   888
apply (simp add: exp_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   889
apply (subst lemma_exp_ext)
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   890
apply (subgoal_tac "DERIV (\<lambda>u. \<Sum>n. of_real (inverse (real (fact n))) * u ^ n) x :> (\<Sum>n. diffs (\<lambda>n. of_real (inverse (real (fact n)))) n * x ^ n)")
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   891
apply (rule_tac [2] K = "of_real (1 + norm x)" in termdiffs)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   892
apply (simp_all only: diffs_of_real scaleR_conv_of_real exp_fdiffs)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   893
apply (rule exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real])+
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   894
apply (simp del: of_real_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   895
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   896
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
   897
lemma isCont_exp [simp]: "isCont exp x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
   898
by (rule DERIV_exp [THEN DERIV_isCont])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
   899
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
   900
29167
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   901
subsubsection {* Properties of the Exponential Function *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   902
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   903
lemma powser_zero:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   904
  fixes f :: "nat \<Rightarrow> 'a::{real_normed_algebra_1}"
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   905
  shows "(\<Sum>n. f n * 0 ^ n) = f 0"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   906
proof -
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   907
  have "(\<Sum>n = 0..<1. f n * 0 ^ n) = (\<Sum>n. f n * 0 ^ n)"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   908
    by (rule sums_unique [OF series_zero], simp add: power_0_left)
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
   909
  thus ?thesis unfolding One_nat_def by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   910
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   911
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   912
lemma exp_zero [simp]: "exp 0 = 1"
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   913
unfolding exp_def by (simp add: scaleR_conv_of_real powser_zero)
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   914
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   915
lemma setsum_cl_ivl_Suc2:
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   916
  "(\<Sum>i=m..Suc n. f i) = (if Suc n < m then 0 else f m + (\<Sum>i=m..n. f (Suc i)))"
28069
ba4de3022862 moved lemma into SetInterval where it belongs
nipkow
parents: 27483
diff changeset
   917
by (simp add: setsum_head_Suc setsum_shift_bounds_cl_Suc_ivl
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   918
         del: setsum_cl_ivl_Suc)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   919
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   920
lemma exp_series_add:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   921
  fixes x y :: "'a::{real_field}"
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   922
  defines S_def: "S \<equiv> \<lambda>x n. x ^ n /\<^sub>R real (fact n)"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   923
  shows "S (x + y) n = (\<Sum>i=0..n. S x i * S y (n - i))"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   924
proof (induct n)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   925
  case 0
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   926
  show ?case
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   927
    unfolding S_def by simp
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   928
next
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   929
  case (Suc n)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   930
  have S_Suc: "\<And>x n. S x (Suc n) = (x * S x n) /\<^sub>R real (Suc n)"
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
   931
    unfolding S_def by (simp del: mult_Suc)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   932
  hence times_S: "\<And>x n. x * S x n = real (Suc n) *\<^sub>R S x (Suc n)"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   933
    by simp
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   934
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   935
  have "real (Suc n) *\<^sub>R S (x + y) (Suc n) = (x + y) * S (x + y) n"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   936
    by (simp only: times_S)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   937
  also have "\<dots> = (x + y) * (\<Sum>i=0..n. S x i * S y (n-i))"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   938
    by (simp only: Suc)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   939
  also have "\<dots> = x * (\<Sum>i=0..n. S x i * S y (n-i))
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   940
                + y * (\<Sum>i=0..n. S x i * S y (n-i))"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   941
    by (rule left_distrib)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   942
  also have "\<dots> = (\<Sum>i=0..n. (x * S x i) * S y (n-i))
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   943
                + (\<Sum>i=0..n. S x i * (y * S y (n-i)))"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   944
    by (simp only: setsum_right_distrib mult_ac)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   945
  also have "\<dots> = (\<Sum>i=0..n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n-i)))
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   946
                + (\<Sum>i=0..n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i)))"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   947
    by (simp add: times_S Suc_diff_le)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   948
  also have "(\<Sum>i=0..n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n-i))) =
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   949
             (\<Sum>i=0..Suc n. real i *\<^sub>R (S x i * S y (Suc n-i)))"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   950
    by (subst setsum_cl_ivl_Suc2, simp)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   951
  also have "(\<Sum>i=0..n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) =
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   952
             (\<Sum>i=0..Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i)))"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   953
    by (subst setsum_cl_ivl_Suc, simp)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   954
  also have "(\<Sum>i=0..Suc n. real i *\<^sub>R (S x i * S y (Suc n-i))) +
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   955
             (\<Sum>i=0..Suc n. real (Suc n-i) *\<^sub>R (S x i * S y (Suc n-i))) =
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   956
             (\<Sum>i=0..Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n-i)))"
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   957
    by (simp only: setsum_addf [symmetric] scaleR_left_distrib [symmetric]
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   958
              real_of_nat_add [symmetric], simp)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   959
  also have "\<dots> = real (Suc n) *\<^sub>R (\<Sum>i=0..Suc n. S x i * S y (Suc n-i))"
23127
56ee8105c002 simplify names of locale interpretations
huffman
parents: 23115
diff changeset
   960
    by (simp only: scaleR_right.setsum)
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   961
  finally show
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   962
    "S (x + y) (Suc n) = (\<Sum>i=0..Suc n. S x i * S y (Suc n - i))"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   963
    by (simp add: scaleR_cancel_left del: setsum_cl_ivl_Suc)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   964
qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   965
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   966
lemma exp_add: "exp (x + y) = exp x * exp y"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   967
unfolding exp_def
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   968
by (simp only: Cauchy_product summable_norm_exp exp_series_add)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   969
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   970
lemma mult_exp_exp: "exp x * exp y = exp (x + y)"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   971
by (rule exp_add [symmetric])
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   972
23241
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   973
lemma exp_of_real: "exp (of_real x) = of_real (exp x)"
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   974
unfolding exp_def
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   975
apply (subst of_real.suminf)
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   976
apply (rule summable_exp_generic)
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   977
apply (simp add: scaleR_conv_of_real)
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   978
done
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   979
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   980
lemma exp_not_eq_zero [simp]: "exp x \<noteq> 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   981
proof
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   982
  have "exp x * exp (- x) = 1" by (simp add: mult_exp_exp)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   983
  also assume "exp x = 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   984
  finally show "False" by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   985
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   986
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   987
lemma exp_minus: "exp (- x) = inverse (exp x)"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   988
by (rule inverse_unique [symmetric], simp add: mult_exp_exp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   989
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   990
lemma exp_diff: "exp (x - y) = exp x / exp y"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   991
  unfolding diff_minus divide_inverse
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   992
  by (simp add: exp_add exp_minus)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   993
29167
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   994
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   995
subsubsection {* Properties of the Exponential Function on Reals *}
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   996
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   997
text {* Comparisons of @{term "exp x"} with zero. *}
29167
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   998
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   999
text{*Proof: because every exponential can be seen as a square.*}
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
  1000
lemma exp_ge_zero [simp]: "0 \<le> exp (x::real)"
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
  1001
proof -
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
  1002
  have "0 \<le> exp (x/2) * exp (x/2)" by simp
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
  1003
  thus ?thesis by (simp add: exp_add [symmetric])
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
  1004
qed
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
  1005
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1006
lemma exp_gt_zero [simp]: "0 < exp (x::real)"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1007
by (simp add: order_less_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1008
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1009
lemma not_exp_less_zero [simp]: "\<not> exp (x::real) < 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1010
by (simp add: not_less)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1011
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1012
lemma not_exp_le_zero [simp]: "\<not> exp (x::real) \<le> 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1013
by (simp add: not_le)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1014
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1015
lemma abs_exp_cancel [simp]: "\<bar>exp x::real\<bar> = exp x"
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1016
by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1017
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1018
lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15241
diff changeset
  1019
apply (induct "n")
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1020
apply (auto simp add: real_of_nat_Suc right_distrib exp_add mult_commute)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1021
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1022
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1023
text {* Strict monotonicity of exponential. *}
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1024
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1025
lemma exp_ge_add_one_self_aux: "0 \<le> (x::real) ==> (1 + x) \<le> exp(x)"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1026
apply (drule order_le_imp_less_or_eq, auto)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1027
apply (simp add: exp_def)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1028
apply (rule real_le_trans)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1029
apply (rule_tac [2] n = 2 and f = "(%n. inverse (real (fact n)) * x ^ n)" in series_pos_le)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1030
apply (auto intro: summable_exp simp add: numeral_2_eq_2 zero_le_mult_iff)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1031
done
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1032
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1033
lemma exp_gt_one: "0 < (x::real) \<Longrightarrow> 1 < exp x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1034
proof -
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1035
  assume x: "0 < x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1036
  hence "1 < 1 + x" by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1037
  also from x have "1 + x \<le> exp x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1038
    by (simp add: exp_ge_add_one_self_aux)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1039
  finally show ?thesis .
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1040
qed
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1041
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1042
lemma exp_less_mono:
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1043
  fixes x y :: real
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1044
  assumes "x < y" shows "exp x < exp y"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1045
proof -
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1046
  from `x < y` have "0 < y - x" by simp
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1047
  hence "1 < exp (y - x)" by (rule exp_gt_one)
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1048
  hence "1 < exp y / exp x" by (simp only: exp_diff)
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1049
  thus "exp x < exp y" by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1050
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1051
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1052
lemma exp_less_cancel: "exp (x::real) < exp y ==> x < y"
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1053
apply (simp add: linorder_not_le [symmetric])
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1054
apply (auto simp add: order_le_less exp_less_mono)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1055
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1056
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1057
lemma exp_less_cancel_iff [iff]: "exp (x::real) < exp y \<longleftrightarrow> x < y"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1058
by (auto intro: exp_less_mono exp_less_cancel)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1059
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1060
lemma exp_le_cancel_iff [iff]: "exp (x::real) \<le> exp y \<longleftrightarrow> x \<le> y"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1061
by (auto simp add: linorder_not_less [symmetric])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1062
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1063
lemma exp_inj_iff [iff]: "exp (x::real) = exp y \<longleftrightarrow> x = y"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1064
by (simp add: order_eq_iff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1065
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1066
text {* Comparisons of @{term "exp x"} with one. *}
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1067
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1068
lemma one_less_exp_iff [simp]: "1 < exp (x::real) \<longleftrightarrow> 0 < x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1069
  using exp_less_cancel_iff [where x=0 and y=x] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1070
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1071
lemma exp_less_one_iff [simp]: "exp (x::real) < 1 \<longleftrightarrow> x < 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1072
  using exp_less_cancel_iff [where x=x and y=0] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1073
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1074
lemma one_le_exp_iff [simp]: "1 \<le> exp (x::real) \<longleftrightarrow> 0 \<le> x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1075
  using exp_le_cancel_iff [where x=0 and y=x] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1076
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1077
lemma exp_le_one_iff [simp]: "exp (x::real) \<le> 1 \<longleftrightarrow> x \<le> 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1078
  using exp_le_cancel_iff [where x=x and y=0] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1079
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1080
lemma exp_eq_one_iff [simp]: "exp (x::real) = 1 \<longleftrightarrow> x = 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1081
  using exp_inj_iff [where x=x and y=0] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1082
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1083
lemma lemma_exp_total: "1 \<le> y ==> \<exists>x. 0 \<le> x & x \<le> y - 1 & exp(x::real) = y"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1084
apply (rule IVT)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1085
apply (auto intro: isCont_exp simp add: le_diff_eq)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1086
apply (subgoal_tac "1 + (y - 1) \<le> exp (y - 1)") 
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1087
apply simp
17014
ad5ceb90877d renamed exp_ge_add_one_self to exp_ge_add_one_self_aux
avigad
parents: 16924
diff changeset
  1088
apply (rule exp_ge_add_one_self_aux, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1089
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1090
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1091
lemma exp_total: "0 < (y::real) ==> \<exists>x. exp x = y"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1092
apply (rule_tac x = 1 and y = y in linorder_cases)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1093
apply (drule order_less_imp_le [THEN lemma_exp_total])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1094
apply (rule_tac [2] x = 0 in exI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1095
apply (frule_tac [3] real_inverse_gt_one)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1096
apply (drule_tac [4] order_less_imp_le [THEN lemma_exp_total], auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1097
apply (rule_tac x = "-x" in exI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1098
apply (simp add: exp_minus)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1099
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1100
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1101
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1102
subsection {* Natural Logarithm *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1103
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1104
definition
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1105
  ln :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1106
  "ln x = (THE u. exp u = x)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1107
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1108
lemma ln_exp [simp]: "ln (exp x) = x"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1109
by (simp add: ln_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1110
22654
c2b6b5a9e136 new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents: 22653
diff changeset
  1111
lemma exp_ln [simp]: "0 < x \<Longrightarrow> exp (ln x) = x"
c2b6b5a9e136 new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents: 22653
diff changeset
  1112
by (auto dest: exp_total)
c2b6b5a9e136 new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents: 22653
diff changeset
  1113
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1114
lemma exp_ln_iff [simp]: "exp (ln x) = x \<longleftrightarrow> 0 < x"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1115
apply (rule iffI)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1116
apply (erule subst, rule exp_gt_zero)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1117
apply (erule exp_ln)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1118
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1119
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1120
lemma ln_unique: "exp y = x \<Longrightarrow> ln x = y"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1121
by (erule subst, rule ln_exp)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1122
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1123
lemma ln_one [simp]: "ln 1 = 0"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1124
by (rule ln_unique, simp)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1125
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1126
lemma ln_mult: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln (x * y) = ln x + ln y"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1127
by (rule ln_unique, simp add: exp_add)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1128
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1129
lemma ln_inverse: "0 < x \<Longrightarrow> ln (inverse x) = - ln x"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1130
by (rule ln_unique, simp add: exp_minus)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1131
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1132
lemma ln_div: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln (x / y) = ln x - ln y"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1133
by (rule ln_unique, simp add: exp_diff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1134
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1135
lemma ln_realpow: "0 < x \<Longrightarrow> ln (x ^ n) = real n * ln x"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1136
by (rule ln_unique, simp add: exp_real_of_nat_mult)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1137
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1138
lemma ln_less_cancel_iff [simp]: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln x < ln y \<longleftrightarrow> x < y"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1139
by (subst exp_less_cancel_iff [symmetric], simp)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1140
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1141
lemma ln_le_cancel_iff [simp]: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln x \<le> ln y \<longleftrightarrow> x \<le> y"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1142
by (simp add: linorder_not_less [symmetric])
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1143
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1144
lemma ln_inj_iff [simp]: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln x = ln y \<longleftrightarrow> x = y"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1145
by (simp add: order_eq_iff)
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1146
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1147
lemma ln_add_one_self_le_self [simp]: "0 \<le> x \<Longrightarrow> ln (1 + x) \<le> x"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1148
apply (rule exp_le_cancel_iff [THEN iffD1])
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1149
apply (simp add: exp_ge_add_one_self_aux)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1150
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1151
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1152
lemma ln_less_self [simp]: "0 < x \<Longrightarrow> ln x < x"
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1153
by (rule order_less_le_trans [where y="ln (1 + x)"]) simp_all
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1154
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1155
lemma ln_ge_zero [simp]:
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1156
  assumes x: "1 \<le> x" shows "0 \<le> ln x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1157
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1158
  have "0 < x" using x by arith
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1159
  hence "exp 0 \<le> exp (ln x)"
22915
bb8a928a6bfa fix proofs
huffman
parents: 22722
diff changeset
  1160
    by (simp add: x)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1161
  thus ?thesis by (simp only: exp_le_cancel_iff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1162
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1163
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1164
lemma ln_ge_zero_imp_ge_one:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1165
  assumes ln: "0 \<le> ln x" 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1166
      and x:  "0 < x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1167
  shows "1 \<le> x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1168
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1169
  from ln have "ln 1 \<le> ln x" by simp
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1170
  thus ?thesis by (simp add: x del: ln_one) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1171
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1172
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1173
lemma ln_ge_zero_iff [simp]: "0 < x ==> (0 \<le> ln x) = (1 \<le> x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1174
by (blast intro: ln_ge_zero ln_ge_zero_imp_ge_one)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1175
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1176
lemma ln_less_zero_iff [simp]: "0 < x ==> (ln x < 0) = (x < 1)"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1177
by (insert ln_ge_zero_iff [of x], arith)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1178
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1179
lemma ln_gt_zero:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1180
  assumes x: "1 < x" shows "0 < ln x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1181
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1182
  have "0 < x" using x by arith
22915
bb8a928a6bfa fix proofs
huffman
parents: 22722
diff changeset
  1183
  hence "exp 0 < exp (ln x)" by (simp add: x)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1184
  thus ?thesis  by (simp only: exp_less_cancel_iff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1185
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1186
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1187
lemma ln_gt_zero_imp_gt_one:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1188
  assumes ln: "0 < ln x" 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1189
      and x:  "0 < x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1190
  shows "1 < x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1191
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1192
  from ln have "ln 1 < ln x" by simp
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1193
  thus ?thesis by (simp add: x del: ln_one) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1194
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1195
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1196
lemma ln_gt_zero_iff [simp]: "0 < x ==> (0 < ln x) = (1 < x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1197
by (blast intro: ln_gt_zero ln_gt_zero_imp_gt_one)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1198
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1199
lemma ln_eq_zero_iff [simp]: "0 < x ==> (ln x = 0) = (x = 1)"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1200
by (insert ln_less_zero_iff [of x] ln_gt_zero_iff [of x], arith)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1201
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1202
lemma ln_less_zero: "[| 0 < x; x < 1 |] ==> ln x < 0"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1203
by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1204
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1205
lemma exp_ln_eq: "exp u = x ==> ln x = u"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1206
by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1207
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1208
lemma isCont_ln: "0 < x \<Longrightarrow> isCont ln x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1209
apply (subgoal_tac "isCont ln (exp (ln x))", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1210
apply (rule isCont_inverse_function [where f=exp], simp_all)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1211
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1212
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1213
lemma DERIV_ln: "0 < x \<Longrightarrow> DERIV ln x :> inverse x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1214
apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1215
apply (erule lemma_DERIV_subst [OF DERIV_exp exp_ln])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1216
apply (simp_all add: abs_if isCont_ln)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1217
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1218
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1219
lemma ln_series: assumes "0 < x" and "x < 2"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1220
  shows "ln x = (\<Sum> n. (-1)^n * (1 / real (n + 1)) * (x - 1)^(Suc n))" (is "ln x = suminf (?f (x - 1))")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1221
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1222
  let "?f' x n" = "(-1)^n * (x - 1)^n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1223
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1224
  have "ln x - suminf (?f (x - 1)) = ln 1 - suminf (?f (1 - 1))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1225
  proof (rule DERIV_isconst3[where x=x])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1226
    fix x :: real assume "x \<in> {0 <..< 2}" hence "0 < x" and "x < 2" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1227
    have "norm (1 - x) < 1" using `0 < x` and `x < 2` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1228
    have "1 / x = 1 / (1 - (1 - x))" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1229
    also have "\<dots> = (\<Sum> n. (1 - x)^n)" using geometric_sums[OF `norm (1 - x) < 1`] by (rule sums_unique)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1230
    also have "\<dots> = suminf (?f' x)" unfolding power_mult_distrib[symmetric] by (rule arg_cong[where f=suminf], rule arg_cong[where f="op ^"], auto)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1231
    finally have "DERIV ln x :> suminf (?f' x)" using DERIV_ln[OF `0 < x`] unfolding real_divide_def by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1232
    moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1233
    have repos: "\<And> h x :: real. h - 1 + x = h + x - 1" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1234
    have "DERIV (\<lambda>x. suminf (?f x)) (x - 1) :> (\<Sum>n. (-1)^n * (1 / real (n + 1)) * real (Suc n) * (x - 1) ^ n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1235
    proof (rule DERIV_power_series')
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1236
      show "x - 1 \<in> {- 1<..<1}" and "(0 :: real) < 1" using `0 < x` `x < 2` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1237
      { fix x :: real assume "x \<in> {- 1<..<1}" hence "norm (-x) < 1" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1238
	show "summable (\<lambda>n. -1 ^ n * (1 / real (n + 1)) * real (Suc n) * x ^ n)"
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  1239
          unfolding One_nat_def
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1240
	  by (auto simp del: power_mult_distrib simp add: power_mult_distrib[symmetric] summable_geometric[OF `norm (-x) < 1`])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1241
      }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1242
    qed
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  1243
    hence "DERIV (\<lambda>x. suminf (?f x)) (x - 1) :> suminf (?f' x)" unfolding One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1244
    hence "DERIV (\<lambda>x. suminf (?f (x - 1))) x :> suminf (?f' x)" unfolding DERIV_iff repos .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1245
    ultimately have "DERIV (\<lambda>x. ln x - suminf (?f (x - 1))) x :> (suminf (?f' x) - suminf (?f' x))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1246
      by (rule DERIV_diff)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1247
    thus "DERIV (\<lambda>x. ln x - suminf (?f (x - 1))) x :> 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1248
  qed (auto simp add: assms)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1249
  thus ?thesis by (auto simp add: suminf_zero)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1250
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1251
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1252
subsection {* Sine and Cosine *}
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1253
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1254
definition
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1255
  sin :: "real => real" where
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1256
  "sin x = (\<Sum>n. (if even(n) then 0 else
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1257
             (-1 ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1258
 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1259
definition
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1260
  cos :: "real => real" where
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1261
  "cos x = (\<Sum>n. (if even(n) then (-1 ^ (n div 2))/(real (fact n)) 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1262
                            else 0) * x ^ n)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1263
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1264
lemma summable_sin: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1265
     "summable (%n.  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1266
           (if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1267
           else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) *  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1268
                x ^ n)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1269
apply (rule_tac g = "(%n. inverse (real (fact n)) * \<bar>x\<bar> ^ n)" in summable_comparison_test)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1270
apply (rule_tac [2] summable_exp)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1271
apply (rule_tac x = 0 in exI)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1272
apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1273
done
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1274
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1275
lemma summable_cos: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1276
      "summable (%n.  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1277
           (if even n then  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1278
           -1 ^ (n div 2)/(real (fact n)) else 0) * x ^ n)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1279
apply (rule_tac g = "(%n. inverse (real (fact n)) * \<bar>x\<bar> ^ n)" in summable_comparison_test)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1280
apply (rule_tac [2] summable_exp)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1281
apply (rule_tac x = 0 in exI)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1282
apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1283
done
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1284
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1285
lemma lemma_STAR_sin:
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1286
     "(if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1287
       else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1288
by (induct "n", auto)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1289
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1290
lemma lemma_STAR_cos:
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1291
     "0 < n -->  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1292
      -1 ^ (n div 2)/(real (fact n)) * 0 ^ n = 0"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1293
by (induct "n", auto)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1294
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1295
lemma lemma_STAR_cos1:
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1296
     "0 < n -->  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1297
      (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1298
by (induct "n", auto)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1299
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1300
lemma lemma_STAR_cos2:
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1301
  "(\<Sum>n=1..<n. if even n then -1 ^ (n div 2)/(real (fact n)) *  0 ^ n 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1302
                         else 0) = 0"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1303
apply (induct "n")
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1304
apply (case_tac [2] "n", auto)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1305
done
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1306
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1307
lemma sin_converges: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1308
      "(%n. (if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1309
            else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) *  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1310
                 x ^ n) sums sin(x)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1311
unfolding sin_def by (rule summable_sin [THEN summable_sums])
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1312
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1313
lemma cos_converges: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1314
      "(%n. (if even n then  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1315
           -1 ^ (n div 2)/(real (fact n))  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1316
           else 0) * x ^ n) sums cos(x)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1317
unfolding cos_def by (rule summable_cos [THEN summable_sums])
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1318
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1319
lemma sin_fdiffs: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1320
      "diffs(%n. if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1321
           else -1 ^ ((n - Suc 0) div 2)/(real (fact n)))  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1322
       = (%n. if even n then  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1323
                 -1 ^ (n div 2)/(real (fact n))  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1324
              else 0)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1325
by (auto intro!: ext 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1326
         simp add: diffs_def divide_inverse real_of_nat_def of_nat_mult
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1327
         simp del: mult_Suc of_nat_Suc)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1328
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1329
lemma sin_fdiffs2: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1330
       "diffs(%n. if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1331
           else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) n  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1332
       = (if even n then  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1333
                 -1 ^ (n div 2)/(real (fact n))  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1334
              else 0)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1335
by (simp only: sin_fdiffs)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1336
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1337
lemma cos_fdiffs: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1338
      "diffs(%n. if even n then  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1339
                 -1 ^ (n div 2)/(real (fact n)) else 0)  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1340
       = (%n. - (if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1341
           else -1 ^ ((n - Suc 0)div 2)/(real (fact n))))"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1342
by (auto intro!: ext 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1343
         simp add: diffs_def divide_inverse odd_Suc_mult_two_ex real_of_nat_def of_nat_mult
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1344
         simp del: mult_Suc of_nat_Suc)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1345
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1346
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1347
lemma cos_fdiffs2: 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1348
      "diffs(%n. if even n then  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1349
                 -1 ^ (n div 2)/(real (fact n)) else 0) n 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1350
       = - (if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1351
           else -1 ^ ((n - Suc 0)div 2)/(real (fact n)))"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1352
by (simp only: cos_fdiffs)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1353
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1354
text{*Now at last we can get the derivatives of exp, sin and cos*}
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1355
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1356
lemma lemma_sin_minus:
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1357
     "- sin x = (\<Sum>n. - ((if even n then 0 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1358
                  else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1359
by (auto intro!: sums_unique sums_minus sin_converges)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1360
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1361
lemma lemma_sin_ext:
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1362
     "sin = (%x. \<Sum>n. 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1363
                   (if even n then 0  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1364
                       else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) *  
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1365
                   x ^ n)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1366
by (auto intro!: ext simp add: sin_def)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1367
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1368
lemma lemma_cos_ext:
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1369
     "cos = (%x. \<Sum>n. 
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1370
                   (if even n then -1 ^ (n div 2)/(real (fact n)) else 0) *
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1371
                   x ^ n)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1372
by (auto intro!: ext simp add: cos_def)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1373
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1374
lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1375
apply (simp add: cos_def)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1376
apply (subst lemma_sin_ext)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1377
apply (auto simp add: sin_fdiffs2 [symmetric])
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1378
apply (rule_tac K = "1 + \<bar>x\<bar>" in termdiffs)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1379
apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1380
done
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1381
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1382
lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1383
apply (subst lemma_cos_ext)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1384
apply (auto simp add: lemma_sin_minus cos_fdiffs2 [symmetric] minus_mult_left)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1385
apply (rule_tac K = "1 + \<bar>x\<bar>" in termdiffs)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1386
apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs diffs_minus)
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1387
done
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1388
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1389
lemma isCont_sin [simp]: "isCont sin x"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1390
by (rule DERIV_sin [THEN DERIV_isCont])
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1391
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1392
lemma isCont_cos [simp]: "isCont cos x"
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1393
by (rule DERIV_cos [THEN DERIV_isCont])
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1394
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1395
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1396
subsection {* Properties of Sine and Cosine *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1397
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1398
lemma sin_zero [simp]: "sin 0 = 0"
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
  1399
unfolding sin_def by (simp add: powser_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1400
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1401
lemma cos_zero [simp]: "cos 0 = 1"
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
  1402
unfolding cos_def by (simp add: powser_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1403
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1404
lemma DERIV_sin_sin_mult [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1405
     "DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1406
by (rule DERIV_mult, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1407
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1408
lemma DERIV_sin_sin_mult2 [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1409
     "DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1410
apply (cut_tac x = x in DERIV_sin_sin_mult)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1411
apply (auto simp add: mult_assoc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1412
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1413
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1414
lemma DERIV_sin_realpow2 [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1415
     "DERIV (%x. (sin x)\<twosuperior>) x :> cos(x) * sin(x) + cos(x) * sin(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1416
by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1417
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1418
lemma DERIV_sin_realpow2a [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1419
     "DERIV (%x. (sin x)\<twosuperior>) x :> 2 * cos(x) * sin(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1420
by (auto simp add: numeral_2_eq_2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1421
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1422
lemma DERIV_cos_cos_mult [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1423
     "DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1424
by (rule DERIV_mult, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1425
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1426
lemma DERIV_cos_cos_mult2 [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1427
     "DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1428
apply (cut_tac x = x in DERIV_cos_cos_mult)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1429
apply (auto simp add: mult_ac)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1430
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1431
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1432
lemma DERIV_cos_realpow2 [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1433
     "DERIV (%x. (cos x)\<twosuperior>) x :> -sin(x) * cos(x) + -sin(x) * cos(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1434
by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1435
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1436
lemma DERIV_cos_realpow2a [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1437
     "DERIV (%x. (cos x)\<twosuperior>) x :> -2 * cos(x) * sin(x)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1438
by (auto simp add: numeral_2_eq_2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1439
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1440
lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1441
by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1442
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1443
lemma DERIV_cos_realpow2b: "DERIV (%x. (cos x)\<twosuperior>) x :> -(2 * cos(x) * sin(x))"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1444
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1445
apply (rule DERIV_cos_realpow2a, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1446
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1447
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1448
(* most useful *)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1449
lemma DERIV_cos_cos_mult3 [simp]:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1450
     "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1451
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1452
apply (rule DERIV_cos_cos_mult2, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1453
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1454
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1455
lemma DERIV_sin_circle_all: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1456
     "\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :>  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1457
             (2*cos(x)*sin(x) - 2*cos(x)*sin(x))"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1458
apply (simp only: diff_minus, safe)
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1459
apply (rule DERIV_add) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1460
apply (auto simp add: numeral_2_eq_2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1461
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1462
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1463
lemma DERIV_sin_circle_all_zero [simp]:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1464
     "\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :> 0"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1465
by (cut_tac DERIV_sin_circle_all, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1466
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1467
lemma sin_cos_squared_add [simp]: "((sin x)\<twosuperior>) + ((cos x)\<twosuperior>) = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1468
apply (cut_tac x = x and y = 0 in DERIV_sin_circle_all_zero [THEN DERIV_isconst_all])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1469
apply (auto simp add: numeral_2_eq_2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1470
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1471
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1472
lemma sin_cos_squared_add2 [simp]: "((cos x)\<twosuperior>) + ((sin x)\<twosuperior>) = 1"
23286
huffman
parents: 23278
diff changeset
  1473
apply (subst add_commute)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1474
apply (rule sin_cos_squared_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1475
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1476
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1477
lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1478
apply (cut_tac x = x in sin_cos_squared_add2)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1479
apply (simp add: power2_eq_square)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1480
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1481
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1482
lemma sin_squared_eq: "(sin x)\<twosuperior> = 1 - (cos x)\<twosuperior>"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1483
apply (rule_tac a1 = "(cos x)\<twosuperior>" in add_right_cancel [THEN iffD1])
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1484
apply simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1485
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1486
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1487
lemma cos_squared_eq: "(cos x)\<twosuperior> = 1 - (sin x)\<twosuperior>"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1488
apply (rule_tac a1 = "(sin x)\<twosuperior>" in add_right_cancel [THEN iffD1])
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1489
apply simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1490
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1491
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
  1492
lemma abs_sin_le_one [simp]: "\<bar>sin x\<bar> \<le> 1"
23097
f4779adcd1a2 simplify some proofs
huffman
parents: 23082
diff changeset
  1493
by (rule power2_le_imp_le, simp_all add: sin_squared_eq)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1494
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1495
lemma sin_ge_minus_one [simp]: "-1 \<le> sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1496
apply (insert abs_sin_le_one [of x]) 
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  1497
apply (simp add: abs_le_iff del: abs_sin_le_one) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1498
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1499
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1500
lemma sin_le_one [simp]: "sin x \<le> 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1501
apply (insert abs_sin_le_one [of x]) 
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  1502
apply (simp add: abs_le_iff del: abs_sin_le_one) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1503
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1504
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
  1505
lemma abs_cos_le_one [simp]: "\<bar>cos x\<bar> \<le> 1"
23097
f4779adcd1a2 simplify some proofs
huffman
parents: 23082
diff changeset
  1506
by (rule power2_le_imp_le, simp_all add: cos_squared_eq)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1507
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1508
lemma cos_ge_minus_one [simp]: "-1 \<le> cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1509
apply (insert abs_cos_le_one [of x]) 
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  1510
apply (simp add: abs_le_iff del: abs_cos_le_one) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1511
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1512
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1513
lemma cos_le_one [simp]: "cos x \<le> 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1514
apply (insert abs_cos_le_one [of x]) 
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  1515
apply (simp add: abs_le_iff del: abs_cos_le_one)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1516
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1517
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1518
lemma DERIV_fun_pow: "DERIV g x :> m ==>  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1519
      DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m"
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  1520
unfolding One_nat_def
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1521
apply (rule lemma_DERIV_subst)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1522
apply (rule_tac f = "(%x. x ^ n)" in DERIV_chain2)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1523
apply (rule DERIV_pow, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1524
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1525
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1526
lemma DERIV_fun_exp:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1527
     "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1528
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1529
apply (rule_tac f = exp in DERIV_chain2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1530
apply (rule DERIV_exp, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1531
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1532
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1533
lemma DERIV_fun_sin:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1534
     "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1535
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1536
apply (rule_tac f = sin in DERIV_chain2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1537
apply (rule DERIV_sin, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1538
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1539
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1540
lemma DERIV_fun_cos:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1541
     "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1542
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1543
apply (rule_tac f = cos in DERIV_chain2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1544
apply (rule DERIV_cos, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1545
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1546
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 23066
diff changeset
  1547
lemmas DERIV_intros = DERIV_ident DERIV_const DERIV_cos DERIV_cmult 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1548
                    DERIV_sin  DERIV_exp  DERIV_inverse DERIV_pow 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1549
                    DERIV_add  DERIV_diff  DERIV_mult  DERIV_minus 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1550
                    DERIV_inverse_fun DERIV_quotient DERIV_fun_pow 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1551
                    DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1552
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1553
(* lemma *)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1554
lemma lemma_DERIV_sin_cos_add:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1555
     "\<forall>x.  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1556
         DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1557
               (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1558
apply (safe, rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1559
apply (best intro!: DERIV_intros intro: DERIV_chain2) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1560
  --{*replaces the old @{text DERIV_tac}*}
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  1561
apply (auto simp add: algebra_simps)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1562
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1563
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1564
lemma sin_cos_add [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1565
     "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1566
      (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1567
apply (cut_tac y = 0 and x = x and y7 = y 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1568
       in lemma_DERIV_sin_cos_add [THEN DERIV_isconst_all])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1569
apply (auto simp add: numeral_2_eq_2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1570
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1571
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1572
lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1573
apply (cut_tac x = x and y = y in sin_cos_add)
22969
bf523cac9a05 tuned proofs
huffman
parents: 22960
diff changeset
  1574
apply (simp del: sin_cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1575
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1576
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1577
lemma cos_add: "cos (x + y) = cos x * cos y - sin x * sin y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1578
apply (cut_tac x = x and y = y in sin_cos_add)
22969
bf523cac9a05 tuned proofs
huffman
parents: 22960
diff changeset
  1579
apply (simp del: sin_cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1580
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1581
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1582
lemma lemma_DERIV_sin_cos_minus:
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1583
    "\<forall>x. DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1584
apply (safe, rule lemma_DERIV_subst)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  1585
apply (best intro!: DERIV_intros intro: DERIV_chain2)
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  1586
apply (simp add: algebra_simps)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1587
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1588
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1589
lemma sin_cos_minus: 
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1590
    "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0"
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1591
apply (cut_tac y = 0 and x = x 
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1592
       in lemma_DERIV_sin_cos_minus [THEN DERIV_isconst_all])
22969
bf523cac9a05 tuned proofs
huffman
parents: 22960
diff changeset
  1593
apply simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1594
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1595
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1596
lemma sin_minus [simp]: "sin (-x) = -sin(x)"
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1597
  using sin_cos_minus [where x=x] by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1598
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1599
lemma cos_minus [simp]: "cos (-x) = cos(x)"
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1600
  using sin_cos_minus [where x=x] by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1601
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1602
lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y"
22969
bf523cac9a05 tuned proofs
huffman
parents: 22960
diff changeset
  1603
by (simp add: diff_minus sin_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1604
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1605
lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1606
by (simp add: sin_diff mult_commute)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1607
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1608
lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y"
22969
bf523cac9a05 tuned proofs
huffman
parents: 22960
diff changeset
  1609
by (simp add: diff_minus cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1610
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1611
lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1612
by (simp add: cos_diff mult_commute)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1613
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1614
lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x"
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1615
  using sin_add [where x=x and y=x] by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1616
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1617
lemma cos_double: "cos(2* x) = ((cos x)\<twosuperior>) - ((sin x)\<twosuperior>)"
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1618
  using cos_add [where x=x and y=x]
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1619
  by (simp add: power2_eq_square)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1620
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1621
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1622
subsection {* The Constant Pi *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1623
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1624
definition
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1625
  pi :: "real" where
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1626
  "pi = 2 * (THE x. 0 \<le> (x::real) & x \<le> 2 & cos x = 0)"
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1627
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1628
text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"}; 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1629
   hence define pi.*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1630
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1631
lemma sin_paired:
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  1632
     "(%n. -1 ^ n /(real (fact (2 * n + 1))) * x ^ (2 * n + 1)) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1633
      sums  sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1634
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1635
  have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2.
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1636
            (if even k then 0
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  1637
             else -1 ^ ((k - Suc 0) div 2) / real (fact k)) *
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1638
            x ^ k) 
23176
40a760921d94 simplify some proofs
huffman
parents: 23127
diff changeset
  1639
	sums sin x"
40a760921d94 simplify some proofs
huffman
parents: 23127
diff changeset
  1640
    unfolding sin_def
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1641
    by (rule sin_converges [THEN sums_summable, THEN sums_group], simp) 
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  1642
  thus ?thesis unfolding One_nat_def by (simp add: mult_ac)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1643
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1644
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1645
text {* FIXME: This is a long, ugly proof! *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1646
lemma sin_gt_zero: "[|0 < x; x < 2 |] ==> 0 < sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1647
apply (subgoal_tac 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1648
       "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2.
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  1649
              -1 ^ k / real (fact (2 * k + 1)) * x ^ (2 * k + 1)) 
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  1650
     sums (\<Sum>n. -1 ^ n / real (fact (2 * n + 1)) * x ^ (2 * n + 1))")
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1651
 prefer 2
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1652
 apply (rule sin_paired [THEN sums_summable, THEN sums_group], simp) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1653
apply (rotate_tac 2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1654
apply (drule sin_paired [THEN sums_unique, THEN ssubst])
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  1655
unfolding One_nat_def
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1656
apply (auto simp del: fact_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1657
apply (frule sums_unique)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1658
apply (auto simp del: fact_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1659
apply (rule_tac n1 = 0 in series_pos_less [THEN [2] order_le_less_trans])
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1660
apply (auto simp del: fact_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1661
apply (erule sums_summable)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1662
apply (case_tac "m=0")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1663
apply (simp (no_asm_simp))
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1664
apply (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x") 
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  1665
apply (simp only: mult_less_cancel_left, simp)  
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  1666
apply (simp (no_asm_simp) add: numeral_2_eq_2 [symmetric] mult_assoc [symmetric])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1667
apply (subgoal_tac "x*x < 2*3", simp) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1668
apply (rule mult_strict_mono)
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1669
apply (auto simp add: real_0_less_add_iff real_of_nat_Suc simp del: fact_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1670
apply (subst fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1671
apply (subst fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1672
apply (subst fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1673
apply (subst fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1674
apply (subst real_of_nat_mult)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1675
apply (subst real_of_nat_mult)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1676
apply (subst real_of_nat_mult)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1677
apply (subst real_of_nat_mult)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  1678
apply (simp (no_asm) add: divide_inverse del: fact_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1679
apply (auto simp add: mult_assoc [symmetric] simp del: fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1680
apply (rule_tac c="real (Suc (Suc (4*m)))" in mult_less_imp_less_right) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1681
apply (auto simp add: mult_assoc simp del: fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1682
apply (rule_tac c="real (Suc (Suc (Suc (4*m))))" in mult_less_imp_less_right) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1683
apply (auto simp add: mult_assoc mult_less_cancel_left simp del: fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1684
apply (subgoal_tac "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * x)") 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1685
apply (erule ssubst)+
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1686
apply (auto simp del: fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1687
apply (subgoal_tac "0 < x ^ (4 * m) ")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1688
 prefer 2 apply (simp only: zero_less_power) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1689
apply (simp (no_asm_simp) add: mult_less_cancel_left)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1690
apply (rule mult_strict_mono)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1691
apply (simp_all (no_asm_simp))
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1692
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1693
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1694
lemma sin_gt_zero1: "[|0 < x; x < 2 |] ==> 0 < sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1695
by (auto intro: sin_gt_zero)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1696
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1697
lemma cos_double_less_one: "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1698
apply (cut_tac x = x in sin_gt_zero1)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1699
apply (auto simp add: cos_squared_eq cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1700
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1701
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1702
lemma cos_paired:
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  1703
     "(%n. -1 ^ n /(real (fact (2 * n))) * x ^ (2 * n)) sums cos x"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1704
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1705
  have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2.
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  1706
            (if even k then -1 ^ (k div 2) / real (fact k) else 0) *
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1707
            x ^ k) 
23176
40a760921d94 simplify some proofs
huffman
parents: 23127
diff changeset
  1708
        sums cos x"
40a760921d94 simplify some proofs
huffman
parents: 23127
diff changeset
  1709
    unfolding cos_def
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1710
    by (rule cos_converges [THEN sums_summable, THEN sums_group], simp) 
23176
40a760921d94 simplify some proofs
huffman
parents: 23127
diff changeset
  1711
  thus ?thesis by (simp add: mult_ac)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1712
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1713
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1714
lemma fact_lemma: "real (n::nat) * 4 = real (4 * n)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1715
by simp
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1716
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1717
lemma cos_two_less_zero [simp]: "cos (2) < 0"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1718
apply (cut_tac x = 2 in cos_paired)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1719
apply (drule sums_minus)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1720
apply (rule neg_less_iff_less [THEN iffD1]) 
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  1721
apply (frule sums_unique, auto)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  1722
apply (rule_tac y =
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  1723
 "\<Sum>n=0..< Suc(Suc(Suc 0)). - (-1 ^ n / (real(fact (2*n))) * 2 ^ (2*n))"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15383
diff changeset
  1724
       in order_less_trans)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  1725
apply (simp (no_asm) add: fact_num_eq_if realpow_num_eq_if del: fact_Suc)
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15546
diff changeset
  1726
apply (simp (no_asm) add: mult_assoc del: setsum_op_ivl_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1727
apply (rule sumr_pos_lt_pair)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1728
apply (erule sums_summable, safe)
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  1729
unfolding One_nat_def
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1730
apply (simp (no_asm) add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] 
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  1731
            del: fact_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1732
apply (rule real_mult_inverse_cancel2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1733
apply (rule real_of_nat_fact_gt_zero)+
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1734
apply (simp (no_asm) add: mult_assoc [symmetric] del: fact_Suc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1735
apply (subst fact_lemma) 
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15383
diff changeset
  1736
apply (subst fact_Suc [of "Suc (Suc (Suc (Suc (Suc (Suc (Suc (4 * d)))))))"])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15383
diff changeset
  1737
apply (simp only: real_of_nat_mult)
23007
e025695d9b0e use mult_strict_mono instead of real_mult_less_mono
huffman
parents: 22998
diff changeset
  1738
apply (rule mult_strict_mono, force)
27483
7c58324cd418 use real_of_nat_ge_zero instead of real_of_nat_fact_ge_zero
huffman
parents: 25875
diff changeset
  1739
  apply (rule_tac [3] real_of_nat_ge_zero)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15383
diff changeset
  1740
 prefer 2 apply force
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1741
apply (rule real_of_nat_less_iff [THEN iffD2])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1742
apply (rule fact_less_mono, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1743
done
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1744
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1745
lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq]
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1746
lemmas cos_two_le_zero [simp] = cos_two_less_zero [THEN order_less_imp_le]
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1747
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1748
lemma cos_is_zero: "EX! x. 0 \<le> x & x \<le> 2 & cos x = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1749
apply (subgoal_tac "\<exists>x. 0 \<le> x & x \<le> 2 & cos x = 0")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1750
apply (rule_tac [2] IVT2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1751
apply (auto intro: DERIV_isCont DERIV_cos)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1752
apply (cut_tac x = xa and y = y in linorder_less_linear)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1753
apply (rule ccontr)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1754
apply (subgoal_tac " (\<forall>x. cos differentiable x) & (\<forall>x. isCont cos x) ")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1755
apply (auto intro: DERIV_cos DERIV_isCont simp add: differentiable_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1756
apply (drule_tac f = cos in Rolle)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1757
apply (drule_tac [5] f = cos in Rolle)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1758
apply (auto dest!: DERIV_cos [THEN DERIV_unique] simp add: differentiable_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1759
apply (drule_tac y1 = xa in order_le_less_trans [THEN sin_gt_zero])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1760
apply (assumption, rule_tac y=y in order_less_le_trans, simp_all) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1761
apply (drule_tac y1 = y in order_le_less_trans [THEN sin_gt_zero], assumption, simp_all) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1762
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1763
    
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1764
lemma pi_half: "pi/2 = (THE x. 0 \<le> x & x \<le> 2 & cos x = 0)"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1765
by (simp add: pi_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1766
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1767
lemma cos_pi_half [simp]: "cos (pi / 2) = 0"
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1768
by (simp add: pi_half cos_is_zero [THEN theI'])
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1769
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1770
lemma pi_half_gt_zero [simp]: "0 < pi / 2"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1771
apply (rule order_le_neq_trans)
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1772
apply (simp add: pi_half cos_is_zero [THEN theI'])
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1773
apply (rule notI, drule arg_cong [where f=cos], simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1774
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1775
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1776
lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric]
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1777
lemmas pi_half_ge_zero [simp] = pi_half_gt_zero [THEN order_less_imp_le]
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1778
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1779
lemma pi_half_less_two [simp]: "pi / 2 < 2"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1780
apply (rule order_le_neq_trans)
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1781
apply (simp add: pi_half cos_is_zero [THEN theI'])
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1782
apply (rule notI, drule arg_cong [where f=cos], simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1783
done
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1784
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1785
lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq]
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1786
lemmas pi_half_le_two [simp] =  pi_half_less_two [THEN order_less_imp_le]
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1787
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1788
lemma pi_gt_zero [simp]: "0 < pi"
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1789
by (insert pi_half_gt_zero, simp)
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1790
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1791
lemma pi_ge_zero [simp]: "0 \<le> pi"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1792
by (rule pi_gt_zero [THEN order_less_imp_le])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1793
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1794
lemma pi_neq_zero [simp]: "pi \<noteq> 0"
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  1795
by (rule pi_gt_zero [THEN less_imp_neq, THEN not_sym])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1796
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1797
lemma pi_not_less_zero [simp]: "\<not> pi < 0"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1798
by (simp add: linorder_not_less)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1799
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1800
lemma minus_pi_half_less_zero: "-(pi/2) < 0"
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1801
by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1802
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1803
lemma m2pi_less_pi: "- (2 * pi) < pi"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1804
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1805
  have "- (2 * pi) < 0" and "0 < pi" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1806
  from order_less_trans[OF this] show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1807
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1808
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1809
lemma sin_pi_half [simp]: "sin(pi/2) = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1810
apply (cut_tac x = "pi/2" in sin_cos_squared_add2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1811
apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two])
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1812
apply (simp add: power2_eq_square)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1813
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1814
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1815
lemma cos_pi [simp]: "cos pi = -1"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  1816
by (cut_tac x = "pi/2" and y = "pi/2" in cos_add, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1817
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1818
lemma sin_pi [simp]: "sin pi = 0"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  1819
by (cut_tac x = "pi/2" and y = "pi/2" in sin_add, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1820
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1821
lemma sin_cos_eq: "sin x = cos (pi/2 - x)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1822
by (simp add: diff_minus cos_add)
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1823
declare sin_cos_eq [symmetric, simp]
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1824
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1825
lemma minus_sin_cos_eq: "-sin x = cos (x + pi/2)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1826
by (simp add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1827
declare minus_sin_cos_eq [symmetric, simp]
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1828
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1829
lemma cos_sin_eq: "cos x = sin (pi/2 - x)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1830
by (simp add: diff_minus sin_add)
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  1831
declare cos_sin_eq [symmetric, simp]
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1832
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1833
lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1834
by (simp add: sin_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1835
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1836
lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1837
by (simp add: sin_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1838
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1839
lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1840
by (simp add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1841
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1842
lemma sin_periodic [simp]: "sin (x + 2*pi) = sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1843
by (simp add: sin_add cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1844
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1845
lemma cos_periodic [simp]: "cos (x + 2*pi) = cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1846
by (simp add: cos_add cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1847
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1848
lemma cos_npi [simp]: "cos (real n * pi) = -1 ^ n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15241
diff changeset
  1849
apply (induct "n")
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1850
apply (auto simp add: real_of_nat_Suc left_distrib)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1851
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1852
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1853
lemma cos_npi2 [simp]: "cos (pi * real n) = -1 ^ n"
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1854
proof -
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1855
  have "cos (pi * real n) = cos (real n * pi)" by (simp only: mult_commute)
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1856
  also have "... = -1 ^ n" by (rule cos_npi) 
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1857
  finally show ?thesis .
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1858
qed
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1859
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1860
lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15241
diff changeset
  1861
apply (induct "n")
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1862
apply (auto simp add: real_of_nat_Suc left_distrib)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1863
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1864
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1865
lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0"
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  1866
by (simp add: mult_commute [of pi]) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1867
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1868
lemma cos_two_pi [simp]: "cos (2 * pi) = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1869
by (simp add: cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1870
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1871
lemma sin_two_pi [simp]: "sin (2 * pi) = 0"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1872
by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1873
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1874
lemma sin_gt_zero2: "[| 0 < x; x < pi/2 |] ==> 0 < sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1875
apply (rule sin_gt_zero, assumption)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1876
apply (rule order_less_trans, assumption)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1877
apply (rule pi_half_less_two)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1878
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1879
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1880
lemma sin_less_zero: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1881
  assumes lb: "- pi/2 < x" and "x < 0" shows "sin x < 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1882
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1883
  have "0 < sin (- x)" using prems by (simp only: sin_gt_zero2) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1884
  thus ?thesis by simp
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1885
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1886
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1887
lemma pi_less_4: "pi < 4"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1888
by (cut_tac pi_half_less_two, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1889
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1890
lemma cos_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1891
apply (cut_tac pi_less_4)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1892
apply (cut_tac f = cos and a = 0 and b = x and y = 0 in IVT2_objl, safe, simp_all)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1893
apply (cut_tac cos_is_zero, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1894
apply (rename_tac y z)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1895
apply (drule_tac x = y in spec)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1896
apply (drule_tac x = "pi/2" in spec, simp) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1897
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1898
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1899
lemma cos_gt_zero_pi: "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1900
apply (rule_tac x = x and y = 0 in linorder_cases)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1901
apply (rule cos_minus [THEN subst])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1902
apply (rule cos_gt_zero)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1903
apply (auto intro: cos_gt_zero)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1904
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1905
 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1906
lemma cos_ge_zero: "[| -(pi/2) \<le> x; x \<le> pi/2 |] ==> 0 \<le> cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1907
apply (auto simp add: order_le_less cos_gt_zero_pi)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1908
apply (subgoal_tac "x = pi/2", auto) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1909
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1910
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1911
lemma sin_gt_zero_pi: "[| 0 < x; x < pi  |] ==> 0 < sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1912
apply (subst sin_cos_eq)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1913
apply (rotate_tac 1)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1914
apply (drule real_sum_of_halves [THEN ssubst])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1915
apply (auto intro!: cos_gt_zero_pi simp del: sin_cos_eq [symmetric])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1916
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1917
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1918
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1919
lemma pi_ge_two: "2 \<le> pi"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1920
proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1921
  assume "\<not> 2 \<le> pi" hence "pi < 2" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1922
  have "\<exists>y > pi. y < 2 \<and> y < 2 * pi"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1923
  proof (cases "2 < 2 * pi")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1924
    case True with dense[OF `pi < 2`] show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1925
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1926
    case False have "pi < 2 * pi" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1927
    from dense[OF this] and False show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1928
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1929
  then obtain y where "pi < y" and "y < 2" and "y < 2 * pi" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1930
  hence "0 < sin y" using sin_gt_zero by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1931
  moreover 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1932
  have "sin y < 0" using sin_gt_zero_pi[of "y - pi"] `pi < y` and `y < 2 * pi` sin_periodic_pi[of "y - pi"] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1933
  ultimately show False by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1934
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1935
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1936
lemma sin_ge_zero: "[| 0 \<le> x; x \<le> pi |] ==> 0 \<le> sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1937
by (auto simp add: order_le_less sin_gt_zero_pi)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1938
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1939
lemma cos_total: "[| -1 \<le> y; y \<le> 1 |] ==> EX! x. 0 \<le> x & x \<le> pi & (cos x = y)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1940
apply (subgoal_tac "\<exists>x. 0 \<le> x & x \<le> pi & cos x = y")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1941
apply (rule_tac [2] IVT2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1942
apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1943
apply (cut_tac x = xa and y = y in linorder_less_linear)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1944
apply (rule ccontr, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1945
apply (drule_tac f = cos in Rolle)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1946
apply (drule_tac [5] f = cos in Rolle)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1947
apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1948
            dest!: DERIV_cos [THEN DERIV_unique] 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1949
            simp add: differentiable_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1950
apply (auto dest: sin_gt_zero_pi [OF order_le_less_trans order_less_le_trans])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1951
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1952
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1953
lemma sin_total:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1954
     "[| -1 \<le> y; y \<le> 1 |] ==> EX! x. -(pi/2) \<le> x & x \<le> pi/2 & (sin x = y)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1955
apply (rule ccontr)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1956
apply (subgoal_tac "\<forall>x. (- (pi/2) \<le> x & x \<le> pi/2 & (sin x = y)) = (0 \<le> (x + pi/2) & (x + pi/2) \<le> pi & (cos (x + pi/2) = -y))")
18585
5d379fe2eb74 replaced swap by contrapos_np;
wenzelm
parents: 17318
diff changeset
  1957
apply (erule contrapos_np)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1958
apply (simp del: minus_sin_cos_eq [symmetric])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1959
apply (cut_tac y="-y" in cos_total, simp) apply simp 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1960
apply (erule ex1E)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1961
apply (rule_tac a = "x - (pi/2)" in ex1I)
23286
huffman
parents: 23278
diff changeset
  1962
apply (simp (no_asm) add: add_assoc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1963
apply (rotate_tac 3)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1964
apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1965
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1966
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1967
lemma reals_Archimedean4:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1968
     "[| 0 < y; 0 \<le> x |] ==> \<exists>n. real n * y \<le> x & x < real (Suc n) * y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1969
apply (auto dest!: reals_Archimedean3)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1970
apply (drule_tac x = x in spec, clarify) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1971
apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1972
 prefer 2 apply (erule LeastI) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1973
apply (case_tac "LEAST m::nat. x < real m * y", simp) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1974
apply (subgoal_tac "~ x < real nat * y")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1975
 prefer 2 apply (rule not_less_Least, simp, force)  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1976
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1977
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1978
(* Pre Isabelle99-2 proof was simpler- numerals arithmetic 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1979
   now causes some unwanted re-arrangements of literals!   *)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1980
lemma cos_zero_lemma:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1981
     "[| 0 \<le> x; cos x = 0 |] ==>  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1982
      \<exists>n::nat. ~even n & x = real n * (pi/2)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1983
apply (drule pi_gt_zero [THEN reals_Archimedean4], safe)
15086
e6a2a98d5ef5 removal of more iff-rules from RealDef.thy
paulson
parents: 15085
diff changeset
  1984
apply (subgoal_tac "0 \<le> x - real n * pi & 
e6a2a98d5ef5 removal of more iff-rules from RealDef.thy
paulson
parents: 15085
diff changeset
  1985
                    (x - real n * pi) \<le> pi & (cos (x - real n * pi) = 0) ")
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  1986
apply (auto simp add: algebra_simps real_of_nat_Suc)
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  1987
 prefer 2 apply (simp add: cos_diff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1988
apply (simp add: cos_diff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1989
apply (subgoal_tac "EX! x. 0 \<le> x & x \<le> pi & cos x = 0")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1990
apply (rule_tac [2] cos_total, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1991
apply (drule_tac x = "x - real n * pi" in spec)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1992
apply (drule_tac x = "pi/2" in spec)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1993
apply (simp add: cos_diff)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1994
apply (rule_tac x = "Suc (2 * n)" in exI)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  1995
apply (simp add: real_of_nat_Suc algebra_simps, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1996
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1997
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1998
lemma sin_zero_lemma:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  1999
     "[| 0 \<le> x; sin x = 0 |] ==>  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2000
      \<exists>n::nat. even n & x = real n * (pi/2)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2001
apply (subgoal_tac "\<exists>n::nat. ~ even n & x + pi/2 = real n * (pi/2) ")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2002
 apply (clarify, rule_tac x = "n - 1" in exI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2003
 apply (force simp add: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib)
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  2004
apply (rule cos_zero_lemma)
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  2005
apply (simp_all add: add_increasing)  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2006
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2007
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2008
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2009
lemma cos_zero_iff:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2010
     "(cos x = 0) =  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2011
      ((\<exists>n::nat. ~even n & (x = real n * (pi/2))) |    
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2012
       (\<exists>n::nat. ~even n & (x = -(real n * (pi/2)))))"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2013
apply (rule iffI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2014
apply (cut_tac linorder_linear [of 0 x], safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2015
apply (drule cos_zero_lemma, assumption+)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2016
apply (cut_tac x="-x" in cos_zero_lemma, simp, simp) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2017
apply (force simp add: minus_equation_iff [of x]) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2018
apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) 
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2019
apply (auto simp add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2020
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2021
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2022
(* ditto: but to a lesser extent *)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2023
lemma sin_zero_iff:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2024
     "(sin x = 0) =  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2025
      ((\<exists>n::nat. even n & (x = real n * (pi/2))) |    
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2026
       (\<exists>n::nat. even n & (x = -(real n * (pi/2)))))"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2027
apply (rule iffI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2028
apply (cut_tac linorder_linear [of 0 x], safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2029
apply (drule sin_zero_lemma, assumption+)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2030
apply (cut_tac x="-x" in sin_zero_lemma, simp, simp, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2031
apply (force simp add: minus_equation_iff [of x]) 
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2032
apply (auto simp add: even_mult_two_ex)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2033
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2034
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2035
lemma cos_monotone_0_pi: assumes "0 \<le> y" and "y < x" and "x \<le> pi"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2036
  shows "cos x < cos y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2037
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2038
  have "- (x - y) < 0" by (auto!)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2039
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2040
  from MVT2[OF `y < x` DERIV_cos[THEN impI, THEN allI]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2041
  obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2042
  hence "0 < z" and "z < pi" by (auto!)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2043
  hence "0 < sin z" using sin_gt_zero_pi by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2044
  hence "cos x - cos y < 0" unfolding cos_diff minus_mult_commute[symmetric] using `- (x - y) < 0` by (rule mult_pos_neg2)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2045
  thus ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2046
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2047
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2048
lemma cos_monotone_0_pi': assumes "0 \<le> y" and "y \<le> x" and "x \<le> pi" shows "cos x \<le> cos y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2049
proof (cases "y < x")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2050
  case True show ?thesis using cos_monotone_0_pi[OF `0 \<le> y` True `x \<le> pi`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2051
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2052
  case False hence "y = x" using `y \<le> x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2053
  thus ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2054
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2055
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2056
lemma cos_monotone_minus_pi_0: assumes "-pi \<le> y" and "y < x" and "x \<le> 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2057
  shows "cos y < cos x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2058
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2059
  have "0 \<le> -x" and "-x < -y" and "-y \<le> pi" by (auto!)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2060
  from cos_monotone_0_pi[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2061
  show ?thesis unfolding cos_minus .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2062
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2063
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2064
lemma cos_monotone_minus_pi_0': assumes "-pi \<le> y" and "y \<le> x" and "x \<le> 0" shows "cos y \<le> cos x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2065
proof (cases "y < x")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2066
  case True show ?thesis using cos_monotone_minus_pi_0[OF `-pi \<le> y` True `x \<le> 0`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2067
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2068
  case False hence "y = x" using `y \<le> x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2069
  thus ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2070
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2071
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2072
lemma sin_monotone_2pi': assumes "- (pi / 2) \<le> y" and "y \<le> x" and "x \<le> pi / 2" shows "sin y \<le> sin x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2073
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2074
  have "0 \<le> y + pi / 2" and "y + pi / 2 \<le> x + pi / 2" and "x + pi /2 \<le> pi" using pi_ge_two by (auto!)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2075
  from cos_monotone_0_pi'[OF this] show ?thesis unfolding minus_sin_cos_eq[symmetric] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2076
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2077
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2078
subsection {* Tangent *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2079
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2080
definition
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2081
  tan :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2082
  "tan x = (sin x)/(cos x)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2083
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2084
lemma tan_zero [simp]: "tan 0 = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2085
by (simp add: tan_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2086
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2087
lemma tan_pi [simp]: "tan pi = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2088
by (simp add: tan_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2089
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2090
lemma tan_npi [simp]: "tan (real (n::nat) * pi) = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2091
by (simp add: tan_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2092
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2093
lemma tan_minus [simp]: "tan (-x) = - tan x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2094
by (simp add: tan_def minus_mult_left)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2095
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2096
lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2097
by (simp add: tan_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2098
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2099
lemma lemma_tan_add1: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2100
      "[| cos x \<noteq> 0; cos y \<noteq> 0 |]  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2101
        ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2102
apply (simp add: tan_def divide_inverse)
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2103
apply (auto simp del: inverse_mult_distrib 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2104
            simp add: inverse_mult_distrib [symmetric] mult_ac)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2105
apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst])
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2106
apply (auto simp del: inverse_mult_distrib 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2107
            simp add: mult_assoc left_diff_distrib cos_add)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  2108
done
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2109
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2110
lemma add_tan_eq: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2111
      "[| cos x \<noteq> 0; cos y \<noteq> 0 |]  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2112
       ==> tan x + tan y = sin(x + y)/(cos x * cos y)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2113
apply (simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2114
apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2115
apply (auto simp add: mult_assoc left_distrib)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2116
apply (simp add: sin_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2117
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2118
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2119
lemma tan_add:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2120
     "[| cos x \<noteq> 0; cos y \<noteq> 0; cos (x + y) \<noteq> 0 |]  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2121
      ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2122
apply (simp (no_asm_simp) add: add_tan_eq lemma_tan_add1)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2123
apply (simp add: tan_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2124
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2125
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2126
lemma tan_double:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2127
     "[| cos x \<noteq> 0; cos (2 * x) \<noteq> 0 |]  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2128
      ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2129
apply (insert tan_add [of x x]) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2130
apply (simp add: mult_2 [symmetric])  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2131
apply (auto simp add: numeral_2_eq_2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2132
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2133
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2134
lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2135
by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2136
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2137
lemma tan_less_zero: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2138
  assumes lb: "- pi/2 < x" and "x < 0" shows "tan x < 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2139
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2140
  have "0 < tan (- x)" using prems by (simp only: tan_gt_zero) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2141
  thus ?thesis by simp
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2142
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2143
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2144
lemma tan_half: fixes x :: real assumes "- (pi / 2) < x" and "x < pi / 2"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2145
  shows "tan x = sin (2 * x) / (cos (2 * x) + 1)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2146
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2147
  from cos_gt_zero_pi[OF `- (pi / 2) < x` `x < pi / 2`]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2148
  have "cos x \<noteq> 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2149
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2150
  have minus_cos_2x: "\<And>X. X - cos (2*x) = X - (cos x) ^ 2 + (sin x) ^ 2" unfolding cos_double by algebra
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2151
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2152
  have "tan x = (tan x + tan x) / 2" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2153
  also have "\<dots> = sin (x + x) / (cos x * cos x) / 2" unfolding add_tan_eq[OF `cos x \<noteq> 0` `cos x \<noteq> 0`] ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2154
  also have "\<dots> = sin (2 * x) / ((cos x) ^ 2 + (cos x) ^ 2 + cos (2*x) - cos (2*x))" unfolding divide_divide_eq_left numeral_2_eq_2 by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2155
  also have "\<dots> = sin (2 * x) / ((cos x) ^ 2 + cos (2*x) + (sin x)^2)" unfolding minus_cos_2x by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2156
  also have "\<dots> = sin (2 * x) / (cos (2*x) + 1)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2157
  finally show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2158
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2159
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2160
lemma lemma_DERIV_tan:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2161
     "cos x \<noteq> 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\<twosuperior>)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2162
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2163
apply (best intro!: DERIV_intros intro: DERIV_chain2) 
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 15077
diff changeset
  2164
apply (auto simp add: divide_inverse numeral_2_eq_2)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2165
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2166
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2167
lemma DERIV_tan [simp]: "cos x \<noteq> 0 ==> DERIV tan x :> inverse((cos x)\<twosuperior>)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2168
by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2169
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2170
lemma isCont_tan [simp]: "cos x \<noteq> 0 ==> isCont tan x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2171
by (rule DERIV_tan [THEN DERIV_isCont])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2172
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2173
lemma LIM_cos_div_sin [simp]: "(%x. cos(x)/sin(x)) -- pi/2 --> 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2174
apply (subgoal_tac "(\<lambda>x. cos x * inverse (sin x)) -- pi * inverse 2 --> 0*1")
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2175
apply (simp add: divide_inverse [symmetric])
22613
2f119f54d150 remove redundant lemmas
huffman
parents: 21404
diff changeset
  2176
apply (rule LIM_mult)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2177
apply (rule_tac [2] inverse_1 [THEN subst])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2178
apply (rule_tac [2] LIM_inverse)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2179
apply (simp_all add: divide_inverse [symmetric]) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2180
apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric]) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2181
apply (blast intro!: DERIV_isCont DERIV_sin DERIV_cos)+
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2182
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2183
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2184
lemma lemma_tan_total: "0 < y ==> \<exists>x. 0 < x & x < pi/2 & y < tan x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2185
apply (cut_tac LIM_cos_div_sin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2186
apply (simp only: LIM_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2187
apply (drule_tac x = "inverse y" in spec, safe, force)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2188
apply (drule_tac ?d1.0 = s in pi_half_gt_zero [THEN [2] real_lbound_gt_zero], safe)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2189
apply (rule_tac x = "(pi/2) - e" in exI)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2190
apply (simp (no_asm_simp))
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2191
apply (drule_tac x = "(pi/2) - e" in spec)
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2192
apply (auto simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2193
apply (rule inverse_less_iff_less [THEN iffD1])
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 15077
diff changeset
  2194
apply (auto simp add: divide_inverse)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2195
apply (rule real_mult_order) 
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2196
apply (subgoal_tac [3] "0 < sin e & 0 < cos e")
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2197
apply (auto intro: cos_gt_zero sin_gt_zero2 simp add: mult_commute) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2198
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2199
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2200
lemma tan_total_pos: "0 \<le> y ==> \<exists>x. 0 \<le> x & x < pi/2 & tan x = y"
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  2201
apply (frule order_le_imp_less_or_eq, safe)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2202
 prefer 2 apply force
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2203
apply (drule lemma_tan_total, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2204
apply (cut_tac f = tan and a = 0 and b = x and y = y in IVT_objl)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2205
apply (auto intro!: DERIV_tan [THEN DERIV_isCont])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2206
apply (drule_tac y = xa in order_le_imp_less_or_eq)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2207
apply (auto dest: cos_gt_zero)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2208
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2209
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2210
lemma lemma_tan_total1: "\<exists>x. -(pi/2) < x & x < (pi/2) & tan x = y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2211
apply (cut_tac linorder_linear [of 0 y], safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2212
apply (drule tan_total_pos)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2213
apply (cut_tac [2] y="-y" in tan_total_pos, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2214
apply (rule_tac [3] x = "-x" in exI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2215
apply (auto intro!: exI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2216
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2217
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2218
lemma tan_total: "EX! x. -(pi/2) < x & x < (pi/2) & tan x = y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2219
apply (cut_tac y = y in lemma_tan_total1, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2220
apply (cut_tac x = xa and y = y in linorder_less_linear, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2221
apply (subgoal_tac [2] "\<exists>z. y < z & z < xa & DERIV tan z :> 0")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2222
apply (subgoal_tac "\<exists>z. xa < z & z < y & DERIV tan z :> 0")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2223
apply (rule_tac [4] Rolle)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2224
apply (rule_tac [2] Rolle)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2225
apply (auto intro!: DERIV_tan DERIV_isCont exI 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2226
            simp add: differentiable_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2227
txt{*Now, simulate TRYALL*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2228
apply (rule_tac [!] DERIV_tan asm_rl)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2229
apply (auto dest!: DERIV_unique [OF _ DERIV_tan]
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  2230
	    simp add: cos_gt_zero_pi [THEN less_imp_neq, THEN not_sym]) 
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2231
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2232
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2233
lemma tan_monotone: assumes "- (pi / 2) < y" and "y < x" and "x < pi / 2"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2234
  shows "tan y < tan x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2235
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2236
  have "\<forall> x'. y \<le> x' \<and> x' \<le> x \<longrightarrow> DERIV tan x' :> inverse (cos x'^2)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2237
  proof (rule allI, rule impI)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2238
    fix x' :: real assume "y \<le> x' \<and> x' \<le> x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2239
    hence "-(pi/2) < x'" and "x' < pi/2" by (auto!)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2240
    from cos_gt_zero_pi[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2241
    have "cos x' \<noteq> 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2242
    thus "DERIV tan x' :> inverse (cos x'^2)" by (rule DERIV_tan)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2243
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2244
  from MVT2[OF `y < x` this] 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2245
  obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<twosuperior>)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2246
  hence "- (pi / 2) < z" and "z < pi / 2" by (auto!)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2247
  hence "0 < cos z" using cos_gt_zero_pi by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2248
  hence inv_pos: "0 < inverse ((cos z)\<twosuperior>)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2249
  have "0 < x - y" using `y < x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2250
  from real_mult_order[OF this inv_pos]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2251
  have "0 < tan x - tan y" unfolding tan_diff by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2252
  thus ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2253
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2254
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2255
lemma tan_monotone': assumes "- (pi / 2) < y" and "y < pi / 2" and "- (pi / 2) < x" and "x < pi / 2"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2256
  shows "(y < x) = (tan y < tan x)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2257
proof
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2258
  assume "y < x" thus "tan y < tan x" using tan_monotone and `- (pi / 2) < y` and `x < pi / 2` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2259
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2260
  assume "tan y < tan x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2261
  show "y < x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2262
  proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2263
    assume "\<not> y < x" hence "x \<le> y" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2264
    hence "tan x \<le> tan y" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2265
    proof (cases "x = y")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2266
      case True thus ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2267
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2268
      case False hence "x < y" using `x \<le> y` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2269
      from tan_monotone[OF `- (pi/2) < x` this `y < pi / 2`] show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2270
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2271
    thus False using `tan y < tan x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2272
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2273
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2274
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2275
lemma tan_inverse: "1 / (tan y) = tan (pi / 2 - y)" unfolding tan_def sin_cos_eq[of y] cos_sin_eq[of y] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2276
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2277
lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2278
  by (simp add: tan_def)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2279
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2280
lemma tan_periodic_nat[simp]: fixes n :: nat shows "tan (x + real n * pi) = tan x" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2281
proof (induct n arbitrary: x)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2282
  case (Suc n)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2283
  have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_plus1 real_of_nat_add real_of_one real_add_mult_distrib by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2284
  show ?case unfolding split_pi_off using Suc by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2285
qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2286
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2287
lemma tan_periodic_int[simp]: fixes i :: int shows "tan (x + real i * pi) = tan x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2288
proof (cases "0 \<le> i")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2289
  case True hence i_nat: "real i = real (nat i)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2290
  show ?thesis unfolding i_nat by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2291
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2292
  case False hence i_nat: "real i = - real (nat (-i))" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2293
  have "tan x = tan (x + real i * pi - real i * pi)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2294
  also have "\<dots> = tan (x + real i * pi)" unfolding i_nat mult_minus_left diff_minus_eq_add by (rule tan_periodic_nat)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2295
  finally show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2296
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2297
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2298
lemma tan_periodic_n[simp]: "tan (x + number_of n * pi) = tan x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2299
  using tan_periodic_int[of _ "number_of n" ] unfolding real_number_of .
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2300
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2301
subsection {* Inverse Trigonometric Functions *}
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2302
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2303
definition
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2304
  arcsin :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2305
  "arcsin y = (THE x. -(pi/2) \<le> x & x \<le> pi/2 & sin x = y)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2306
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2307
definition
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2308
  arccos :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2309
  "arccos y = (THE x. 0 \<le> x & x \<le> pi & cos x = y)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2310
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2311
definition     
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2312
  arctan :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2313
  "arctan y = (THE x. -(pi/2) < x & x < pi/2 & tan x = y)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2314
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2315
lemma arcsin:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2316
     "[| -1 \<le> y; y \<le> 1 |]  
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2317
      ==> -(pi/2) \<le> arcsin y &  
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2318
           arcsin y \<le> pi/2 & sin(arcsin y) = y"
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2319
unfolding arcsin_def by (rule theI' [OF sin_total])
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2320
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2321
lemma arcsin_pi:
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2322
     "[| -1 \<le> y; y \<le> 1 |]  
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2323
      ==> -(pi/2) \<le> arcsin y & arcsin y \<le> pi & sin(arcsin y) = y"
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2324
apply (drule (1) arcsin)
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2325
apply (force intro: order_trans)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2326
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2327
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2328
lemma sin_arcsin [simp]: "[| -1 \<le> y; y \<le> 1 |] ==> sin(arcsin y) = y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2329
by (blast dest: arcsin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2330
      
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2331
lemma arcsin_bounded:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2332
     "[| -1 \<le> y; y \<le> 1 |] ==> -(pi/2) \<le> arcsin y & arcsin y \<le> pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2333
by (blast dest: arcsin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2334
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2335
lemma arcsin_lbound: "[| -1 \<le> y; y \<le> 1 |] ==> -(pi/2) \<le> arcsin y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2336
by (blast dest: arcsin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2337
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2338
lemma arcsin_ubound: "[| -1 \<le> y; y \<le> 1 |] ==> arcsin y \<le> pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2339
by (blast dest: arcsin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2340
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2341
lemma arcsin_lt_bounded:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2342
     "[| -1 < y; y < 1 |] ==> -(pi/2) < arcsin y & arcsin y < pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2343
apply (frule order_less_imp_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2344
apply (frule_tac y = y in order_less_imp_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2345
apply (frule arcsin_bounded)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2346
apply (safe, simp)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2347
apply (drule_tac y = "arcsin y" in order_le_imp_less_or_eq)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2348
apply (drule_tac [2] y = "pi/2" in order_le_imp_less_or_eq, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2349
apply (drule_tac [!] f = sin in arg_cong, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2350
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2351
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2352
lemma arcsin_sin: "[|-(pi/2) \<le> x; x \<le> pi/2 |] ==> arcsin(sin x) = x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2353
apply (unfold arcsin_def)
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2354
apply (rule the1_equality)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2355
apply (rule sin_total, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2356
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2357
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2358
lemma arccos:
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2359
     "[| -1 \<le> y; y \<le> 1 |]  
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2360
      ==> 0 \<le> arccos y & arccos y \<le> pi & cos(arccos y) = y"
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2361
unfolding arccos_def by (rule theI' [OF cos_total])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2362
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2363
lemma cos_arccos [simp]: "[| -1 \<le> y; y \<le> 1 |] ==> cos(arccos y) = y"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2364
by (blast dest: arccos)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2365
      
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2366
lemma arccos_bounded: "[| -1 \<le> y; y \<le> 1 |] ==> 0 \<le> arccos y & arccos y \<le> pi"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2367
by (blast dest: arccos)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2368
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2369
lemma arccos_lbound: "[| -1 \<le> y; y \<le> 1 |] ==> 0 \<le> arccos y"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2370
by (blast dest: arccos)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2371
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2372
lemma arccos_ubound: "[| -1 \<le> y; y \<le> 1 |] ==> arccos y \<le> pi"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2373
by (blast dest: arccos)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2374
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2375
lemma arccos_lt_bounded:
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2376
     "[| -1 < y; y < 1 |]  
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2377
      ==> 0 < arccos y & arccos y < pi"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2378
apply (frule order_less_imp_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2379
apply (frule_tac y = y in order_less_imp_le)
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2380
apply (frule arccos_bounded, auto)
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2381
apply (drule_tac y = "arccos y" in order_le_imp_less_or_eq)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2382
apply (drule_tac [2] y = pi in order_le_imp_less_or_eq, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2383
apply (drule_tac [!] f = cos in arg_cong, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2384
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2385
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2386
lemma arccos_cos: "[|0 \<le> x; x \<le> pi |] ==> arccos(cos x) = x"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2387
apply (simp add: arccos_def)
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2388
apply (auto intro!: the1_equality cos_total)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2389
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2390
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2391
lemma arccos_cos2: "[|x \<le> 0; -pi \<le> x |] ==> arccos(cos x) = -x"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2392
apply (simp add: arccos_def)
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2393
apply (auto intro!: the1_equality cos_total)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2394
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2395
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2396
lemma cos_arcsin: "\<lbrakk>-1 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> cos (arcsin x) = sqrt (1 - x\<twosuperior>)"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2397
apply (subgoal_tac "x\<twosuperior> \<le> 1")
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2398
apply (rule power2_eq_imp_eq)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2399
apply (simp add: cos_squared_eq)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2400
apply (rule cos_ge_zero)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2401
apply (erule (1) arcsin_lbound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2402
apply (erule (1) arcsin_ubound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2403
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2404
apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> \<le> 1\<twosuperior>", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2405
apply (rule power_mono, simp, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2406
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2407
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2408
lemma sin_arccos: "\<lbrakk>-1 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> sin (arccos x) = sqrt (1 - x\<twosuperior>)"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2409
apply (subgoal_tac "x\<twosuperior> \<le> 1")
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2410
apply (rule power2_eq_imp_eq)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2411
apply (simp add: sin_squared_eq)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2412
apply (rule sin_ge_zero)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2413
apply (erule (1) arccos_lbound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2414
apply (erule (1) arccos_ubound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2415
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2416
apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> \<le> 1\<twosuperior>", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2417
apply (rule power_mono, simp, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2418
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2419
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2420
lemma arctan [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2421
     "- (pi/2) < arctan y  & arctan y < pi/2 & tan (arctan y) = y"
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2422
unfolding arctan_def by (rule theI' [OF tan_total])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2423
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2424
lemma tan_arctan: "tan(arctan y) = y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2425
by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2426
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2427
lemma arctan_bounded: "- (pi/2) < arctan y  & arctan y < pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2428
by (auto simp only: arctan)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2429
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2430
lemma arctan_lbound: "- (pi/2) < arctan y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2431
by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2432
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2433
lemma arctan_ubound: "arctan y < pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2434
by (auto simp only: arctan)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2435
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2436
lemma arctan_tan: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2437
      "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2438
apply (unfold arctan_def)
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2439
apply (rule the1_equality)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2440
apply (rule tan_total, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2441
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2442
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2443
lemma arctan_zero_zero [simp]: "arctan 0 = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2444
by (insert arctan_tan [of 0], simp)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2445
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2446
lemma cos_arctan_not_zero [simp]: "cos(arctan x) \<noteq> 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2447
apply (auto simp add: cos_zero_iff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2448
apply (case_tac "n")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2449
apply (case_tac [3] "n")
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2450
apply (cut_tac [2] y = x in arctan_ubound)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2451
apply (cut_tac [4] y = x in arctan_lbound) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2452
apply (auto simp add: real_of_nat_Suc left_distrib mult_less_0_iff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2453
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2454
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2455
lemma tan_sec: "cos x \<noteq> 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2456
apply (rule power_inverse [THEN subst])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2457
apply (rule_tac c1 = "(cos x)\<twosuperior>" in real_mult_right_cancel [THEN iffD1])
22960
114cf1906681 remove redundant lemmas
huffman
parents: 22956
diff changeset
  2458
apply (auto dest: field_power_not_zero
20516
2d2e1d323a05 realpow_divide -> power_divide
huffman
parents: 20432
diff changeset
  2459
        simp add: power_mult_distrib left_distrib power_divide tan_def 
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  2460
                  mult_assoc power_inverse [symmetric])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2461
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2462
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2463
lemma isCont_inverse_function2:
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2464
  fixes f g :: "real \<Rightarrow> real" shows
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2465
  "\<lbrakk>a < x; x < b;
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2466
    \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z;
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2467
    \<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk>
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2468
   \<Longrightarrow> isCont g (f x)"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2469
apply (rule isCont_inverse_function
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2470
       [where f=f and d="min (x - a) (b - x)"])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2471
apply (simp_all add: abs_le_iff)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2472
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2473
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2474
lemma isCont_arcsin: "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> isCont arcsin x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2475
apply (subgoal_tac "isCont arcsin (sin (arcsin x))", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2476
apply (rule isCont_inverse_function2 [where f=sin])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2477
apply (erule (1) arcsin_lt_bounded [THEN conjunct1])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2478
apply (erule (1) arcsin_lt_bounded [THEN conjunct2])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2479
apply (fast intro: arcsin_sin, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2480
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2481
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2482
lemma isCont_arccos: "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> isCont arccos x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2483
apply (subgoal_tac "isCont arccos (cos (arccos x))", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2484
apply (rule isCont_inverse_function2 [where f=cos])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2485
apply (erule (1) arccos_lt_bounded [THEN conjunct1])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2486
apply (erule (1) arccos_lt_bounded [THEN conjunct2])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2487
apply (fast intro: arccos_cos, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2488
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2489
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2490
lemma isCont_arctan: "isCont arctan x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2491
apply (rule arctan_lbound [of x, THEN dense, THEN exE], clarify)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2492
apply (rule arctan_ubound [of x, THEN dense, THEN exE], clarify)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2493
apply (subgoal_tac "isCont arctan (tan (arctan x))", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2494
apply (erule (1) isCont_inverse_function2 [where f=tan])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2495
apply (clarify, rule arctan_tan)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2496
apply (erule (1) order_less_le_trans)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2497
apply (erule (1) order_le_less_trans)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2498
apply (clarify, rule isCont_tan)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2499
apply (rule less_imp_neq [symmetric])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2500
apply (rule cos_gt_zero_pi)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2501
apply (erule (1) order_less_le_trans)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2502
apply (erule (1) order_le_less_trans)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2503
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2504
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2505
lemma DERIV_arcsin:
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2506
  "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> DERIV arcsin x :> inverse (sqrt (1 - x\<twosuperior>))"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2507
apply (rule DERIV_inverse_function [where f=sin and a="-1" and b="1"])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2508
apply (rule lemma_DERIV_subst [OF DERIV_sin])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2509
apply (simp add: cos_arcsin)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2510
apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> < 1\<twosuperior>", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2511
apply (rule power_strict_mono, simp, simp, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2512
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2513
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2514
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2515
apply (erule (1) isCont_arcsin)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2516
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2517
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2518
lemma DERIV_arccos:
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2519
  "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> DERIV arccos x :> inverse (- sqrt (1 - x\<twosuperior>))"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2520
apply (rule DERIV_inverse_function [where f=cos and a="-1" and b="1"])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2521
apply (rule lemma_DERIV_subst [OF DERIV_cos])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2522
apply (simp add: sin_arccos)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2523
apply (subgoal_tac "\<bar>x\<bar>\<twosuperior> < 1\<twosuperior>", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2524
apply (rule power_strict_mono, simp, simp, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2525
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2526
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2527
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2528
apply (erule (1) isCont_arccos)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2529
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2530
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2531
lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<twosuperior>)"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2532
apply (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2533
apply (rule lemma_DERIV_subst [OF DERIV_tan])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2534
apply (rule cos_arctan_not_zero)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2535
apply (simp add: power_inverse tan_sec [symmetric])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2536
apply (subgoal_tac "0 < 1 + x\<twosuperior>", simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2537
apply (simp add: add_pos_nonneg)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2538
apply (simp, simp, simp, rule isCont_arctan)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2539
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2540
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2541
subsection {* More Theorems about Sin and Cos *}
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2542
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2543
lemma cos_45: "cos (pi / 4) = sqrt 2 / 2"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2544
proof -
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2545
  let ?c = "cos (pi / 4)" and ?s = "sin (pi / 4)"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2546
  have nonneg: "0 \<le> ?c"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2547
    by (rule cos_ge_zero, rule order_trans [where y=0], simp_all)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2548
  have "0 = cos (pi / 4 + pi / 4)"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2549
    by simp
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2550
  also have "cos (pi / 4 + pi / 4) = ?c\<twosuperior> - ?s\<twosuperior>"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2551
    by (simp only: cos_add power2_eq_square)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2552
  also have "\<dots> = 2 * ?c\<twosuperior> - 1"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2553
    by (simp add: sin_squared_eq)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2554
  finally have "?c\<twosuperior> = (sqrt 2 / 2)\<twosuperior>"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2555
    by (simp add: power_divide)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2556
  thus ?thesis
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2557
    using nonneg by (rule power2_eq_imp_eq) simp
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2558
qed
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2559
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2560
lemma cos_30: "cos (pi / 6) = sqrt 3 / 2"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2561
proof -
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2562
  let ?c = "cos (pi / 6)" and ?s = "sin (pi / 6)"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2563
  have pos_c: "0 < ?c"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2564
    by (rule cos_gt_zero, simp, simp)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2565
  have "0 = cos (pi / 6 + pi / 6 + pi / 6)"
23066
26a9157b620a new field_combine_numerals simproc, which uses fractions as coefficients
huffman
parents: 23053
diff changeset
  2566
    by simp
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2567
  also have "\<dots> = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2568
    by (simp only: cos_add sin_add)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2569
  also have "\<dots> = ?c * (?c\<twosuperior> - 3 * ?s\<twosuperior>)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  2570
    by (simp add: algebra_simps power2_eq_square)
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2571
  finally have "?c\<twosuperior> = (sqrt 3 / 2)\<twosuperior>"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2572
    using pos_c by (simp add: sin_squared_eq power_divide)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2573
  thus ?thesis
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2574
    using pos_c [THEN order_less_imp_le]
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2575
    by (rule power2_eq_imp_eq) simp
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2576
qed
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2577
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2578
lemma sin_45: "sin (pi / 4) = sqrt 2 / 2"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2579
proof -
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2580
  have "sin (pi / 4) = cos (pi / 2 - pi / 4)" by (rule sin_cos_eq)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2581
  also have "pi / 2 - pi / 4 = pi / 4" by simp
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2582
  also have "cos (pi / 4) = sqrt 2 / 2" by (rule cos_45)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2583
  finally show ?thesis .
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2584
qed
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2585
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2586
lemma sin_60: "sin (pi / 3) = sqrt 3 / 2"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2587
proof -
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2588
  have "sin (pi / 3) = cos (pi / 2 - pi / 3)" by (rule sin_cos_eq)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2589
  also have "pi / 2 - pi / 3 = pi / 6" by simp
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2590
  also have "cos (pi / 6) = sqrt 3 / 2" by (rule cos_30)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2591
  finally show ?thesis .
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2592
qed
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2593
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2594
lemma cos_60: "cos (pi / 3) = 1 / 2"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2595
apply (rule power2_eq_imp_eq)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2596
apply (simp add: cos_squared_eq sin_60 power_divide)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2597
apply (rule cos_ge_zero, rule order_trans [where y=0], simp_all)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2598
done
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2599
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2600
lemma sin_30: "sin (pi / 6) = 1 / 2"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2601
proof -
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2602
  have "sin (pi / 6) = cos (pi / 2 - pi / 6)" by (rule sin_cos_eq)
23066
26a9157b620a new field_combine_numerals simproc, which uses fractions as coefficients
huffman
parents: 23053
diff changeset
  2603
  also have "pi / 2 - pi / 6 = pi / 3" by simp
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2604
  also have "cos (pi / 3) = 1 / 2" by (rule cos_60)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2605
  finally show ?thesis .
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2606
qed
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2607
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2608
lemma tan_30: "tan (pi / 6) = 1 / sqrt 3"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2609
unfolding tan_def by (simp add: sin_30 cos_30)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2610
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2611
lemma tan_45: "tan (pi / 4) = 1"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2612
unfolding tan_def by (simp add: sin_45 cos_45)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2613
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2614
lemma tan_60: "tan (pi / 3) = sqrt 3"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2615
unfolding tan_def by (simp add: sin_60 cos_60)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  2616
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  2617
text{*NEEDED??*}
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2618
lemma [simp]:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2619
     "sin (x + 1 / 2 * real (Suc m) * pi) =  
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2620
      cos (x + 1 / 2 * real  (m) * pi)"
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2621
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2622
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  2623
text{*NEEDED??*}
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2624
lemma [simp]:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2625
     "sin (x + real (Suc m) * pi / 2) =  
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2626
      cos (x + real (m) * pi / 2)"
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2627
by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2628
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2629
lemma DERIV_sin_add [simp]: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2630
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2631
apply (rule_tac f = sin and g = "%x. x + k" in DERIV_chain2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2632
apply (best intro!: DERIV_intros intro: DERIV_chain2)+
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2633
apply (simp (no_asm))
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2634
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2635
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2636
lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n"
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2637
proof -
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2638
  have "sin ((real n + 1/2) * pi) = cos (real n * pi)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  2639
    by (auto simp add: algebra_simps sin_add)
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2640
  thus ?thesis
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2641
    by (simp add: real_of_nat_Suc left_distrib add_divide_distrib 
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2642
                  mult_commute [of pi])
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2643
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2644
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2645
lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2646
by (simp add: cos_double mult_assoc power_add [symmetric] numeral_2_eq_2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2647
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2648
lemma cos_3over2_pi [simp]: "cos (3 / 2 * pi) = 0"
23066
26a9157b620a new field_combine_numerals simproc, which uses fractions as coefficients
huffman
parents: 23053
diff changeset
  2649
apply (subgoal_tac "cos (pi + pi/2) = 0", simp)
26a9157b620a new field_combine_numerals simproc, which uses fractions as coefficients
huffman
parents: 23053
diff changeset
  2650
apply (subst cos_add, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2651
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2652
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2653
lemma sin_2npi [simp]: "sin (2 * real (n::nat) * pi) = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2654
by (auto simp add: mult_assoc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2655
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2656
lemma sin_3over2_pi [simp]: "sin (3 / 2 * pi) = - 1"
23066
26a9157b620a new field_combine_numerals simproc, which uses fractions as coefficients
huffman
parents: 23053
diff changeset
  2657
apply (subgoal_tac "sin (pi + pi/2) = - 1", simp)
26a9157b620a new field_combine_numerals simproc, which uses fractions as coefficients
huffman
parents: 23053
diff changeset
  2658
apply (subst sin_add, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2659
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2660
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2661
(*NEEDED??*)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2662
lemma [simp]:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2663
     "cos(x + 1 / 2 * real(Suc m) * pi) = -sin (x + 1 / 2 * real m * pi)"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2664
apply (simp only: cos_add sin_add real_of_nat_Suc right_distrib left_distrib minus_mult_right, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2665
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2666
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2667
(*NEEDED??*)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2668
lemma [simp]: "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2669
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib add_divide_distrib, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2670
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2671
lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2672
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib add_divide_distrib, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2673
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2674
lemma DERIV_cos_add [simp]: "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2675
apply (rule lemma_DERIV_subst)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2676
apply (rule_tac f = cos and g = "%x. x + k" in DERIV_chain2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2677
apply (best intro!: DERIV_intros intro: DERIV_chain2)+
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2678
apply (simp (no_asm))
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2679
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2680
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
  2681
lemma sin_zero_abs_cos_one: "sin x = 0 ==> \<bar>cos x\<bar> = 1"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2682
by (auto simp add: sin_zero_iff even_mult_two_ex)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2683
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2684
lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2685
by (cut_tac x = x in sin_cos_squared_add3, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2686
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2687
subsection {* Machins formula *}
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2688
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2689
lemma tan_total_pi4: assumes "\<bar>x\<bar> < 1"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2690
  shows "\<exists> z. - (pi / 4) < z \<and> z < pi / 4 \<and> tan z = x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2691
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2692
  obtain z where "- (pi / 2) < z" and "z < pi / 2" and "tan z = x" using tan_total by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2693
  have "tan (pi / 4) = 1" and "tan (- (pi / 4)) = - 1" using tan_45 tan_minus by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2694
  have "z \<noteq> pi / 4" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2695
  proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2696
    assume "\<not> (z \<noteq> pi / 4)" hence "z = pi / 4" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2697
    have "tan z = 1" unfolding `z = pi / 4` `tan (pi / 4) = 1` ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2698
    thus False unfolding `tan z = x` using `\<bar>x\<bar> < 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2699
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2700
  have "z \<noteq> - (pi / 4)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2701
  proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2702
    assume "\<not> (z \<noteq> - (pi / 4))" hence "z = - (pi / 4)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2703
    have "tan z = - 1" unfolding `z = - (pi / 4)` `tan (- (pi / 4)) = - 1` ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2704
    thus False unfolding `tan z = x` using `\<bar>x\<bar> < 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2705
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2706
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2707
  have "z < pi / 4"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2708
  proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2709
    assume "\<not> (z < pi / 4)" hence "pi / 4 < z" using `z \<noteq> pi / 4` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2710
    have "- (pi / 2) < pi / 4" using m2pi_less_pi by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2711
    from tan_monotone[OF this `pi / 4 < z` `z < pi / 2`] 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2712
    have "1 < x" unfolding `tan z = x` `tan (pi / 4) = 1` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2713
    thus False using `\<bar>x\<bar> < 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2714
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2715
  moreover 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2716
  have "-(pi / 4) < z"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2717
  proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2718
    assume "\<not> (-(pi / 4) < z)" hence "z < - (pi / 4)" using `z \<noteq> - (pi / 4)` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2719
    have "-(pi / 4) < pi / 2" using m2pi_less_pi by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2720
    from tan_monotone[OF `-(pi / 2) < z` `z < -(pi / 4)` this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2721
    have "x < - 1" unfolding `tan z = x` `tan (-(pi / 4)) = - 1` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2722
    thus False using `\<bar>x\<bar> < 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2723
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2724
  ultimately show ?thesis using `tan z = x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2725
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2726
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2727
lemma arctan_add: assumes "\<bar>x\<bar> \<le> 1" and "\<bar>y\<bar> < 1"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2728
  shows "arctan x + arctan y = arctan ((x + y) / (1 - x * y))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2729
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2730
  obtain y' where "-(pi/4) < y'" and "y' < pi/4" and "tan y' = y" using tan_total_pi4[OF `\<bar>y\<bar> < 1`] by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2731
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2732
  have "pi / 4 < pi / 2" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2733
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2734
  have "\<exists> x'. -(pi/4) \<le> x' \<and> x' \<le> pi/4 \<and> tan x' = x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2735
  proof (cases "\<bar>x\<bar> < 1")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2736
    case True from tan_total_pi4[OF this] obtain x' where "-(pi/4) < x'" and "x' < pi/4" and "tan x' = x" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2737
    hence "-(pi/4) \<le> x'" and "x' \<le> pi/4" and "tan x' = x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2738
    thus ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2739
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2740
    case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2741
    show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2742
    proof (cases "x = 1")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2743
      case True hence "tan (pi/4) = x" using tan_45 by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2744
      moreover 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2745
      have "- pi \<le> pi" unfolding minus_le_self_iff by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2746
      hence "-(pi/4) \<le> pi/4" and "pi/4 \<le> pi/4" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2747
      ultimately show ?thesis by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2748
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2749
      case False hence "x = -1" using `\<not> \<bar>x\<bar> < 1` and `\<bar>x\<bar> \<le> 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2750
      hence "tan (-(pi/4)) = x" using tan_45 tan_minus by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2751
      moreover 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2752
      have "- pi \<le> pi" unfolding minus_le_self_iff by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2753
      hence "-(pi/4) \<le> pi/4" and "-(pi/4) \<le> -(pi/4)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2754
      ultimately show ?thesis by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2755
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2756
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2757
  then obtain x' where "-(pi/4) \<le> x'" and "x' \<le> pi/4" and "tan x' = x" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2758
  hence "-(pi/2) < x'" and "x' < pi/2" using order_le_less_trans[OF `x' \<le> pi/4` `pi / 4 < pi / 2`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2759
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2760
  have "cos x' \<noteq> 0" using cos_gt_zero_pi[THEN less_imp_neq] and `-(pi/2) < x'` and `x' < pi/2` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2761
  moreover have "cos y' \<noteq> 0" using cos_gt_zero_pi[THEN less_imp_neq] and `-(pi/4) < y'` and `y' < pi/4` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2762
  ultimately have "cos x' * cos y' \<noteq> 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2763
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2764
  have divide_nonzero_divide: "\<And> A B C :: real. C \<noteq> 0 \<Longrightarrow> (A / C) / (B / C) = A / B" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2765
  have divide_mult_commute: "\<And> A B C D :: real. A * B / (C * D) = (A / C) * (B / D)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2766
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2767
  have "tan (x' + y') = sin (x' + y') / (cos x' * cos y' - sin x' * sin y')" unfolding tan_def cos_add ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2768
  also have "\<dots> = (tan x' + tan y') / ((cos x' * cos y' - sin x' * sin y') / (cos x' * cos y'))" unfolding add_tan_eq[OF `cos x' \<noteq> 0` `cos y' \<noteq> 0`] divide_nonzero_divide[OF `cos x' * cos y' \<noteq> 0`] ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2769
  also have "\<dots> = (tan x' + tan y') / (1 - tan x' * tan y')" unfolding tan_def diff_divide_distrib divide_self[OF `cos x' * cos y' \<noteq> 0`] unfolding divide_mult_commute ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2770
  finally have tan_eq: "tan (x' + y') = (x + y) / (1 - x * y)" unfolding `tan y' = y` `tan x' = x` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2771
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2772
  have "arctan (tan (x' + y')) = x' + y'" using `-(pi/4) < y'` `-(pi/4) \<le> x'` `y' < pi/4` and `x' \<le> pi/4` by (auto intro!: arctan_tan)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2773
  moreover have "arctan (tan (x')) = x'" using `-(pi/2) < x'` and `x' < pi/2` by (auto intro!: arctan_tan)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2774
  moreover have "arctan (tan (y')) = y'" using `-(pi/4) < y'` and `y' < pi/4` by (auto intro!: arctan_tan)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2775
  ultimately have "arctan x + arctan y = arctan (tan (x' + y'))" unfolding `tan y' = y` [symmetric] `tan x' = x`[symmetric] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2776
  thus "arctan x + arctan y = arctan ((x + y) / (1 - x * y))" unfolding tan_eq .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2777
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2778
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2779
lemma arctan1_eq_pi4: "arctan 1 = pi / 4" unfolding tan_45[symmetric] by (rule arctan_tan, auto simp add: m2pi_less_pi)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2780
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2781
theorem machin: "pi / 4 = 4 * arctan (1/5) - arctan (1 / 239)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2782
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2783
  have "\<bar>1 / 5\<bar> < (1 :: real)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2784
  from arctan_add[OF less_imp_le[OF this] this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2785
  have "2 * arctan (1 / 5) = arctan (5 / 12)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2786
  moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2787
  have "\<bar>5 / 12\<bar> < (1 :: real)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2788
  from arctan_add[OF less_imp_le[OF this] this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2789
  have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2790
  moreover 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2791
  have "\<bar>1\<bar> \<le> (1::real)" and "\<bar>1 / 239\<bar> < (1::real)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2792
  from arctan_add[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2793
  have "arctan 1 + arctan (1 / 239) = arctan (120 / 119)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2794
  ultimately have "arctan 1 + arctan (1 / 239) = 4 * arctan (1 / 5)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2795
  thus ?thesis unfolding arctan1_eq_pi4 by algebra
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2796
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2797
subsection {* Introducing the arcus tangens power series *}
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2798
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2799
lemma monoseq_arctan_series: fixes x :: real
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2800
  assumes "\<bar>x\<bar> \<le> 1" shows "monoseq (\<lambda> n. 1 / real (n*2+1) * x^(n*2+1))" (is "monoseq ?a")
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  2801
proof (cases "x = 0") case True thus ?thesis unfolding monoseq_def One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2802
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2803
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2804
  have "norm x \<le> 1" and "x \<le> 1" and "-1 \<le> x" using assms by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2805
  show "monoseq ?a"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2806
  proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2807
    { fix n fix x :: real assume "0 \<le> x" and "x \<le> 1"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2808
      have "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \<le> 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2809
      proof (rule mult_mono)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2810
	show "1 / real (Suc (Suc n * 2)) \<le> 1 / real (Suc (n * 2))" by (rule frac_le) simp_all
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2811
	show "0 \<le> 1 / real (Suc (n * 2))" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2812
	show "x ^ Suc (Suc n * 2) \<le> x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: `0 \<le> x` `x \<le> 1`)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2813
	show "0 \<le> x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: `0 \<le> x`)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2814
      qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2815
    } note mono = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2816
    
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2817
    show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2818
    proof (cases "0 \<le> x")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2819
      case True from mono[OF this `x \<le> 1`, THEN allI]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2820
      show ?thesis unfolding Suc_plus1[symmetric] by (rule mono_SucI2)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2821
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2822
      case False hence "0 \<le> -x" and "-x \<le> 1" using `-1 \<le> x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2823
      from mono[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2824
      have "\<And>n. 1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \<ge> 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" using `0 \<le> -x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2825
      thus ?thesis unfolding Suc_plus1[symmetric] by (rule mono_SucI1[OF allI])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2826
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2827
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2828
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2829
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2830
lemma zeroseq_arctan_series: fixes x :: real
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2831
  assumes "\<bar>x\<bar> \<le> 1" shows "(\<lambda> n. 1 / real (n*2+1) * x^(n*2+1)) ----> 0" (is "?a ----> 0")
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  2832
proof (cases "x = 0") case True thus ?thesis unfolding One_nat_def by (auto simp add: LIMSEQ_const)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2833
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2834
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2835
  have "norm x \<le> 1" and "x \<le> 1" and "-1 \<le> x" using assms by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2836
  show "?a ----> 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2837
  proof (cases "\<bar>x\<bar> < 1")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2838
    case True hence "norm x < 1" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2839
    from LIMSEQ_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF `norm x < 1`, THEN LIMSEQ_Suc]]
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  2840
    have "(\<lambda>n. 1 / real (n + 1) * x ^ (n + 1)) ----> 0"
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  2841
      unfolding inverse_eq_divide Suc_plus1 by simp
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  2842
    then show ?thesis using pos2 by (rule LIMSEQ_linear)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2843
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2844
    case False hence "x = -1 \<or> x = 1" using `\<bar>x\<bar> \<le> 1` by auto
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  2845
    hence n_eq: "\<And> n. x ^ (n * 2 + 1) = x" unfolding One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2846
    from LIMSEQ_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] LIMSEQ_const[of x]]
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  2847
    show ?thesis unfolding n_eq Suc_plus1 by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2848
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2849
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2850
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2851
lemma summable_arctan_series: fixes x :: real and n :: nat
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2852
  assumes "\<bar>x\<bar> \<le> 1" shows "summable (\<lambda> k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "summable (?c x)")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2853
  by (rule summable_Leibniz(1), rule zeroseq_arctan_series[OF assms], rule monoseq_arctan_series[OF assms])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2854
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2855
lemma less_one_imp_sqr_less_one: fixes x :: real assumes "\<bar>x\<bar> < 1" shows "x^2 < 1"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2856
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2857
  from mult_mono1[OF less_imp_le[OF `\<bar>x\<bar> < 1`] abs_ge_zero[of x]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2858
  have "\<bar> x^2 \<bar> < 1" using `\<bar> x \<bar> < 1` unfolding numeral_2_eq_2 power_Suc2 by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2859
  thus ?thesis using zero_le_power2 by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2860
qed 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2861
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2862
lemma DERIV_arctan_series: assumes "\<bar> x \<bar> < 1"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2863
  shows "DERIV (\<lambda> x'. \<Sum> k. (-1)^k * (1 / real (k*2+1) * x' ^ (k*2+1))) x :> (\<Sum> k. (-1)^k * x^(k*2))" (is "DERIV ?arctan _ :> ?Int")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2864
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2865
  let "?f n" = "if even n then (-1)^(n div 2) * 1 / real (Suc n) else 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2866
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2867
  { fix n :: nat assume "even n" hence "2 * (n div 2) = n" by presburger } note n_even=this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2868
  have if_eq: "\<And> n x'. ?f n * real (Suc n) * x'^n = (if even n then (-1)^(n div 2) * x'^(2 * (n div 2)) else 0)" using n_even by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2869
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2870
  { fix x :: real assume "\<bar>x\<bar> < 1" hence "x^2 < 1" by (rule less_one_imp_sqr_less_one)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2871
    have "summable (\<lambda> n. -1 ^ n * (x^2) ^n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2872
      by (rule summable_Leibniz(1), auto intro!: LIMSEQ_realpow_zero monoseq_realpow `x^2 < 1` order_less_imp_le[OF `x^2 < 1`])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2873
    hence "summable (\<lambda> n. -1 ^ n * x^(2*n))" unfolding power_mult .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2874
  } note summable_Integral = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2875
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2876
  { fix f :: "nat \<Rightarrow> real"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2877
    have "\<And> x. f sums x = (\<lambda> n. if even n then f (n div 2) else 0) sums x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2878
    proof
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2879
      fix x :: real assume "f sums x" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2880
      from sums_if[OF sums_zero this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2881
      show "(\<lambda> n. if even n then f (n div 2) else 0) sums x" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2882
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2883
      fix x :: real assume "(\<lambda> n. if even n then f (n div 2) else 0) sums x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2884
      from LIMSEQ_linear[OF this[unfolded sums_def] pos2, unfolded sum_split_even_odd[unfolded mult_commute]]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2885
      show "f sums x" unfolding sums_def by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2886
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2887
    hence "op sums f = op sums (\<lambda> n. if even n then f (n div 2) else 0)" ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2888
  } note sums_even = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2889
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2890
  have Int_eq: "(\<Sum> n. ?f n * real (Suc n) * x^n) = ?Int" unfolding if_eq mult_commute[of _ 2] suminf_def sums_even[of "\<lambda> n. -1 ^ n * x ^ (2 * n)", symmetric]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2891
    by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2892
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2893
  { fix x :: real
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2894
    have if_eq': "\<And> n. (if even n then -1 ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n = 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2895
      (if even n then -1 ^ (n div 2) * (1 / real (Suc (2 * (n div 2))) * x ^ Suc (2 * (n div 2))) else 0)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2896
      using n_even by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2897
    have idx_eq: "\<And> n. n * 2 + 1 = Suc (2 * n)" by auto 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2898
    have "(\<Sum> n. ?f n * x^(Suc n)) = ?arctan x" unfolding if_eq' idx_eq suminf_def sums_even[of "\<lambda> n. -1 ^ n * (1 / real (Suc (2 * n)) * x ^ Suc (2 * n))", symmetric]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2899
      by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2900
  } note arctan_eq = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2901
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2902
  have "DERIV (\<lambda> x. \<Sum> n. ?f n * x^(Suc n)) x :> (\<Sum> n. ?f n * real (Suc n) * x^n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2903
  proof (rule DERIV_power_series')
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2904
    show "x \<in> {- 1 <..< 1}" using `\<bar> x \<bar> < 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2905
    { fix x' :: real assume x'_bounds: "x' \<in> {- 1 <..< 1}"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2906
      hence "\<bar>x'\<bar> < 1" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2907
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2908
      let ?S = "\<Sum> n. (-1)^n * x'^(2 * n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2909
      show "summable (\<lambda> n. ?f n * real (Suc n) * x'^n)" unfolding if_eq
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2910
	by (rule sums_summable[where l="0 + ?S"], rule sums_if, rule sums_zero, rule summable_sums, rule summable_Integral[OF `\<bar>x'\<bar> < 1`])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2911
    }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2912
  qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2913
  thus ?thesis unfolding Int_eq arctan_eq .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2914
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2915
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2916
lemma arctan_series: assumes "\<bar> x \<bar> \<le> 1"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2917
  shows "arctan x = (\<Sum> k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "_ = suminf (\<lambda> n. ?c x n)")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2918
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2919
  let "?c' x n" = "(-1)^n * x^(n*2)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2920
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2921
  { fix r x :: real assume "0 < r" and "r < 1" and "\<bar> x \<bar> < r"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2922
    have "\<bar>x\<bar> < 1" using `r < 1` and `\<bar>x\<bar> < r` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2923
    from DERIV_arctan_series[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2924
    have "DERIV (\<lambda> x. suminf (?c x)) x :> (suminf (?c' x))" .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2925
  } note DERIV_arctan_suminf = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2926
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2927
  { fix x :: real assume "\<bar>x\<bar> \<le> 1" note summable_Leibniz[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2928
  note arctan_series_borders = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2929
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2930
  { fix x :: real assume "\<bar>x\<bar> < 1" have "arctan x = (\<Sum> k. ?c x k)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2931
  proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2932
    obtain r where "\<bar>x\<bar> < r" and "r < 1" using dense[OF `\<bar>x\<bar> < 1`] by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2933
    hence "0 < r" and "-r < x" and "x < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2934
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2935
    have suminf_eq_arctan_bounded: "\<And> x a b. \<lbrakk> -r < a ; b < r ; a < b ; a \<le> x ; x \<le> b \<rbrakk> \<Longrightarrow> suminf (?c x) - arctan x = suminf (?c a) - arctan a"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2936
    proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2937
      fix x a b assume "-r < a" and "b < r" and "a < b" and "a \<le> x" and "x \<le> b"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2938
      hence "\<bar>x\<bar> < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2939
      show "suminf (?c x) - arctan x = suminf (?c a) - arctan a"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2940
      proof (rule DERIV_isconst2[of "a" "b"])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2941
	show "a < b" and "a \<le> x" and "x \<le> b" using `a < b` `a \<le> x` `x \<le> b` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2942
	have "\<forall> x. -r < x \<and> x < r \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2943
	proof (rule allI, rule impI)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2944
	  fix x assume "-r < x \<and> x < r" hence "\<bar>x\<bar> < r" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2945
	  hence "\<bar>x\<bar> < 1" using `r < 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2946
	  have "\<bar> - (x^2) \<bar> < 1" using less_one_imp_sqr_less_one[OF `\<bar>x\<bar> < 1`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2947
	  hence "(\<lambda> n. (- (x^2)) ^ n) sums (1 / (1 - (- (x^2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2948
	  hence "(?c' x) sums (1 / (1 - (- (x^2))))" unfolding power_mult_distrib[symmetric] power_mult nat_mult_commute[of _ 2] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2949
	  hence suminf_c'_eq_geom: "inverse (1 + x^2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2950
	  have "DERIV (\<lambda> x. suminf (?c x)) x :> (inverse (1 + x^2))" unfolding suminf_c'_eq_geom
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2951
	    by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2952
	  from DERIV_add_minus[OF this DERIV_arctan]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2953
	  show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0" unfolding diff_minus by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2954
	qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2955
	hence DERIV_in_rball: "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `-r < a` `b < r` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2956
	thus "\<forall> y. a < y \<and> y < b \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) y :> 0" using `\<bar>x\<bar> < r` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2957
	show "\<forall> y. a \<le> y \<and> y \<le> b \<longrightarrow> isCont (\<lambda> x. suminf (?c x) - arctan x) y" using DERIV_in_rball DERIV_isCont by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2958
      qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2959
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2960
    
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2961
    have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2962
      unfolding Suc_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2963
    
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2964
    have "suminf (?c x) - arctan x = 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2965
    proof (cases "x = 0")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2966
      case True thus ?thesis using suminf_arctan_zero by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2967
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2968
      case False hence "0 < \<bar>x\<bar>" and "- \<bar>x\<bar> < \<bar>x\<bar>" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2969
      have "suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>) = suminf (?c 0) - arctan 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2970
	by (rule suminf_eq_arctan_bounded[where x="0" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>", symmetric], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2971
      moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2972
      have "suminf (?c x) - arctan x = suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2973
	by (rule suminf_eq_arctan_bounded[where x="x" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>"], auto simp add: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>`)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2974
      ultimately 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2975
      show ?thesis using suminf_arctan_zero by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2976
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2977
    thus ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2978
  qed } note when_less_one = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2979
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2980
  show "arctan x = suminf (\<lambda> n. ?c x n)"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2981
  proof (cases "\<bar>x\<bar> < 1")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2982
    case True thus ?thesis by (rule when_less_one)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2983
  next case False hence "\<bar>x\<bar> = 1" using `\<bar>x\<bar> \<le> 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2984
    let "?a x n" = "\<bar>1 / real (n*2+1) * x^(n*2+1)\<bar>"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2985
    let "?diff x n" = "\<bar> arctan x - (\<Sum> i = 0..<n. ?c x i)\<bar>"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2986
    { fix n :: nat
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2987
      have "0 < (1 :: real)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2988
      moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2989
      { fix x :: real assume "0 < x" and "x < 1" hence "\<bar>x\<bar> \<le> 1" and "\<bar>x\<bar> < 1" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2990
	from `0 < x` have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2991
	note bounds = mp[OF arctan_series_borders(2)[OF `\<bar>x\<bar> \<le> 1`] this, unfolded when_less_one[OF `\<bar>x\<bar> < 1`, symmetric], THEN spec]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2992
	have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos, auto simp only: zero_less_power[OF `0 < x`], auto)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2993
	hence a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2994
        have "?diff x n \<le> ?a x n"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2995
	proof (cases "even n")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2996
	  case True hence sgn_pos: "(-1)^n = (1::real)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2997
	  from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2998
	  from bounds[of m, unfolded this atLeastAtMost_iff]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2999
	  have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n + 1. (?c x i)) - (\<Sum>i = 0..<n. (?c x i))" by auto
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3000
	  also have "\<dots> = ?c x n" unfolding One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3001
	  also have "\<dots> = ?a x n" unfolding sgn_pos a_pos by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3002
	  finally show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3003
	next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3004
	  case False hence sgn_neg: "(-1)^n = (-1::real)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3005
	  from `odd n` obtain m where m_def: "2 * m + 1 = n" unfolding odd_Suc_mult_two_ex by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3006
	  hence m_plus: "2 * (m + 1) = n + 1" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3007
	  from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3008
	  have "\<bar>arctan x - (\<Sum>i = 0..<n. (?c x i))\<bar> \<le> (\<Sum>i = 0..<n. (?c x i)) - (\<Sum>i = 0..<n+1. (?c x i))" by auto
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3009
	  also have "\<dots> = - ?c x n" unfolding One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3010
	  also have "\<dots> = ?a x n" unfolding sgn_neg a_pos by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3011
	  finally show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3012
	qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3013
        hence "0 \<le> ?a x n - ?diff x n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3014
      }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3015
      hence "\<forall> x \<in> { 0 <..< 1 }. 0 \<le> ?a x n - ?diff x n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3016
      moreover have "\<And>x. isCont (\<lambda> x. ?a x n - ?diff x n) x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3017
	unfolding real_diff_def divide_inverse
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3018
	by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3019
      ultimately have "0 \<le> ?a 1 n - ?diff 1 n" by (rule LIM_less_bound)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3020
      hence "?diff 1 n \<le> ?a 1 n" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3021
    }
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3022
    have "?a 1 ----> 0"
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3023
      unfolding LIMSEQ_rabs_zero power_one divide_inverse One_nat_def
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3024
      by (auto intro!: LIMSEQ_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3025
    have "?diff 1 ----> 0"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3026
    proof (rule LIMSEQ_I)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3027
      fix r :: real assume "0 < r"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3028
      obtain N :: nat where N_I: "\<And> n. N \<le> n \<Longrightarrow> ?a 1 n < r" using LIMSEQ_D[OF `?a 1 ----> 0` `0 < r`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3029
      { fix n assume "N \<le> n" from `?diff 1 n \<le> ?a 1 n` N_I[OF this]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3030
	have "norm (?diff 1 n - 0) < r" by auto }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3031
      thus "\<exists> N. \<forall> n \<ge> N. norm (?diff 1 n - 0) < r" by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3032
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3033
    from this[unfolded LIMSEQ_rabs_zero real_diff_def add_commute[of "arctan 1"], THEN LIMSEQ_add_const, of "- arctan 1", THEN LIMSEQ_minus]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3034
    have "(?c 1) sums (arctan 1)" unfolding sums_def by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3035
    hence "arctan 1 = (\<Sum> i. ?c 1 i)" by (rule sums_unique)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3036
    
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3037
    show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3038
    proof (cases "x = 1", simp add: `arctan 1 = (\<Sum> i. ?c 1 i)`)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3039
      assume "x \<noteq> 1" hence "x = -1" using `\<bar>x\<bar> = 1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3040
      
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3041
      have "- (pi / 2) < 0" using pi_gt_zero by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3042
      have "- (2 * pi) < 0" using pi_gt_zero by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3043
      
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3044
      have c_minus_minus: "\<And> i. ?c (- 1) i = - ?c 1 i" unfolding One_nat_def by auto
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3045
    
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3046
      have "arctan (- 1) = arctan (tan (-(pi / 4)))" unfolding tan_45 tan_minus ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3047
      also have "\<dots> = - (pi / 4)" by (rule arctan_tan, auto simp add: order_less_trans[OF `- (pi / 2) < 0` pi_gt_zero])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3048
      also have "\<dots> = - (arctan (tan (pi / 4)))" unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric], auto simp add: order_less_trans[OF `- (2 * pi) < 0` pi_gt_zero])
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3049
      also have "\<dots> = - (arctan 1)" unfolding tan_45 ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3050
      also have "\<dots> = - (\<Sum> i. ?c 1 i)" using `arctan 1 = (\<Sum> i. ?c 1 i)` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3051
      also have "\<dots> = (\<Sum> i. ?c (- 1) i)" using suminf_minus[OF sums_summable[OF `(?c 1) sums (arctan 1)`]] unfolding c_minus_minus by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3052
      finally show ?thesis using `x = -1` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3053
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3054
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3055
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3056
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3057
lemma arctan_half: fixes x :: real
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3058
  shows "arctan x = 2 * arctan (x / (1 + sqrt(1 + x^2)))"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3059
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3060
  obtain y where low: "- (pi / 2) < y" and high: "y < pi / 2" and y_eq: "tan y = x" using tan_total by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3061
  hence low2: "- (pi / 2) < y / 2" and high2: "y / 2 < pi / 2" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3062
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3063
  have divide_nonzero_divide: "\<And> A B C :: real. C \<noteq> 0 \<Longrightarrow> A / B = (A / C) / (B / C)" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3064
  
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3065
  have "0 < cos y" using cos_gt_zero_pi[OF low high] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3066
  hence "cos y \<noteq> 0" and cos_sqrt: "sqrt ((cos y) ^ 2) = cos y" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3067
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3068
  have "1 + (tan y)^2 = 1 + sin y^2 / cos y^2" unfolding tan_def power_divide ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3069
  also have "\<dots> = cos y^2 / cos y^2 + sin y^2 / cos y^2" using `cos y \<noteq> 0` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3070
  also have "\<dots> = 1 / cos y^2" unfolding add_divide_distrib[symmetric] sin_cos_squared_add2 ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3071
  finally have "1 + (tan y)^2 = 1 / cos y^2" .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3072
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3073
  have "sin y / (cos y + 1) = tan y / ((cos y + 1) / cos y)" unfolding tan_def divide_nonzero_divide[OF `cos y \<noteq> 0`, symmetric] ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3074
  also have "\<dots> = tan y / (1 + 1 / cos y)" using `cos y \<noteq> 0` unfolding add_divide_distrib by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3075
  also have "\<dots> = tan y / (1 + 1 / sqrt(cos y^2))" unfolding cos_sqrt ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3076
  also have "\<dots> = tan y / (1 + sqrt(1 / cos y^2))" unfolding real_sqrt_divide by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3077
  finally have eq: "sin y / (cos y + 1) = tan y / (1 + sqrt(1 + (tan y)^2))" unfolding `1 + (tan y)^2 = 1 / cos y^2` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3078
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3079
  have "arctan x = y" using arctan_tan low high y_eq by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3080
  also have "\<dots> = 2 * (arctan (tan (y/2)))" using arctan_tan[OF low2 high2] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3081
  also have "\<dots> = 2 * (arctan (sin y / (cos y + 1)))" unfolding tan_half[OF low2 high2] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3082
  finally show ?thesis unfolding eq `tan y = x` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3083
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3084
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3085
lemma arctan_monotone: assumes "x < y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3086
  shows "arctan x < arctan y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3087
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3088
  obtain z where "-(pi / 2) < z" and "z < pi / 2" and "tan z = x" using tan_total by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3089
  obtain w where "-(pi / 2) < w" and "w < pi / 2" and "tan w = y" using tan_total by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3090
  have "z < w" unfolding tan_monotone'[OF `-(pi / 2) < z` `z < pi / 2` `-(pi / 2) < w` `w < pi / 2`] `tan z = x` `tan w = y` using `x < y` .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3091
  thus ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3092
    unfolding `tan z = x`[symmetric] arctan_tan[OF `-(pi / 2) < z` `z < pi / 2`]
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3093
    unfolding `tan w = y`[symmetric] arctan_tan[OF `-(pi / 2) < w` `w < pi / 2`] .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3094
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3095
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3096
lemma arctan_monotone': assumes "x \<le> y" shows "arctan x \<le> arctan y"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3097
proof (cases "x = y") 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3098
  case False hence "x < y" using `x \<le> y` by auto from arctan_monotone[OF this] show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3099
qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3100
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3101
lemma arctan_minus: "arctan (- x) = - arctan x" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3102
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3103
  obtain y where "- (pi / 2) < y" and "y < pi / 2" and "tan y = x" using tan_total by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3104
  thus ?thesis unfolding `tan y = x`[symmetric] tan_minus[symmetric] using arctan_tan[of y] arctan_tan[of "-y"] by auto 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3105
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3106
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3107
lemma arctan_inverse: assumes "x \<noteq> 0" shows "arctan (1 / x) = sgn x * pi / 2 - arctan x"
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3108
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3109
  obtain y where "- (pi / 2) < y" and "y < pi / 2" and "tan y = x" using tan_total by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3110
  hence "y = arctan x" unfolding `tan y = x`[symmetric] using arctan_tan by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3111
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3112
  { fix y x :: real assume "0 < y" and "y < pi /2" and "y = arctan x" and "tan y = x" hence "- (pi / 2) < y" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3113
    have "tan y > 0" using tan_monotone'[OF _ _ `- (pi / 2) < y` `y < pi / 2`, of 0] tan_zero `0 < y` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3114
    hence "x > 0" using `tan y = x` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3115
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3116
    have "- (pi / 2) < pi / 2 - y" using `y > 0` `y < pi / 2` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3117
    moreover have "pi / 2 - y < pi / 2" using `y > 0` `y < pi / 2` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3118
    ultimately have "arctan (1 / x) = pi / 2 - y" unfolding `tan y = x`[symmetric] tan_inverse using arctan_tan by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3119
    hence "arctan (1 / x) = sgn x * pi / 2 - arctan x" unfolding `y = arctan x` real_sgn_pos[OF `x > 0`] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3120
  } note pos_y = this
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3121
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3122
  show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3123
  proof (cases "y > 0")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3124
    case True from pos_y[OF this `y < pi / 2` `y = arctan x` `tan y = x`] show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3125
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3126
    case False hence "y \<le> 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3127
    moreover have "y \<noteq> 0" 
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3128
    proof (rule ccontr)
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3129
      assume "\<not> y \<noteq> 0" hence "y = 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3130
      have "x = 0" unfolding `tan y = x`[symmetric] `y = 0` tan_zero ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3131
      thus False using `x \<noteq> 0` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3132
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3133
    ultimately have "y < 0" by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3134
    hence "0 < - y" and "-y < pi / 2" using `- (pi / 2) < y` by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3135
    moreover have "-y = arctan (-x)" unfolding arctan_minus `y = arctan x` ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3136
    moreover have "tan (-y) = -x" unfolding tan_minus `tan y = x` ..
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3137
    ultimately have "arctan (1 / -x) = sgn (-x) * pi / 2 - arctan (-x)" using pos_y by blast
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3138
    hence "arctan (- (1 / x)) = - (sgn x * pi / 2 - arctan x)" unfolding arctan_minus[of x] divide_minus_right sgn_minus by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3139
    thus ?thesis unfolding arctan_minus neg_equal_iff_equal .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3140
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3141
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3142
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3143
theorem pi_series: "pi / 4 = (\<Sum> k. (-1)^k * 1 / real (k*2+1))" (is "_ = ?SUM")
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3144
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3145
  have "pi / 4 = arctan 1" using arctan1_eq_pi4 by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3146
  also have "\<dots> = ?SUM" using arctan_series[of 1] by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3147
  finally show ?thesis by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3148
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3149
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3150
subsection {* Existence of Polar Coordinates *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3151
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3152
lemma cos_x_y_le_one: "\<bar>x / sqrt (x\<twosuperior> + y\<twosuperior>)\<bar> \<le> 1"
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3153
apply (rule power2_le_imp_le [OF _ zero_le_one])
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3154
apply (simp add: abs_divide power_divide divide_le_eq not_sum_power2_lt_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3155
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3156
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3157
lemma cos_arccos_abs: "\<bar>y\<bar> \<le> 1 \<Longrightarrow> cos (arccos y) = y"
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3158
by (simp add: abs_le_iff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3159
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3160
lemma sin_arccos_abs: "\<bar>y\<bar> \<le> 1 \<Longrightarrow> sin (arccos y) = sqrt (1 - y\<twosuperior>)"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3161
by (simp add: sin_arccos abs_le_iff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3162
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3163
lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one]
15228
4d332d10fa3d revised simprules for division
paulson
parents: 15140
diff changeset
  3164
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3165
lemmas sin_arccos_lemma1 = sin_arccos_abs [OF cos_x_y_le_one]
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3166
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  3167
lemma polar_ex1:
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3168
     "0 < y ==> \<exists>r a. x = r * cos a & y = r * sin a"
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  3169
apply (rule_tac x = "sqrt (x\<twosuperior> + y\<twosuperior>)" in exI)
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3170
apply (rule_tac x = "arccos (x / sqrt (x\<twosuperior> + y\<twosuperior>))" in exI)
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3171
apply (simp add: cos_arccos_lemma1)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3172
apply (simp add: sin_arccos_lemma1)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3173
apply (simp add: power_divide)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3174
apply (simp add: real_sqrt_mult [symmetric])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3175
apply (simp add: right_diff_distrib)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3176
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3177
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  3178
lemma polar_ex2:
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3179
     "y < 0 ==> \<exists>r a. x = r * cos a & y = r * sin a"
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3180
apply (insert polar_ex1 [where x=x and y="-y"], simp, clarify)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3181
apply (rule_tac x = r in exI)
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3182
apply (rule_tac x = "-a" in exI, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3183
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3184
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3185
lemma polar_Ex: "\<exists>r a. x = r * cos a & y = r * sin a"
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3186
apply (rule_tac x=0 and y=y in linorder_cases)
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3187
apply (erule polar_ex1)
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3188
apply (rule_tac x=x in exI, rule_tac x=0 in exI, simp)
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3189
apply (erule polar_ex2)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3190
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3191
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3192
end