src/HOL/Transcendental.thy
author hoelzl
Tue, 26 Mar 2013 12:20:59 +0100
changeset 51527 bd62e7ff103b
parent 51482 80efd8c49f52
child 51641 cd05e9fcc63d
permissions -rw-r--r--
move Ln.thy and Log.thy to Transcendental.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
     1
(*  Title:      HOL/Transcendental.thy
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
     2
    Author:     Jacques D. Fleuriot, University of Cambridge, University of Edinburgh
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
     3
    Author:     Lawrence C Paulson
51527
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
     4
    Author:     Jeremy Avigad
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
     5
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
    27
     "(\<Sum>p=0..<Suc n. (x ^ p) * y ^ (Suc n - p)) =
23082
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
33549
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
    30
         del: setsum_op_ivl_Suc)
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
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)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
    40
apply (simp add: distrib_left del: setsum_op_ivl_Suc)
34974
18b41bba42b5 new theory Algebras.thy for generic algebraic structures
haftmann
parents: 33667
diff changeset
    41
apply (subst mult_left_commute [of "x - y"])
23082
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:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
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:
44726
8478eab380e9 generalize some lemmas
huffman
parents: 44725
diff changeset
    59
  fixes x z :: "'a::real_normed_field"
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)"
44726
8478eab380e9 generalize some lemmas
huffman
parents: 44725
diff changeset
    70
    by (rule convergent_Cauchy)
20849
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   117
  "(\<Sum> i = 0 ..< 2 * n. if even i then f i else g i) =
29803
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)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   121
  have "(\<Sum> i = 0 ..< 2 * Suc n. if even i then f i else g i) =
29803
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   138
    have sum_eq: "?SUM (2 * (m div 2)) = setsum g { 0 ..< m div 2 }"
29803
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`] .
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   166
  {
29803
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
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 38642
diff changeset
   169
    have if_eq: "\<And>B T E. (if \<not> B then T else E) = (if B then E else T)" 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
   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`] .
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   172
    from this[unfolded sums_def, THEN LIMSEQ_Suc]
29803
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]
31148
7ba7c1f8bc22 Cleaned up Parity a little
nipkow
parents: 31017
diff changeset
   176
                even_Suc Suc_m1 if_eq .
29803
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"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
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>
29803
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
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"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   206
    with `a ----> 0`[THEN LIMSEQ_D]
29803
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
51477
2990382dc066 clean up lemma_nest_unique and renamed to nested_sequence_unique
hoelzl
parents: 50347
diff changeset
   212
  show ?thesis by (rule nested_sequence_unique)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   213
qed
29803
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   230
29803
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   236
    with `?f ----> l`[THEN LIMSEQ_D]
29803
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   239
    from `0 < r` `?g ----> l`[THEN LIMSEQ_D]
29803
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")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   246
        case True from even_nat_div_two_times_two[OF this]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   247
        have n_eq: "2 * (n div 2) = n" unfolding numeral_2_eq_2[symmetric] by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   248
        with `n \<ge> 2 * f_no` have "n div 2 \<ge> f_no" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   249
        from f[OF this]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   250
        show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost .
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   251
      next
35213
b9866ad4e3be fix looping call to simplifier
huffman
parents: 35038
diff changeset
   252
        case False hence "even (n - 1)" by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   253
        from even_nat_div_two_times_two[OF this]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   254
        have n_eq: "2 * ((n - 1) div 2) = n - 1" unfolding numeral_2_eq_2[symmetric] by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   255
        hence range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   256
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   257
        from n_eq `n \<ge> 2 * g_no` have "(n - 1) div 2 \<ge> g_no" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   258
        from g[OF this]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   259
        show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost range_eq .
29803
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
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44319
diff changeset
   299
    note leibniz = summable_Leibniz'[OF _ ge0, of "\<lambda>x. x", OF tendsto_minus[OF `a ----> 0`, unfolded minus_zero] monotone]
29803
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
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   307
29803
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
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44319
diff changeset
   313
    moreover have ?f and ?g using leibniz(3,5)[unfolded mult_minus_right setsum_negf move_minus, THEN tendsto_minus_cancel] 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
   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:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   341
  fixes x :: "'a::{real_normed_vector, ring_1}"
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   342
  shows "summable (%n. (diffs c)(n) * (x ^ n)) ==>
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   343
      (%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums
15546
5188ce7316b7 suminf -> \<Sum>
nipkow
parents: 15544
diff changeset
   344
         (\<Sum>n. (diffs c)(n) * (x ^ n))"
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   345
unfolding diffs_def
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   346
apply (drule summable_sums)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   347
apply (rule sums_Suc_imp, simp_all)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   348
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   349
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   350
lemma lemma_termdiff1:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   351
  fixes z :: "'a :: {monoid_mult,comm_ring}" shows
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   352
  "(\<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
   353
   (\<Sum>p=0..<m. (z ^ p) * (((z + h) ^ (m - p)) - (z ^ (m - p))))"
41550
efa734d9b221 eliminated global prems;
wenzelm
parents: 38642
diff changeset
   354
by(auto simp add: algebra_simps power_add [symmetric])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   355
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   356
lemma sumr_diff_mult_const2:
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   357
  "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
   358
by (simp add: setsum_subtractf)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   359
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   360
lemma lemma_termdiff2:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   361
  fixes h :: "'a :: {field}"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   362
  assumes h: "h \<noteq> 0" shows
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   363
  "((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
   364
   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
   365
        (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs")
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   366
apply (subgoal_tac "h * ?lhs = h * ?rhs", simp add: h)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   367
apply (simp add: right_diff_distrib diff_divide_distrib h)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   368
apply (simp add: mult_assoc [symmetric])
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   369
apply (cases "n", simp)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   370
apply (simp add: lemma_realpow_diff_sumr2 h
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   371
                 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
   372
            del: power_Suc setsum_op_ivl_Suc of_nat_Suc)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   373
apply (subst lemma_realpow_rev_sumr)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   374
apply (subst sumr_diff_mult_const2)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   375
apply simp
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   376
apply (simp only: lemma_termdiff1 setsum_right_distrib)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   377
apply (rule setsum_cong [OF refl])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
   378
apply (simp add: diff_minus [symmetric] less_iff_Suc_add)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   379
apply (clarify)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   380
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
   381
            del: setsum_op_ivl_Suc power_Suc)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   382
apply (subst mult_assoc [symmetric], subst power_add [symmetric])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   383
apply (simp add: mult_ac)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   384
done
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   385
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   386
lemma real_setsum_nat_ivl_bounded2:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34974
diff changeset
   387
  fixes K :: "'a::linordered_semidom"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   388
  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
   389
  assumes K: "0 \<le> K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   390
  shows "setsum f {0..<n-k} \<le> of_nat n * K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   391
apply (rule order_trans [OF setsum_mono])
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   392
apply (rule f, simp)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   393
apply (simp add: mult_right_mono K)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   394
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   395
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   396
lemma lemma_termdiff3:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   397
  fixes h z :: "'a::{real_normed_field}"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   398
  assumes 1: "h \<noteq> 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   399
  assumes 2: "norm z \<le> K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   400
  assumes 3: "norm (z + h) \<le> K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   401
  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
   402
          \<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
   403
proof -
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   404
  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
   405
        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
   406
          (z + h) ^ q * z ^ (n - 2 - q)) * norm h"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   407
    apply (subst lemma_termdiff2 [OF 1])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   408
    apply (subst norm_mult)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   409
    apply (rule mult_commute)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   410
    done
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   411
  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
   412
  proof (rule mult_right_mono [OF _ norm_ge_zero])
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   413
    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
   414
    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
   415
      apply (erule subst)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   416
      apply (simp only: norm_mult norm_power power_add)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   417
      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
   418
      done
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   419
    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
   420
              (z + h) ^ q * z ^ (n - 2 - q))
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   421
          \<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
   422
      apply (intro
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   423
         order_trans [OF norm_setsum]
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   424
         real_setsum_nat_ivl_bounded2
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   425
         mult_nonneg_nonneg
47489
04e7d09ade7a tuned some proofs;
huffman
parents: 47108
diff changeset
   426
         of_nat_0_le_iff
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   427
         zero_le_power K)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   428
      apply (rule le_Kn, simp)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   429
      done
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   430
  qed
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   431
  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
   432
    by (simp only: mult_assoc)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   433
  finally show ?thesis .
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   434
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   435
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   436
lemma lemma_termdiff4:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   437
  fixes f :: "'a::{real_normed_field} \<Rightarrow>
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   438
              'b::real_normed_vector"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   439
  assumes k: "0 < (k::real)"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   440
  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
   441
  shows "f -- 0 --> 0"
31338
d41a8ba25b67 generalize constants from Lim.thy to class metric_space
huffman
parents: 31271
diff changeset
   442
unfolding LIM_eq diff_0_right
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   443
proof (safe)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   444
  let ?h = "of_real (k / 2)::'a"
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   445
  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
   446
  hence "norm (f ?h) \<le> K * norm ?h" by (rule le)
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   447
  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
   448
  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
   449
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   450
  fix r::real assume r: "0 < r"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   451
  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
   452
  proof (cases)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   453
    assume "K = 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   454
    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
   455
      by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   456
    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
   457
  next
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   458
    assume K_neq_zero: "K \<noteq> 0"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   459
    with zero_le_K have K: "0 < K" by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   460
    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
   461
    proof (rule exI, safe)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   462
      from k r K show "0 < min k (r * inverse K / 2)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   463
        by (simp add: mult_pos_pos positive_imp_inverse_positive)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   464
    next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   465
      fix x::'a
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   466
      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
   467
      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
   468
        by simp_all
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   469
      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
   470
      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
   471
        by (rule mult_strict_left_mono)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   472
      also have "\<dots> = r / 2"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   473
        using K_neq_zero by simp
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   474
      also have "r / 2 < r"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   475
        using r by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   476
      finally show "norm (f x) < r" .
20860
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
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   479
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   480
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
   481
lemma lemma_termdiff5:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   482
  fixes g :: "'a::{real_normed_field} \<Rightarrow>
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   483
              nat \<Rightarrow> 'b::banach"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   484
  assumes k: "0 < (k::real)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   485
  assumes f: "summable f"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   486
  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
   487
  shows "(\<lambda>h. suminf (g h)) -- 0 --> 0"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   488
proof (rule lemma_termdiff4 [OF k])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   489
  fix h::'a assume "h \<noteq> 0" and "norm h < k"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   490
  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
   491
    by (simp add: le)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   492
  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
   493
    by simp
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   494
  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
   495
    by (rule summable_mult2)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   496
  ultimately have C: "summable (\<lambda>n. norm (g h n))"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   497
    by (rule summable_comparison_test)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   498
  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
   499
    by (rule summable_norm)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   500
  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
   501
    by (rule summable_le)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   502
  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
   503
    by (rule suminf_mult2 [symmetric])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   504
  finally show "norm (suminf (g h)) \<le> suminf f * norm h" .
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   505
qed
15077
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
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   508
text{* FIXME: Long proofs*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   509
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   510
lemma termdiffs_aux:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   511
  fixes x :: "'a::{real_normed_field,banach}"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   512
  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
   513
  assumes 2: "norm x < norm K"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   514
  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
   515
             - of_nat n * x ^ (n - Suc 0))) -- 0 --> 0"
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   516
proof -
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   517
  from dense [OF 2]
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   518
  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
   519
  from norm_ge_zero r1 have r: "0 < r"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   520
    by (rule order_le_less_trans)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   521
  hence r_neq_0: "r \<noteq> 0" by simp
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   522
  show ?thesis
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   523
  proof (rule lemma_termdiff5)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   524
    show "0 < r - norm x" using r1 by simp
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   525
  next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   526
    from r r2 have "norm (of_real r::'a) < norm K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   527
      by simp
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   528
    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
   529
      by (rule powser_insidea)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   530
    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
   531
      using r
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   532
      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
   533
    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
   534
      by (rule diffs_equiv [THEN sums_summable])
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   535
    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
   536
      = (\<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
   537
      apply (rule ext)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   538
      apply (simp add: diffs_def)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   539
      apply (case_tac n, simp_all add: r_neq_0)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   540
      done
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   541
    finally have "summable
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   542
      (\<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
   543
      by (rule diffs_equiv [THEN sums_summable])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   544
    also have
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   545
      "(\<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
   546
           r ^ (n - Suc 0)) =
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   547
       (\<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
   548
      apply (rule ext)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   549
      apply (case_tac "n", simp)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   550
      apply (case_tac "nat", simp)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   551
      apply (simp add: r_neq_0)
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   552
      done
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   553
    finally show
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   554
      "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
   555
  next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   556
    fix h::'a and n::nat
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   557
    assume h: "h \<noteq> 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   558
    assume "norm h < r - norm x"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   559
    hence "norm x + norm h < r" by simp
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   560
    with norm_triangle_ineq have xh: "norm (x + h) < r"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   561
      by (rule order_le_less_trans)
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   562
    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
   563
          \<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
   564
      apply (simp only: norm_mult mult_assoc)
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   565
      apply (rule mult_left_mono [OF _ norm_ge_zero])
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   566
      apply (simp (no_asm) add: mult_assoc [symmetric])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   567
      apply (rule lemma_termdiff3)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   568
      apply (rule h)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   569
      apply (rule r1 [THEN order_less_imp_le])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   570
      apply (rule xh [THEN order_less_imp_le])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   571
      done
20849
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   572
  qed
389cd9c8cfe1 rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents: 20692
diff changeset
   573
qed
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
   574
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   575
lemma termdiffs:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   576
  fixes K x :: "'a::{real_normed_field,banach}"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   577
  assumes 1: "summable (\<lambda>n. c n * K ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   578
  assumes 2: "summable (\<lambda>n. (diffs c) n * K ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   579
  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
   580
  assumes 4: "norm x < norm K"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   581
  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
   582
unfolding deriv_def
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   583
proof (rule LIM_zero_cancel)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   584
  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
   585
            - suminf (\<lambda>n. diffs c n * x ^ n)) -- 0 --> 0"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   586
  proof (rule LIM_equal2)
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   587
    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
   588
  next
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   589
    fix h :: 'a
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   590
    assume "h \<noteq> 0"
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   591
    assume "norm (h - 0) < norm K - norm x"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   592
    hence "norm x + norm h < norm K" by simp
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   593
    hence 5: "norm (x + h) < norm K"
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   594
      by (rule norm_triangle_ineq [THEN order_le_less_trans])
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   595
    have A: "summable (\<lambda>n. c n * x ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   596
      by (rule powser_inside [OF 1 4])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   597
    have B: "summable (\<lambda>n. c n * (x + h) ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   598
      by (rule powser_inside [OF 1 5])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   599
    have C: "summable (\<lambda>n. diffs c n * x ^ n)"
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   600
      by (rule powser_inside [OF 2 4])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   601
    show "((\<Sum>n. c n * (x + h) ^ n) - (\<Sum>n. c n * x ^ n)) / h
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   602
             - (\<Sum>n. diffs c n * x ^ n) =
23082
ffef77eed382 generalize powerseries and termdiffs lemmas using axclasses
huffman
parents: 23069
diff changeset
   603
          (\<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
   604
      apply (subst sums_unique [OF diffs_equiv [OF C]])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   605
      apply (subst suminf_diff [OF B A])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   606
      apply (subst suminf_divide [symmetric])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   607
      apply (rule summable_diff [OF B A])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   608
      apply (subst suminf_diff)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   609
      apply (rule summable_divide)
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   610
      apply (rule summable_diff [OF B A])
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   611
      apply (rule sums_summable [OF diffs_equiv [OF C]])
29163
e72d07a878f8 clean up some proofs; remove unused lemmas
huffman
parents: 28952
diff changeset
   612
      apply (rule arg_cong [where f="suminf"], rule ext)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
   613
      apply (simp add: algebra_simps)
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   614
      done
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   615
  next
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   616
    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
   617
               of_nat n * x ^ (n - Suc 0))) -- 0 --> 0"
20860
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   618
        by (rule termdiffs_aux [OF 3 4])
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
qed
1a8efd618190 reorganize and speed up termdiffs proofs
huffman
parents: 20849
diff changeset
   621
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   622
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   623
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
   624
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   625
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
   626
  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
   627
  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
   628
  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
   629
  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
   630
  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
   631
  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
   632
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
   633
  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
   634
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   635
  obtain N_L where N_L: "\<And> n. N_L \<le> n \<Longrightarrow> \<bar> \<Sum> i. L (i + n) \<bar> < r/3"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   636
    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
   637
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   638
  obtain N_f' where N_f': "\<And> n. N_f' \<le> n \<Longrightarrow> \<bar> \<Sum> i. f' x0 (i + n) \<bar> < r/3"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   639
    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
   640
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   641
  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
   642
  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
   643
    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
   644
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   645
  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
   646
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   647
  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
   648
  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
   649
  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
   650
  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
   651
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   652
  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
   653
  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
   654
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   655
  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
   656
  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
   657
    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
   658
    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
   659
      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
   660
      then obtain n where "x = ?s n" and "n \<in> {0..<?N}" using image_iff[THEN iffD1] by blast
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   661
      from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF `0 < ?r`, unfolded real_norm_def]
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   662
      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
   663
      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
   664
      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
   665
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   666
  qed auto
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
  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
   669
  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
   670
    by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   671
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   672
  { 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
   673
    hence x_in_I: "x0 + x \<in> { a <..< b }" using S_a S_b by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   674
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   675
    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
   676
    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
   677
    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
   678
    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
   679
    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
   680
    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
   681
    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
   682
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   683
    { fix n
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   684
      have "\<bar> ?diff (n + ?N) x \<bar> \<le> L (n + ?N) * \<bar> (x0 + x) - x0 \<bar> / \<bar> x \<bar>"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   685
        using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] unfolding abs_divide .
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   686
      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
   687
    } 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
   688
    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
   689
    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
   690
    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
   691
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   692
    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
   693
    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
   694
    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
   695
      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
   696
      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
   697
      also have "S \<le> S'" using `S \<le> S'` .
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   698
      also have "S' \<le> ?s n" unfolding S'_def
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   699
      proof (rule Min_le_iff[THEN iffD2])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   700
        have "?s n \<in> (?s ` {0..<?N}) \<and> ?s n \<le> ?s n" using `n \<in> { 0 ..< ?N}` by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   701
        thus "\<exists> a \<in> (?s ` {0..<?N}). a \<le> ?s n" by blast
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   702
      qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   703
      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
   704
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   705
      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
   706
      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
   707
      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
   708
      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
   709
    qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   710
    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
   711
    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
   712
    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
   713
    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
   714
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   715
    from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]]
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   716
    have "\<bar> (suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0) \<bar> =
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   717
                    \<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
   718
    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
   719
    also have "\<dots> \<le> ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   720
    also have "\<dots> < r /3 + r/3 + r/3"
36842
99745a4b9cc9 fix some linarith_split_limit warnings
huffman
parents: 36824
diff changeset
   721
      using `?diff_part < r/3` `?L_part \<le> r/3` and `?f'_part < r/3`
99745a4b9cc9 fix some linarith_split_limit warnings
huffman
parents: 36824
diff changeset
   722
      by (rule add_strict_mono [OF add_less_le_mono])
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   723
    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
   724
      by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   725
  } thus "\<exists> s > 0. \<forall> x. x \<noteq> 0 \<and> norm (x - 0) < s \<longrightarrow>
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   726
      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
   727
    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
   728
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   729
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   730
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
   731
  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
   732
  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
   733
  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
   734
  (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
   735
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   736
  { 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
   737
    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
   738
    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
   739
    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
   740
      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
   741
      proof -
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   742
        have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using `0 < R'` `0 < R` `R' < R` by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   743
        hence in_Rball: "(R' + R) / 2 \<in> {-R <..< R}" using `R' < R` by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   744
        have "norm R' < norm ((R' + R) / 2)" using `0 < R'` `0 < R` `R' < R` by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   745
        from powser_insidea[OF converges[OF in_Rball] this] show ?thesis 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
   746
      qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   747
      { fix n x y assume "x \<in> {-R' <..< R'}" and "y \<in> {-R' <..< R'}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   748
        show "\<bar>?f x n - ?f y n\<bar> \<le> \<bar>f n * real (Suc n) * R'^n\<bar> * \<bar>x-y\<bar>"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   749
        proof -
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   750
          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>"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   751
            unfolding right_diff_distrib[symmetric] lemma_realpow_diff_sumr2 abs_mult by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   752
          also have "\<dots> \<le> (\<bar>f n\<bar> * \<bar>x-y\<bar>) * (\<bar>real (Suc n)\<bar> * \<bar>R' ^ n\<bar>)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   753
          proof (rule mult_left_mono)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   754
            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)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   755
            also have "\<dots> \<le> (\<Sum>p = 0..<Suc n. R' ^ n)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   756
            proof (rule setsum_mono)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   757
              fix p assume "p \<in> {0..<Suc n}" hence "p \<le> n" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   758
              { fix n fix x :: real assume "x \<in> {-R'<..<R'}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   759
                hence "\<bar>x\<bar> \<le> R'"  by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   760
                hence "\<bar>x^n\<bar> \<le> R'^n" unfolding power_abs by (rule power_mono, auto)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   761
              } from mult_mono[OF this[OF `x \<in> {-R'<..<R'}`, of p] this[OF `y \<in> {-R'<..<R'}`, of "n-p"]] `0 < R'`
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   762
              have "\<bar>x^p * y^(n-p)\<bar> \<le> R'^p * R'^(n-p)" unfolding abs_mult by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   763
              thus "\<bar>x^p * y^(n-p)\<bar> \<le> R'^n" unfolding power_add[symmetric] using `p \<le> n` by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   764
            qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   765
            also have "\<dots> = real (Suc n) * R' ^ n" unfolding setsum_constant card_atLeastLessThan real_of_nat_def by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   766
            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'`]]] .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   767
            show "0 \<le> \<bar>f n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult[symmetric] by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   768
          qed
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
   769
          also have "\<dots> = \<bar>f n * real (Suc n) * R' ^ n\<bar> * \<bar>x - y\<bar>" unfolding abs_mult mult_assoc[symmetric] by algebra
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   770
          finally show ?thesis .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   771
        qed }
31881
eba74a5790d2 use DERIV_intros
hoelzl
parents: 31880
diff changeset
   772
      { fix n show "DERIV (\<lambda> x. ?f x n) x0 :> (?f' x0 n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   773
          by (auto intro!: DERIV_intros simp del: power_Suc) }
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   774
      { fix x assume "x \<in> {-R' <..< R'}" hence "R' \<in> {-R <..< R}" and "norm x < norm R'" using assms `R' < R` by auto
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   775
        have "summable (\<lambda> n. f n * x^n)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   776
        proof (rule summable_le2[THEN conjunct1, OF _ powser_insidea[OF converges[OF `R' \<in> {-R <..< R}`] `norm x < norm R'`]], rule allI)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   777
          fix n
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   778
          have le: "\<bar>f n\<bar> * 1 \<le> \<bar>f n\<bar> * real (Suc n)" by (rule mult_left_mono, auto)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   779
          show "\<bar>f n * x ^ n\<bar> \<le> norm (f n * real (Suc n) * x ^ n)" unfolding real_norm_def abs_mult
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   780
            by (rule mult_right_mono, auto simp add: le[unfolded mult_1_right])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   781
        qed
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
   782
        from this[THEN summable_mult2[where c=x], unfolded mult_assoc, unfolded mult_commute]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
   783
        show "summable (?f x)" 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
   784
      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
   785
      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
   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
  } 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
   788
  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
   789
  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
   790
  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
   791
  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
   792
    case True
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   793
    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
   794
    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
   795
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   796
    case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   797
    have "- ?R < 0" using assms by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   798
    also have "\<dots> \<le> x0" using False 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
   799
    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
   800
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   801
  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
   802
  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
   803
  show ?thesis .
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
   804
qed
29695
171146a93106 Added real related theorems from Fact.thy
chaieb
parents: 29667
diff changeset
   805
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
   806
subsection {* Exponential Function *}
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   807
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
   808
definition exp :: "'a \<Rightarrow> 'a::{real_normed_field,banach}" where
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
   809
  "exp = (\<lambda>x. \<Sum>n. x ^ n /\<^sub>R real (fact n))"
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   810
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   811
lemma summable_exp_generic:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   812
  fixes x :: "'a::{real_normed_algebra_1,banach}"
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   813
  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
   814
  shows "summable S"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   815
proof -
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   816
  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
   817
    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
   818
  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
   819
    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
   820
  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
   821
    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
   822
  from r1 show ?thesis
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   823
  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
   824
    fix n :: nat
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   825
    assume n: "N \<le> n"
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   826
    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
   827
      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
   828
    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
   829
      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
   830
    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
   831
      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
   832
    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
   833
      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
   834
    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
   835
      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
   836
    thus "norm (S (Suc n)) \<le> r * norm (S n)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35213
diff changeset
   837
      by (simp add: S_Suc inverse_eq_divide)
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   838
  qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   839
qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   840
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   841
lemma summable_norm_exp:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   842
  fixes x :: "'a::{real_normed_algebra_1,banach}"
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   843
  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
   844
proof (rule summable_norm_comparison_test [OF exI, rule_format])
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   845
  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
   846
    by (rule summable_exp_generic)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   847
next
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   848
  fix n show "norm (x ^ n /\<^sub>R real (fact n)) \<le> norm x ^ n /\<^sub>R real (fact n)"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35213
diff changeset
   849
    by (simp add: norm_power_ineq)
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   850
qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   851
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   852
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
   853
by (insert summable_exp_generic [where x=x], simp)
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   854
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   855
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
   856
unfolding exp_def by (rule summable_exp_generic [THEN summable_sums])
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   857
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
   858
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
   859
lemma exp_fdiffs:
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   860
      "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
   861
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
   862
         del: mult_Suc of_nat_Suc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   863
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   864
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
   865
by (simp add: diffs_def)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   866
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   867
lemma DERIV_exp [simp]: "DERIV exp x :> exp(x)"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
   868
unfolding exp_def scaleR_conv_of_real
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
   869
apply (rule DERIV_cong)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
   870
apply (rule termdiffs [where K="of_real (1 + norm x)"])
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   871
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
   872
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
   873
apply (simp del: of_real_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   874
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   875
51527
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
   876
declare DERIV_exp[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
   877
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   878
lemma isCont_exp: "isCont exp x"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   879
  by (rule DERIV_exp [THEN DERIV_isCont])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   880
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   881
lemma isCont_exp' [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. exp (f x)) a"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   882
  by (rule isCont_o2 [OF _ isCont_exp])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   883
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   884
lemma tendsto_exp [tendsto_intros]:
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   885
  "(f ---> a) F \<Longrightarrow> ((\<lambda>x. exp (f x)) ---> exp a) F"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
   886
  by (rule isCont_tendsto_compose [OF isCont_exp])
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
   887
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
   888
lemma continuous_exp [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. exp (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
   889
  unfolding continuous_def by (rule tendsto_exp)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
   890
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
   891
lemma continuous_on_exp [continuous_on_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. exp (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
   892
  unfolding continuous_on_def by (auto intro: tendsto_exp)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
   893
29167
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   894
subsubsection {* Properties of the Exponential Function *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   895
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   896
lemma powser_zero:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   897
  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
   898
  shows "(\<Sum>n. f n * 0 ^ n) = f 0"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   899
proof -
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   900
  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
   901
    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
   902
  thus ?thesis unfolding One_nat_def by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   903
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   904
23278
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   905
lemma exp_zero [simp]: "exp 0 = 1"
375335bf619f clean up proofs of exp_zero, sin_zero, cos_zero
huffman
parents: 23255
diff changeset
   906
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
   907
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   908
lemma setsum_cl_ivl_Suc2:
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   909
  "(\<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
   910
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
   911
         del: setsum_cl_ivl_Suc)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   912
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   913
lemma exp_series_add:
31017
2c227493ea56 stripped class recpower further
haftmann
parents: 30273
diff changeset
   914
  fixes x y :: "'a::{real_field}"
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   915
  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
   916
  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
   917
proof (induct n)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   918
  case 0
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   919
  show ?case
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   920
    unfolding S_def by simp
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   921
next
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   922
  case (Suc n)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   923
  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
   924
    unfolding S_def by (simp del: mult_Suc)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   925
  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
   926
    by simp
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   927
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   928
  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
   929
    by (simp only: times_S)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   930
  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
   931
    by (simp only: Suc)
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   932
  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
   933
                + y * (\<Sum>i=0..n. S x i * S y (n-i))"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
   934
    by (rule distrib_right)
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   935
  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
   936
                + (\<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
   937
    by (simp only: setsum_right_distrib mult_ac)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   938
  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
   939
                + (\<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
   940
    by (simp add: times_S Suc_diff_le)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   941
  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
   942
             (\<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
   943
    by (subst setsum_cl_ivl_Suc2, simp)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   944
  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
   945
             (\<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
   946
    by (subst setsum_cl_ivl_Suc, simp)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   947
  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
   948
             (\<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
   949
             (\<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
   950
    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
   951
              real_of_nat_add [symmetric], simp)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 23477
diff changeset
   952
  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
   953
    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
   954
  finally show
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   955
    "S (x + y) (Suc n) = (\<Sum>i=0..Suc n. S x i * S y (Suc n - i))"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35213
diff changeset
   956
    by (simp del: setsum_cl_ivl_Suc)
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   957
qed
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   958
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   959
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
   960
unfolding exp_def
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   961
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
   962
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   963
lemma mult_exp_exp: "exp x * exp y = exp (x + y)"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   964
by (rule exp_add [symmetric])
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   965
23241
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   966
lemma exp_of_real: "exp (of_real x) = of_real (exp x)"
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   967
unfolding exp_def
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 43335
diff changeset
   968
apply (subst suminf_of_real)
23241
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   969
apply (rule summable_exp_generic)
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   970
apply (simp add: scaleR_conv_of_real)
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   971
done
5f12b40a95bf add lemma exp_of_real
huffman
parents: 23177
diff changeset
   972
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   973
lemma exp_not_eq_zero [simp]: "exp x \<noteq> 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   974
proof
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   975
  have "exp x * exp (- x) = 1" by (simp add: mult_exp_exp)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   976
  also assume "exp x = 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   977
  finally show "False" by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   978
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   979
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   980
lemma exp_minus: "exp (- x) = inverse (exp x)"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   981
by (rule inverse_unique [symmetric], simp add: mult_exp_exp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   982
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   983
lemma exp_diff: "exp (x - y) = exp x / exp y"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   984
  unfolding diff_minus divide_inverse
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   985
  by (simp add: exp_add exp_minus)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
   986
29167
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   987
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   988
subsubsection {* Properties of the Exponential Function on Reals *}
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   989
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
   990
text {* Comparisons of @{term "exp x"} with zero. *}
29167
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   991
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   992
text{*Proof: because every exponential can be seen as a square.*}
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   993
lemma exp_ge_zero [simp]: "0 \<le> exp (x::real)"
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   994
proof -
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   995
  have "0 \<le> exp (x/2) * exp (x/2)" by simp
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   996
  thus ?thesis by (simp add: exp_add [symmetric])
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   997
qed
37a952bb9ebc rearranged subsections; cleaned up some proofs
huffman
parents: 29166
diff changeset
   998
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
   999
lemma exp_gt_zero [simp]: "0 < exp (x::real)"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1000
by (simp add: order_less_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1001
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1002
lemma not_exp_less_zero [simp]: "\<not> exp (x::real) < 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1003
by (simp add: not_less)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1004
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1005
lemma not_exp_le_zero [simp]: "\<not> exp (x::real) \<le> 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1006
by (simp add: not_le)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1007
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1008
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
  1009
by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1010
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1011
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
  1012
apply (induct "n")
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  1013
apply (auto simp add: real_of_nat_Suc distrib_left exp_add mult_commute)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1014
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1015
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1016
text {* Strict monotonicity of exponential. *}
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1017
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1018
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
  1019
apply (drule order_le_imp_less_or_eq, auto)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1020
apply (simp add: exp_def)
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
  1021
apply (rule order_trans)
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1022
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
  1023
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
  1024
done
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1025
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1026
lemma exp_gt_one: "0 < (x::real) \<Longrightarrow> 1 < exp x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1027
proof -
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1028
  assume x: "0 < x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1029
  hence "1 < 1 + x" by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1030
  also from x have "1 + x \<le> exp x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1031
    by (simp add: exp_ge_add_one_self_aux)
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1032
  finally show ?thesis .
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1033
qed
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1034
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1035
lemma exp_less_mono:
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1036
  fixes x y :: real
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1037
  assumes "x < y" shows "exp x < exp y"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1038
proof -
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1039
  from `x < y` have "0 < y - x" by simp
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1040
  hence "1 < exp (y - x)" by (rule exp_gt_one)
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  1041
  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
  1042
  thus "exp x < exp y" by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1043
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1044
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1045
lemma exp_less_cancel: "exp (x::real) < exp y ==> x < y"
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1046
apply (simp add: linorder_not_le [symmetric])
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1047
apply (auto simp add: order_le_less exp_less_mono)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1048
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1049
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1050
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
  1051
by (auto intro: exp_less_mono exp_less_cancel)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1052
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1053
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
  1054
by (auto simp add: linorder_not_less [symmetric])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1055
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1056
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
  1057
by (simp add: order_eq_iff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1058
29170
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1059
text {* Comparisons of @{term "exp x"} with one. *}
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1060
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1061
lemma one_less_exp_iff [simp]: "1 < exp (x::real) \<longleftrightarrow> 0 < x"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1062
  using exp_less_cancel_iff [where x=0 and y=x] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1063
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1064
lemma exp_less_one_iff [simp]: "exp (x::real) < 1 \<longleftrightarrow> x < 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1065
  using exp_less_cancel_iff [where x=x and y=0] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1066
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1067
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
  1068
  using exp_le_cancel_iff [where x=0 and y=x] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1069
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1070
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
  1071
  using exp_le_cancel_iff [where x=x and y=0] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1072
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1073
lemma exp_eq_one_iff [simp]: "exp (x::real) = 1 \<longleftrightarrow> x = 0"
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1074
  using exp_inj_iff [where x=x and y=0] by simp
dad3933c88dd clean up lemmas about exp
huffman
parents: 29167
diff changeset
  1075
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1076
lemma lemma_exp_total: "1 \<le> y ==> \<exists>x. 0 \<le> x & x \<le> y - 1 & exp(x::real) = y"
44755
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1077
proof (rule IVT)
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1078
  assume "1 \<le> y"
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1079
  hence "0 \<le> y - 1" by simp
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1080
  hence "1 + (y - 1) \<le> exp (y - 1)" by (rule exp_ge_add_one_self_aux)
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1081
  thus "y \<le> exp (y - 1)" by simp
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1082
qed (simp_all add: le_diff_eq)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1083
23115
4615b2078592 generalized exp to work over any complete field; new proof of exp_add
huffman
parents: 23112
diff changeset
  1084
lemma exp_total: "0 < (y::real) ==> \<exists>x. exp x = y"
44755
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1085
proof (rule linorder_le_cases [of 1 y])
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1086
  assume "1 \<le> y" thus "\<exists>x. exp x = y"
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1087
    by (fast dest: lemma_exp_total)
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1088
next
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1089
  assume "0 < y" and "y \<le> 1"
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1090
  hence "1 \<le> inverse y" by (simp add: one_le_inverse_iff)
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1091
  then obtain x where "exp x = inverse y" by (fast dest: lemma_exp_total)
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1092
  hence "exp (- x) = y" by (simp add: exp_minus)
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1093
  thus "\<exists>x. exp x = y" ..
257ac9da021f convert some proofs to Isar-style
huffman
parents: 44746
diff changeset
  1094
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1095
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1096
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  1097
subsection {* Natural Logarithm *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1098
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1099
definition ln :: "real \<Rightarrow> real" where
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1100
  "ln x = (THE u. exp u = x)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1101
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  1102
lemma ln_exp [simp]: "ln (exp x) = x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1103
  by (simp add: ln_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1104
22654
c2b6b5a9e136 new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents: 22653
diff changeset
  1105
lemma exp_ln [simp]: "0 < x \<Longrightarrow> exp (ln x) = x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1106
  by (auto dest: exp_total)
22654
c2b6b5a9e136 new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents: 22653
diff changeset
  1107
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1108
lemma exp_ln_iff [simp]: "exp (ln x) = x \<longleftrightarrow> 0 < x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1109
  by (metis exp_gt_zero exp_ln)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1110
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1111
lemma ln_unique: "exp y = x \<Longrightarrow> ln x = y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1112
  by (erule subst, rule ln_exp)
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1113
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1114
lemma ln_one [simp]: "ln 1 = 0"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1115
  by (rule ln_unique, simp)
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1116
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1117
lemma ln_mult: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln (x * y) = ln x + ln y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1118
  by (rule ln_unique, simp add: exp_add)
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1119
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1120
lemma ln_inverse: "0 < x \<Longrightarrow> ln (inverse x) = - ln x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1121
  by (rule ln_unique, simp add: exp_minus)
29171
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_div: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln (x / y) = ln x - ln y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1124
  by (rule ln_unique, simp add: exp_diff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1125
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1126
lemma ln_realpow: "0 < x \<Longrightarrow> ln (x ^ n) = real n * ln x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1127
  by (rule ln_unique, simp add: exp_real_of_nat_mult)
29171
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_less_cancel_iff [simp]: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln x < ln y \<longleftrightarrow> x < y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1130
  by (subst exp_less_cancel_iff [symmetric], simp)
29171
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_le_cancel_iff [simp]: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln x \<le> ln y \<longleftrightarrow> x \<le> y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1133
  by (simp add: linorder_not_less [symmetric])
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1134
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1135
lemma ln_inj_iff [simp]: "\<lbrakk>0 < x; 0 < y\<rbrakk> \<Longrightarrow> ln x = ln y \<longleftrightarrow> x = y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1136
  by (simp add: order_eq_iff)
29171
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_add_one_self_le_self [simp]: "0 \<le> x \<Longrightarrow> ln (1 + x) \<le> x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1139
  apply (rule exp_le_cancel_iff [THEN iffD1])
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1140
  apply (simp add: exp_ge_add_one_self_aux)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1141
  done
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1142
29171
5eff800a695f clean up lemmas about ln
huffman
parents: 29170
diff changeset
  1143
lemma ln_less_self [simp]: "0 < x \<Longrightarrow> ln x < x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1144
  by (rule order_less_le_trans [where y="ln (1 + x)"]) simp_all
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1145
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1146
lemma ln_ge_zero [simp]: "1 \<le> x \<Longrightarrow> 0 \<le> ln x"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1147
  using ln_le_cancel_iff [of 1 x] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1148
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1149
lemma ln_ge_zero_imp_ge_one: "\<lbrakk>0 \<le> ln x; 0 < x\<rbrakk> \<Longrightarrow> 1 \<le> x"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1150
  using ln_le_cancel_iff [of 1 x] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1151
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1152
lemma ln_ge_zero_iff [simp]: "0 < x \<Longrightarrow> (0 \<le> ln x) = (1 \<le> x)"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1153
  using ln_le_cancel_iff [of 1 x] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1154
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1155
lemma ln_less_zero_iff [simp]: "0 < x \<Longrightarrow> (ln x < 0) = (x < 1)"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1156
  using ln_less_cancel_iff [of x 1] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1157
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1158
lemma ln_gt_zero: "1 < x \<Longrightarrow> 0 < ln x"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1159
  using ln_less_cancel_iff [of 1 x] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1160
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1161
lemma ln_gt_zero_imp_gt_one: "\<lbrakk>0 < ln x; 0 < x\<rbrakk> \<Longrightarrow> 1 < x"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1162
  using ln_less_cancel_iff [of 1 x] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1163
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1164
lemma ln_gt_zero_iff [simp]: "0 < x \<Longrightarrow> (0 < ln x) = (1 < x)"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1165
  using ln_less_cancel_iff [of 1 x] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1166
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1167
lemma ln_eq_zero_iff [simp]: "0 < x \<Longrightarrow> (ln x = 0) = (x = 1)"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1168
  using ln_inj_iff [of x 1] by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1169
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1170
lemma ln_less_zero: "\<lbrakk>0 < x; x < 1\<rbrakk> \<Longrightarrow> ln x < 0"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1171
  by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1172
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1173
lemma isCont_ln: "0 < x \<Longrightarrow> isCont ln x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1174
  apply (subgoal_tac "isCont ln (exp (ln x))", simp)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1175
  apply (rule isCont_inverse_function [where f=exp], simp_all)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1176
  done
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1177
45915
0e5a87b772f9 tendsto lemmas for ln and powr
huffman
parents: 45309
diff changeset
  1178
lemma tendsto_ln [tendsto_intros]:
0e5a87b772f9 tendsto lemmas for ln and powr
huffman
parents: 45309
diff changeset
  1179
  "\<lbrakk>(f ---> a) F; 0 < a\<rbrakk> \<Longrightarrow> ((\<lambda>x. ln (f x)) ---> ln a) F"
0e5a87b772f9 tendsto lemmas for ln and powr
huffman
parents: 45309
diff changeset
  1180
  by (rule isCont_tendsto_compose [OF isCont_ln])
0e5a87b772f9 tendsto lemmas for ln and powr
huffman
parents: 45309
diff changeset
  1181
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1182
lemma continuous_ln:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1183
  "continuous F f \<Longrightarrow> 0 < f (Lim F (\<lambda>x. x)) \<Longrightarrow> continuous F (\<lambda>x. ln (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1184
  unfolding continuous_def by (rule tendsto_ln)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1185
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1186
lemma isCont_ln' [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1187
  "continuous (at x) f \<Longrightarrow> 0 < f x \<Longrightarrow> continuous (at x) (\<lambda>x. ln (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1188
  unfolding continuous_at by (rule tendsto_ln)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1189
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1190
lemma continuous_within_ln [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1191
  "continuous (at x within s) f \<Longrightarrow> 0 < f x \<Longrightarrow> continuous (at x within s) (\<lambda>x. ln (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1192
  unfolding continuous_within by (rule tendsto_ln)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1193
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1194
lemma continuous_on_ln [continuous_on_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1195
  "continuous_on s f \<Longrightarrow> (\<forall>x\<in>s. 0 < f x) \<Longrightarrow> continuous_on s (\<lambda>x. ln (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1196
  unfolding continuous_on_def by (auto intro: tendsto_ln)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  1197
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1198
lemma DERIV_ln: "0 < x \<Longrightarrow> DERIV ln x :> inverse x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1199
  apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"])
44317
b7e9fa025f15 remove redundant lemma lemma_DERIV_subst in favor of DERIV_cong
huffman
parents: 44316
diff changeset
  1200
  apply (erule DERIV_cong [OF DERIV_exp exp_ln])
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1201
  apply (simp_all add: abs_if isCont_ln)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  1202
  done
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  1203
33667
958dc9f03611 A little rationalisation
paulson
parents: 33549
diff changeset
  1204
lemma DERIV_ln_divide: "0 < x ==> DERIV ln x :> 1 / x"
958dc9f03611 A little rationalisation
paulson
parents: 33549
diff changeset
  1205
  by (rule DERIV_ln[THEN DERIV_cong], simp, simp add: divide_inverse)
958dc9f03611 A little rationalisation
paulson
parents: 33549
diff changeset
  1206
51527
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1207
declare DERIV_ln_divide[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1208
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1209
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
  1210
  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
  1211
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1212
  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
  1213
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1214
  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
  1215
  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
  1216
    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
  1217
    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
  1218
    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
  1219
    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
  1220
    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)
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
  1221
    finally have "DERIV ln x :> suminf (?f' x)" using DERIV_ln[OF `0 < x`] unfolding divide_inverse 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
  1222
    moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1223
    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
  1224
    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
  1225
    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
  1226
      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
  1227
      { fix x :: real assume "x \<in> {- 1<..<1}" hence "norm (-x) < 1" by auto
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  1228
        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
  1229
          unfolding One_nat_def
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35213
diff changeset
  1230
          by (auto simp add: power_mult_distrib[symmetric] summable_geometric[OF `norm (-x) < 1`])
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1231
      }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  1232
    qed
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  1233
    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
  1234
    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
  1235
    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
  1236
      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
  1237
    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
  1238
  qed (auto simp add: assms)
44289
d81d09cdab9c optimize some proofs
huffman
parents: 44282
diff changeset
  1239
  thus ?thesis 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
  1240
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  1241
50326
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1242
lemma exp_first_two_terms: "exp x = 1 + x + (\<Sum> n. inverse(fact (n+2)) * (x ^ (n+2)))"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1243
proof -
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1244
  have "exp x = suminf (%n. inverse(fact n) * (x ^ n))"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1245
    by (simp add: exp_def)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1246
  also from summable_exp have "... = (\<Sum> n::nat = 0 ..< 2. inverse(fact n) * (x ^ n)) +
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1247
      (\<Sum> n. inverse(fact(n+2)) * (x ^ (n+2)))" (is "_ = ?a + _")
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1248
    by (rule suminf_split_initial_segment)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1249
  also have "?a = 1 + x"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1250
    by (simp add: numeral_2_eq_2)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1251
  finally show ?thesis .
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1252
qed
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1253
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1254
lemma exp_bound: "0 <= (x::real) ==> x <= 1 ==> exp x <= 1 + x + x^2"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1255
proof -
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1256
  assume a: "0 <= x"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1257
  assume b: "x <= 1"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1258
  { fix n :: nat
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1259
    have "2 * 2 ^ n \<le> fact (n + 2)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1260
      by (induct n, simp, simp)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1261
    hence "real ((2::nat) * 2 ^ n) \<le> real (fact (n + 2))"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1262
      by (simp only: real_of_nat_le_iff)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1263
    hence "2 * 2 ^ n \<le> real (fact (n + 2))"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1264
      by simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1265
    hence "inverse (fact (n + 2)) \<le> inverse (2 * 2 ^ n)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1266
      by (rule le_imp_inverse_le) simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1267
    hence "inverse (fact (n + 2)) \<le> 1/2 * (1/2)^n"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1268
      by (simp add: inverse_mult_distrib power_inverse)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1269
    hence "inverse (fact (n + 2)) * (x^n * x\<twosuperior>) \<le> 1/2 * (1/2)^n * (1 * x\<twosuperior>)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1270
      by (rule mult_mono)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1271
        (rule mult_mono, simp_all add: power_le_one a b mult_nonneg_nonneg)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1272
    hence "inverse (fact (n + 2)) * x ^ (n + 2) \<le> (x\<twosuperior>/2) * ((1/2)^n)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1273
      unfolding power_add by (simp add: mult_ac del: fact_Suc) }
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1274
  note aux1 = this
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1275
  have "(\<lambda>n. x\<twosuperior> / 2 * (1 / 2) ^ n) sums (x\<twosuperior> / 2 * (1 / (1 - 1 / 2)))"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1276
    by (intro sums_mult geometric_sums, simp)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1277
  hence aux2: "(\<lambda>n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1278
    by simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1279
  have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <= x^2"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1280
  proof -
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1281
    have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <=
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1282
        suminf (%n. (x^2/2) * ((1/2)^n))"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1283
      apply (rule summable_le)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1284
      apply (rule allI, rule aux1)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1285
      apply (rule summable_exp [THEN summable_ignore_initial_segment])
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1286
      by (rule sums_summable, rule aux2)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1287
    also have "... = x^2"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1288
      by (rule sums_unique [THEN sym], rule aux2)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1289
    finally show ?thesis .
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1290
  qed
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1291
  thus ?thesis unfolding exp_first_two_terms by auto
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1292
qed
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1293
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1294
lemma ln_one_minus_pos_upper_bound: "0 <= x ==> x < 1 ==> ln (1 - x) <= - x"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1295
proof -
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1296
  assume a: "0 <= (x::real)" and b: "x < 1"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1297
  have "(1 - x) * (1 + x + x^2) = (1 - x^3)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1298
    by (simp add: algebra_simps power2_eq_square power3_eq_cube)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1299
  also have "... <= 1"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1300
    by (auto simp add: a)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1301
  finally have "(1 - x) * (1 + x + x ^ 2) <= 1" .
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1302
  moreover have c: "0 < 1 + x + x\<twosuperior>"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1303
    by (simp add: add_pos_nonneg a)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1304
  ultimately have "1 - x <= 1 / (1 + x + x^2)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1305
    by (elim mult_imp_le_div_pos)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1306
  also have "... <= 1 / exp x"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1307
    apply (rule divide_left_mono)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1308
    apply (rule exp_bound, rule a)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1309
    apply (rule b [THEN less_imp_le])
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1310
    apply simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1311
    apply (rule mult_pos_pos)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1312
    apply (rule c)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1313
    apply simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1314
    done
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1315
  also have "... = exp (-x)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1316
    by (auto simp add: exp_minus divide_inverse)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1317
  finally have "1 - x <= exp (- x)" .
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1318
  also have "1 - x = exp (ln (1 - x))"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1319
  proof -
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1320
    have "0 < 1 - x"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1321
      by (insert b, auto)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1322
    thus ?thesis
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1323
      by (auto simp only: exp_ln_iff [THEN sym])
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1324
  qed
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1325
  finally have "exp (ln (1 - x)) <= exp (- x)" .
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1326
  thus ?thesis by (auto simp only: exp_le_cancel_iff)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1327
qed
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1328
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1329
lemma exp_ge_add_one_self [simp]: "1 + (x::real) <= exp x"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1330
  apply (case_tac "0 <= x")
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1331
  apply (erule exp_ge_add_one_self_aux)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1332
  apply (case_tac "x <= -1")
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1333
  apply (subgoal_tac "1 + x <= 0")
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1334
  apply (erule order_trans)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1335
  apply simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1336
  apply simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1337
  apply (subgoal_tac "1 + x = exp(ln (1 + x))")
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1338
  apply (erule ssubst)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1339
  apply (subst exp_le_cancel_iff)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1340
  apply (subgoal_tac "ln (1 - (- x)) <= - (- x)")
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1341
  apply simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1342
  apply (rule ln_one_minus_pos_upper_bound)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1343
  apply auto
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1344
done
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1345
51527
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1346
lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> x - x^2 <= ln (1 + x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1347
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1348
  assume a: "0 <= x" and b: "x <= 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1349
  have "exp (x - x^2) = exp x / exp (x^2)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1350
    by (rule exp_diff)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1351
  also have "... <= (1 + x + x^2) / exp (x ^2)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1352
    apply (rule divide_right_mono) 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1353
    apply (rule exp_bound)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1354
    apply (rule a, rule b)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1355
    apply simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1356
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1357
  also have "... <= (1 + x + x^2) / (1 + x^2)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1358
    apply (rule divide_left_mono)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1359
    apply (simp add: exp_ge_add_one_self_aux)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1360
    apply (simp add: a)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1361
    apply (simp add: mult_pos_pos add_pos_nonneg)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1362
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1363
  also from a have "... <= 1 + x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1364
    by (simp add: field_simps add_strict_increasing zero_le_mult_iff)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1365
  finally have "exp (x - x^2) <= 1 + x" .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1366
  also have "... = exp (ln (1 + x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1367
  proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1368
    from a have "0 < 1 + x" by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1369
    thus ?thesis
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1370
      by (auto simp only: exp_ln_iff [THEN sym])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1371
  qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1372
  finally have "exp (x - x ^ 2) <= exp (ln (1 + x))" .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1373
  thus ?thesis by (auto simp only: exp_le_cancel_iff)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1374
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1375
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1376
lemma aux5: "x < 1 ==> ln(1 - x) = - ln(1 + x / (1 - x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1377
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1378
  assume a: "x < 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1379
  have "ln(1 - x) = - ln(1 / (1 - x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1380
  proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1381
    have "ln(1 - x) = - (- ln (1 - x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1382
      by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1383
    also have "- ln(1 - x) = ln 1 - ln(1 - x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1384
      by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1385
    also have "... = ln(1 / (1 - x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1386
      apply (rule ln_div [THEN sym])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1387
      by (insert a, auto)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1388
    finally show ?thesis .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1389
  qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1390
  also have " 1 / (1 - x) = 1 + x / (1 - x)" using a by(simp add:field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1391
  finally show ?thesis .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1392
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1393
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1394
lemma ln_one_minus_pos_lower_bound: "0 <= x ==> x <= (1 / 2) ==> 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1395
    - x - 2 * x^2 <= ln (1 - x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1396
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1397
  assume a: "0 <= x" and b: "x <= (1 / 2)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1398
  from b have c: "x < 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1399
    by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1400
  then have "ln (1 - x) = - ln (1 + x / (1 - x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1401
    by (rule aux5)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1402
  also have "- (x / (1 - x)) <= ..."
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1403
  proof - 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1404
    have "ln (1 + x / (1 - x)) <= x / (1 - x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1405
      apply (rule ln_add_one_self_le_self)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1406
      apply (rule divide_nonneg_pos)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1407
      by (insert a c, auto) 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1408
    thus ?thesis
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1409
      by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1410
  qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1411
  also have "- (x / (1 - x)) = -x / (1 - x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1412
    by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1413
  finally have d: "- x / (1 - x) <= ln (1 - x)" .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1414
  have "0 < 1 - x" using a b by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1415
  hence e: "-x - 2 * x^2 <= - x / (1 - x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1416
    using mult_right_le_one_le[of "x*x" "2*x"] a b
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1417
    by (simp add:field_simps power2_eq_square)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1418
  from e d show "- x - 2 * x^2 <= ln (1 - x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1419
    by (rule order_trans)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1420
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1421
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1422
lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1423
  apply (subgoal_tac "ln (1 + x) \<le> ln (exp x)", simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1424
  apply (subst ln_le_cancel_iff)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1425
  apply auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1426
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1427
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1428
lemma abs_ln_one_plus_x_minus_x_bound_nonneg:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1429
    "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1430
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1431
  assume x: "0 <= x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1432
  assume x1: "x <= 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1433
  from x have "ln (1 + x) <= x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1434
    by (rule ln_add_one_self_le_self)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1435
  then have "ln (1 + x) - x <= 0" 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1436
    by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1437
  then have "abs(ln(1 + x) - x) = - (ln(1 + x) - x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1438
    by (rule abs_of_nonpos)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1439
  also have "... = x - ln (1 + x)" 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1440
    by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1441
  also have "... <= x^2"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1442
  proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1443
    from x x1 have "x - x^2 <= ln (1 + x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1444
      by (intro ln_one_plus_pos_lower_bound)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1445
    thus ?thesis
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1446
      by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1447
  qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1448
  finally show ?thesis .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1449
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1450
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1451
lemma abs_ln_one_plus_x_minus_x_bound_nonpos:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1452
    "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1453
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1454
  assume a: "-(1 / 2) <= x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1455
  assume b: "x <= 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1456
  have "abs(ln (1 + x) - x) = x - ln(1 - (-x))" 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1457
    apply (subst abs_of_nonpos)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1458
    apply simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1459
    apply (rule ln_add_one_self_le_self2)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1460
    using a apply auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1461
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1462
  also have "... <= 2 * x^2"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1463
    apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))")
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1464
    apply (simp add: algebra_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1465
    apply (rule ln_one_minus_pos_lower_bound)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1466
    using a b apply auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1467
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1468
  finally show ?thesis .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1469
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1470
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1471
lemma abs_ln_one_plus_x_minus_x_bound:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1472
    "abs x <= 1 / 2 ==> abs(ln (1 + x) - x) <= 2 * x^2"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1473
  apply (case_tac "0 <= x")
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1474
  apply (rule order_trans)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1475
  apply (rule abs_ln_one_plus_x_minus_x_bound_nonneg)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1476
  apply auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1477
  apply (rule abs_ln_one_plus_x_minus_x_bound_nonpos)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1478
  apply auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1479
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1480
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1481
lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)"  
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1482
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1483
  assume x: "exp 1 <= x" "x <= y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1484
  moreover have "0 < exp (1::real)" by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1485
  ultimately have a: "0 < x" and b: "0 < y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1486
    by (fast intro: less_le_trans order_trans)+
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1487
  have "x * ln y - x * ln x = x * (ln y - ln x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1488
    by (simp add: algebra_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1489
  also have "... = x * ln(y / x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1490
    by (simp only: ln_div a b)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1491
  also have "y / x = (x + (y - x)) / x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1492
    by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1493
  also have "... = 1 + (y - x) / x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1494
    using x a by (simp add: field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1495
  also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1496
    apply (rule mult_left_mono)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1497
    apply (rule ln_add_one_self_le_self)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1498
    apply (rule divide_nonneg_pos)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1499
    using x a apply simp_all
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1500
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1501
  also have "... = y - x" using a by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1502
  also have "... = (y - x) * ln (exp 1)" by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1503
  also have "... <= (y - x) * ln x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1504
    apply (rule mult_left_mono)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1505
    apply (subst ln_le_cancel_iff)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1506
    apply fact
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1507
    apply (rule a)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1508
    apply (rule x)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1509
    using x apply simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1510
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1511
  also have "... = y * ln x - x * ln x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1512
    by (rule left_diff_distrib)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1513
  finally have "x * ln y <= y * ln x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1514
    by arith
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1515
  then have "ln y <= (y * ln x) / x" using a by (simp add: field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1516
  also have "... = y * (ln x / x)" by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1517
  finally show ?thesis using b by (simp add: field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1518
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1519
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1520
lemma ln_le_minus_one:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1521
  "0 < x \<Longrightarrow> ln x \<le> x - 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1522
  using exp_ge_add_one_self[of "ln x"] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1523
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1524
lemma ln_eq_minus_one:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1525
  assumes "0 < x" "ln x = x - 1" shows "x = 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1526
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1527
  let "?l y" = "ln y - y + 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1528
  have D: "\<And>x. 0 < x \<Longrightarrow> DERIV ?l x :> (1 / x - 1)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1529
    by (auto intro!: DERIV_intros)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1530
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1531
  show ?thesis
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1532
  proof (cases rule: linorder_cases)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1533
    assume "x < 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1534
    from dense[OF `x < 1`] obtain a where "x < a" "a < 1" by blast
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1535
    from `x < a` have "?l x < ?l a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1536
    proof (rule DERIV_pos_imp_increasing, safe)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1537
      fix y assume "x \<le> y" "y \<le> a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1538
      with `0 < x` `a < 1` have "0 < 1 / y - 1" "0 < y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1539
        by (auto simp: field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1540
      with D show "\<exists>z. DERIV ?l y :> z \<and> 0 < z"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1541
        by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1542
    qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1543
    also have "\<dots> \<le> 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1544
      using ln_le_minus_one `0 < x` `x < a` by (auto simp: field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1545
    finally show "x = 1" using assms by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1546
  next
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1547
    assume "1 < x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1548
    from dense[OF `1 < x`] obtain a where "1 < a" "a < x" by blast
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1549
    from `a < x` have "?l x < ?l a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1550
    proof (rule DERIV_neg_imp_decreasing, safe)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1551
      fix y assume "a \<le> y" "y \<le> x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1552
      with `1 < a` have "1 / y - 1 < 0" "0 < y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1553
        by (auto simp: field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1554
      with D show "\<exists>z. DERIV ?l y :> z \<and> z < 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1555
        by blast
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1556
    qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1557
    also have "\<dots> \<le> 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1558
      using ln_le_minus_one `1 < a` by (auto simp: field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1559
    finally show "x = 1" using assms by auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1560
  qed simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1561
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1562
50326
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1563
lemma exp_at_bot: "(exp ---> (0::real)) at_bot"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1564
  unfolding tendsto_Zfun_iff
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1565
proof (rule ZfunI, simp add: eventually_at_bot_dense)
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1566
  fix r :: real assume "0 < r"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1567
  { fix x assume "x < ln r"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1568
    then have "exp x < exp (ln r)"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1569
      by simp
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1570
    with `0 < r` have "exp x < r"
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1571
      by simp }
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1572
  then show "\<exists>k. \<forall>n<k. exp n < r" by auto
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1573
qed
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1574
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1575
lemma exp_at_top: "LIM x at_top. exp x :: real :> at_top"
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  1576
  by (rule filterlim_at_top_at_top[where Q="\<lambda>x. True" and P="\<lambda>x. 0 < x" and g="ln"])
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  1577
     (auto intro: eventually_gt_at_top)
50326
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1578
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1579
lemma ln_at_0: "LIM x at_right 0. ln x :> at_bot"
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  1580
  by (rule filterlim_at_bot_at_right[where Q="\<lambda>x. 0 < x" and P="\<lambda>x. True" and g="exp"])
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  1581
     (auto simp: eventually_within)
50326
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1582
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1583
lemma ln_at_top: "LIM x at_top. ln x :> at_top"
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  1584
  by (rule filterlim_at_top_at_top[where Q="\<lambda>x. 0 < x" and P="\<lambda>x. True" and g="exp"])
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  1585
     (auto intro: eventually_gt_at_top)
50326
b5afeccab2db add filterlim rules for exp and ln to infinity
hoelzl
parents: 49962
diff changeset
  1586
50347
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1587
lemma tendsto_power_div_exp_0: "((\<lambda>x. x ^ k / exp x) ---> (0::real)) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1588
proof (induct k)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1589
  show "((\<lambda>x. x ^ 0 / exp x) ---> (0::real)) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1590
    by (simp add: inverse_eq_divide[symmetric])
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1591
       (metis filterlim_compose[OF tendsto_inverse_0] exp_at_top filterlim_mono
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1592
              at_top_le_at_infinity order_refl)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1593
next
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1594
  case (Suc k)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1595
  show ?case
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1596
  proof (rule lhospital_at_top_at_top)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1597
    show "eventually (\<lambda>x. DERIV (\<lambda>x. x ^ Suc k) x :> (real (Suc k) * x^k)) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1598
      by eventually_elim (intro DERIV_intros, simp, simp)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1599
    show "eventually (\<lambda>x. DERIV exp x :> exp x) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1600
      by eventually_elim (auto intro!: DERIV_intros)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1601
    show "eventually (\<lambda>x. exp x \<noteq> 0) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1602
      by auto
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1603
    from tendsto_mult[OF tendsto_const Suc, of "real (Suc k)"]
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1604
    show "((\<lambda>x. real (Suc k) * x ^ k / exp x) ---> 0) at_top"
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1605
      by simp
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1606
  qed (rule exp_at_top)
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1607
qed
77e3effa50b6 prove tendsto_power_div_exp_0
hoelzl
parents: 50346
diff changeset
  1608
51527
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1609
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1610
definition
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1611
  powr  :: "[real,real] => real"     (infixr "powr" 80) where
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1612
    --{*exponentation with real exponent*}
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1613
  "x powr a = exp(a * ln x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1614
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1615
definition
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1616
  log :: "[real,real] => real" where
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1617
    --{*logarithm of @{term x} to base @{term a}*}
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1618
  "log a x = ln x / ln a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1619
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1620
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1621
lemma tendsto_log [tendsto_intros]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1622
  "\<lbrakk>(f ---> a) F; (g ---> b) F; 0 < a; a \<noteq> 1; 0 < b\<rbrakk> \<Longrightarrow> ((\<lambda>x. log (f x) (g x)) ---> log a b) F"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1623
  unfolding log_def by (intro tendsto_intros) auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1624
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1625
lemma continuous_log:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1626
  assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\<lambda>x. x))" and "f (Lim F (\<lambda>x. x)) \<noteq> 1" and "0 < g (Lim F (\<lambda>x. x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1627
  shows "continuous F (\<lambda>x. log (f x) (g x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1628
  using assms unfolding continuous_def by (rule tendsto_log)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1629
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1630
lemma continuous_at_within_log[continuous_intros]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1631
  assumes "continuous (at a within s) f" "continuous (at a within s) g" and "0 < f a" and "f a \<noteq> 1" and "0 < g a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1632
  shows "continuous (at a within s) (\<lambda>x. log (f x) (g x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1633
  using assms unfolding continuous_within by (rule tendsto_log)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1634
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1635
lemma isCont_log[continuous_intros, simp]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1636
  assumes "isCont f a" "isCont g a" "0 < f a" "f a \<noteq> 1" "0 < g a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1637
  shows "isCont (\<lambda>x. log (f x) (g x)) a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1638
  using assms unfolding continuous_at by (rule tendsto_log)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1639
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1640
lemma continuous_on_log[continuous_on_intros]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1641
  assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. 0 < f x" "\<forall>x\<in>s. f x \<noteq> 1" "\<forall>x\<in>s. 0 < g x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1642
  shows "continuous_on s (\<lambda>x. log (f x) (g x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1643
  using assms unfolding continuous_on_def by (fast intro: tendsto_log)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1644
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1645
lemma powr_one_eq_one [simp]: "1 powr a = 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1646
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1647
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1648
lemma powr_zero_eq_one [simp]: "x powr 0 = 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1649
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1650
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1651
lemma powr_one_gt_zero_iff [simp]: "(x powr 1 = x) = (0 < x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1652
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1653
declare powr_one_gt_zero_iff [THEN iffD2, simp]
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1654
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1655
lemma powr_mult: 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1656
      "[| 0 < x; 0 < y |] ==> (x * y) powr a = (x powr a) * (y powr a)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1657
by (simp add: powr_def exp_add [symmetric] ln_mult distrib_left)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1658
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1659
lemma powr_gt_zero [simp]: "0 < x powr a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1660
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1661
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1662
lemma powr_ge_pzero [simp]: "0 <= x powr y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1663
by (rule order_less_imp_le, rule powr_gt_zero)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1664
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1665
lemma powr_not_zero [simp]: "x powr a \<noteq> 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1666
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1667
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1668
lemma powr_divide:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1669
     "[| 0 < x; 0 < y |] ==> (x / y) powr a = (x powr a)/(y powr a)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1670
apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1671
apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1672
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1673
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1674
lemma powr_divide2: "x powr a / x powr b = x powr (a - b)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1675
  apply (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1676
  apply (subst exp_diff [THEN sym])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1677
  apply (simp add: left_diff_distrib)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1678
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1679
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1680
lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1681
by (simp add: powr_def exp_add [symmetric] distrib_right)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1682
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1683
lemma powr_mult_base:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1684
  "0 < x \<Longrightarrow>x * x powr y = x powr (1 + y)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1685
using assms by (auto simp: powr_add)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1686
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1687
lemma powr_powr: "(x powr a) powr b = x powr (a * b)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1688
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1689
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1690
lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1691
by (simp add: powr_powr mult_commute)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1692
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1693
lemma powr_minus: "x powr (-a) = inverse (x powr a)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1694
by (simp add: powr_def exp_minus [symmetric])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1695
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1696
lemma powr_minus_divide: "x powr (-a) = 1/(x powr a)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1697
by (simp add: divide_inverse powr_minus)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1698
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1699
lemma powr_less_mono: "[| a < b; 1 < x |] ==> x powr a < x powr b"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1700
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1701
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1702
lemma powr_less_cancel: "[| x powr a < x powr b; 1 < x |] ==> a < b"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1703
by (simp add: powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1704
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1705
lemma powr_less_cancel_iff [simp]: "1 < x ==> (x powr a < x powr b) = (a < b)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1706
by (blast intro: powr_less_cancel powr_less_mono)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1707
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1708
lemma powr_le_cancel_iff [simp]: "1 < x ==> (x powr a \<le> x powr b) = (a \<le> b)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1709
by (simp add: linorder_not_less [symmetric])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1710
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1711
lemma log_ln: "ln x = log (exp(1)) x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1712
by (simp add: log_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1713
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1714
lemma DERIV_log: assumes "x > 0" shows "DERIV (\<lambda>y. log b y) x :> 1 / (ln b * x)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1715
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1716
  def lb \<equiv> "1 / ln b"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1717
  moreover have "DERIV (\<lambda>y. lb * ln y) x :> lb / x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1718
    using `x > 0` by (auto intro!: DERIV_intros)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1719
  ultimately show ?thesis
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1720
    by (simp add: log_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1721
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1722
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1723
lemmas DERIV_log[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1724
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1725
lemma powr_log_cancel [simp]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1726
     "[| 0 < a; a \<noteq> 1; 0 < x |] ==> a powr (log a x) = x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1727
by (simp add: powr_def log_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1728
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1729
lemma log_powr_cancel [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a (a powr y) = y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1730
by (simp add: log_def powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1731
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1732
lemma log_mult: 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1733
     "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y |]  
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1734
      ==> log a (x * y) = log a x + log a y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1735
by (simp add: log_def ln_mult divide_inverse distrib_right)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1736
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1737
lemma log_eq_div_ln_mult_log: 
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1738
     "[| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]  
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1739
      ==> log a x = (ln b/ln a) * log b x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1740
by (simp add: log_def divide_inverse)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1741
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1742
text{*Base 10 logarithms*}
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1743
lemma log_base_10_eq1: "0 < x ==> log 10 x = (ln (exp 1) / ln 10) * ln x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1744
by (simp add: log_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1745
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1746
lemma log_base_10_eq2: "0 < x ==> log 10 x = (log 10 (exp 1)) * ln x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1747
by (simp add: log_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1748
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1749
lemma log_one [simp]: "log a 1 = 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1750
by (simp add: log_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1751
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1752
lemma log_eq_one [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a a = 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1753
by (simp add: log_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1754
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1755
lemma log_inverse:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1756
     "[| 0 < a; a \<noteq> 1; 0 < x |] ==> log a (inverse x) = - log a x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1757
apply (rule_tac a1 = "log a x" in add_left_cancel [THEN iffD1])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1758
apply (simp add: log_mult [symmetric])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1759
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1760
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1761
lemma log_divide:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1762
     "[|0 < a; a \<noteq> 1; 0 < x; 0 < y|] ==> log a (x/y) = log a x - log a y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1763
by (simp add: log_mult divide_inverse log_inverse)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1764
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1765
lemma log_less_cancel_iff [simp]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1766
     "[| 1 < a; 0 < x; 0 < y |] ==> (log a x < log a y) = (x < y)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1767
apply safe
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1768
apply (rule_tac [2] powr_less_cancel)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1769
apply (drule_tac a = "log a x" in powr_less_mono, auto)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1770
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1771
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1772
lemma log_inj: assumes "1 < b" shows "inj_on (log b) {0 <..}"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1773
proof (rule inj_onI, simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1774
  fix x y assume pos: "0 < x" "0 < y" and *: "log b x = log b y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1775
  show "x = y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1776
  proof (cases rule: linorder_cases)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1777
    assume "x < y" hence "log b x < log b y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1778
      using log_less_cancel_iff[OF `1 < b`] pos by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1779
    thus ?thesis using * by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1780
  next
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1781
    assume "y < x" hence "log b y < log b x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1782
      using log_less_cancel_iff[OF `1 < b`] pos by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1783
    thus ?thesis using * by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1784
  qed simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1785
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1786
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1787
lemma log_le_cancel_iff [simp]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1788
     "[| 1 < a; 0 < x; 0 < y |] ==> (log a x \<le> log a y) = (x \<le> y)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1789
by (simp add: linorder_not_less [symmetric])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1790
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1791
lemma zero_less_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < log a x \<longleftrightarrow> 1 < x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1792
  using log_less_cancel_iff[of a 1 x] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1793
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1794
lemma zero_le_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 \<le> log a x \<longleftrightarrow> 1 \<le> x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1795
  using log_le_cancel_iff[of a 1 x] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1796
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1797
lemma log_less_zero_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x < 0 \<longleftrightarrow> x < 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1798
  using log_less_cancel_iff[of a x 1] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1799
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1800
lemma log_le_zero_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x \<le> 0 \<longleftrightarrow> x \<le> 1"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1801
  using log_le_cancel_iff[of a x 1] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1802
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1803
lemma one_less_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 1 < log a x \<longleftrightarrow> a < x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1804
  using log_less_cancel_iff[of a a x] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1805
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1806
lemma one_le_log_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 1 \<le> log a x \<longleftrightarrow> a \<le> x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1807
  using log_le_cancel_iff[of a a x] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1808
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1809
lemma log_less_one_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x < 1 \<longleftrightarrow> x < a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1810
  using log_less_cancel_iff[of a x a] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1811
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1812
lemma log_le_one_cancel_iff[simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> log a x \<le> 1 \<longleftrightarrow> x \<le> a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1813
  using log_le_cancel_iff[of a x a] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1814
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1815
lemma powr_realpow: "0 < x ==> x powr (real n) = x^n"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1816
  apply (induct n, simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1817
  apply (subgoal_tac "real(Suc n) = real n + 1")
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1818
  apply (erule ssubst)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1819
  apply (subst powr_add, simp, simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1820
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1821
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1822
lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 else x powr (real n))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1823
  apply (case_tac "x = 0", simp, simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1824
  apply (rule powr_realpow [THEN sym], simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1825
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1826
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1827
lemma powr_int:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1828
  assumes "x > 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1829
  shows "x powr i = (if i \<ge> 0 then x ^ nat i else 1 / x ^ nat (-i))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1830
proof cases
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1831
  assume "i < 0"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1832
  have r: "x powr i = 1 / x powr (-i)" by (simp add: powr_minus field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1833
  show ?thesis using `i < 0` `x > 0` by (simp add: r field_simps powr_realpow[symmetric])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1834
qed (simp add: assms powr_realpow[symmetric])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1835
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1836
lemma powr_numeral: "0 < x \<Longrightarrow> x powr numeral n = x^numeral n"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1837
  using powr_realpow[of x "numeral n"] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1838
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1839
lemma powr_neg_numeral: "0 < x \<Longrightarrow> x powr neg_numeral n = 1 / x^numeral n"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1840
  using powr_int[of x "neg_numeral n"] by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1841
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1842
lemma root_powr_inverse:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1843
  "0 < n \<Longrightarrow> 0 < x \<Longrightarrow> root n x = x powr (1/n)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1844
  by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1845
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1846
lemma ln_powr: "0 < x ==> 0 < y ==> ln(x powr y) = y * ln x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1847
by (unfold powr_def, simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1848
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1849
lemma log_powr: "0 < x ==> 0 \<le> y ==> log b (x powr y) = y * log b x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1850
  apply (case_tac "y = 0")
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1851
  apply force
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1852
  apply (auto simp add: log_def ln_powr field_simps)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1853
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1854
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1855
lemma log_nat_power: "0 < x ==> log b (x^n) = real n * log b x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1856
  apply (subst powr_realpow [symmetric])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1857
  apply (auto simp add: log_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1858
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1859
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1860
lemma ln_bound: "1 <= x ==> ln x <= x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1861
  apply (subgoal_tac "ln(1 + (x - 1)) <= x - 1")
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1862
  apply simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1863
  apply (rule ln_add_one_self_le_self, simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1864
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1865
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1866
lemma powr_mono: "a <= b ==> 1 <= x ==> x powr a <= x powr b"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1867
  apply (case_tac "x = 1", simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1868
  apply (case_tac "a = b", simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1869
  apply (rule order_less_imp_le)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1870
  apply (rule powr_less_mono, auto)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1871
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1872
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1873
lemma ge_one_powr_ge_zero: "1 <= x ==> 0 <= a ==> 1 <= x powr a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1874
  apply (subst powr_zero_eq_one [THEN sym])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1875
  apply (rule powr_mono, assumption+)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1876
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1877
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1878
lemma powr_less_mono2: "0 < a ==> 0 < x ==> x < y ==> x powr a <
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1879
    y powr a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1880
  apply (unfold powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1881
  apply (rule exp_less_mono)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1882
  apply (rule mult_strict_left_mono)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1883
  apply (subst ln_less_cancel_iff, assumption)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1884
  apply (rule order_less_trans)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1885
  prefer 2
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1886
  apply assumption+
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1887
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1888
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1889
lemma powr_less_mono2_neg: "a < 0 ==> 0 < x ==> x < y ==> y powr a <
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1890
    x powr a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1891
  apply (unfold powr_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1892
  apply (rule exp_less_mono)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1893
  apply (rule mult_strict_left_mono_neg)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1894
  apply (subst ln_less_cancel_iff)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1895
  apply assumption
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1896
  apply (rule order_less_trans)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1897
  prefer 2
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1898
  apply assumption+
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1899
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1900
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1901
lemma powr_mono2: "0 <= a ==> 0 < x ==> x <= y ==> x powr a <= y powr a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1902
  apply (case_tac "a = 0", simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1903
  apply (case_tac "x = y", simp)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1904
  apply (rule order_less_imp_le)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1905
  apply (rule powr_less_mono2, auto)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1906
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1907
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1908
lemma powr_inj:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1909
  "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> a powr x = a powr y \<longleftrightarrow> x = y"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1910
  unfolding powr_def exp_inj_iff by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1911
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1912
lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1913
  apply (rule mult_imp_le_div_pos)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1914
  apply (assumption)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1915
  apply (subst mult_commute)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1916
  apply (subst ln_powr [THEN sym])
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1917
  apply auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1918
  apply (rule ln_bound)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1919
  apply (erule ge_one_powr_ge_zero)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1920
  apply (erule order_less_imp_le)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1921
done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1922
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1923
lemma ln_powr_bound2:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1924
  assumes "1 < x" and "0 < a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1925
  shows "(ln x) powr a <= (a powr a) * x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1926
proof -
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1927
  from assms have "ln x <= (x powr (1 / a)) / (1 / a)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1928
    apply (intro ln_powr_bound)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1929
    apply (erule order_less_imp_le)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1930
    apply (rule divide_pos_pos)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1931
    apply simp_all
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1932
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1933
  also have "... = a * (x powr (1 / a))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1934
    by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1935
  finally have "(ln x) powr a <= (a * (x powr (1 / a))) powr a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1936
    apply (intro powr_mono2)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1937
    apply (rule order_less_imp_le, rule assms)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1938
    apply (rule ln_gt_zero)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1939
    apply (rule assms)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1940
    apply assumption
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1941
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1942
  also have "... = (a powr a) * ((x powr (1 / a)) powr a)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1943
    apply (rule powr_mult)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1944
    apply (rule assms)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1945
    apply (rule powr_gt_zero)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1946
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1947
  also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1948
    by (rule powr_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1949
  also have "... = x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1950
    apply simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1951
    apply (subgoal_tac "a ~= 0")
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1952
    using assms apply auto
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1953
    done
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1954
  finally show ?thesis .
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1955
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1956
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1957
lemma tendsto_powr [tendsto_intros]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1958
  "\<lbrakk>(f ---> a) F; (g ---> b) F; 0 < a\<rbrakk> \<Longrightarrow> ((\<lambda>x. f x powr g x) ---> a powr b) F"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1959
  unfolding powr_def by (intro tendsto_intros)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1960
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1961
lemma continuous_powr:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1962
  assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\<lambda>x. x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1963
  shows "continuous F (\<lambda>x. (f x) powr (g x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1964
  using assms unfolding continuous_def by (rule tendsto_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1965
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1966
lemma continuous_at_within_powr[continuous_intros]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1967
  assumes "continuous (at a within s) f" "continuous (at a within s) g" and "0 < f a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1968
  shows "continuous (at a within s) (\<lambda>x. (f x) powr (g x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1969
  using assms unfolding continuous_within by (rule tendsto_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1970
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1971
lemma isCont_powr[continuous_intros, simp]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1972
  assumes "isCont f a" "isCont g a" "0 < f a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1973
  shows "isCont (\<lambda>x. (f x) powr g x) a"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1974
  using assms unfolding continuous_at by (rule tendsto_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1975
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1976
lemma continuous_on_powr[continuous_on_intros]:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1977
  assumes "continuous_on s f" "continuous_on s g" and "\<forall>x\<in>s. 0 < f x"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1978
  shows "continuous_on s (\<lambda>x. (f x) powr (g x))"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1979
  using assms unfolding continuous_on_def by (fast intro: tendsto_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1980
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1981
(* FIXME: generalize by replacing d by with g x and g ---> d? *)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1982
lemma tendsto_zero_powrI:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1983
  assumes "eventually (\<lambda>x. 0 < f x ) F" and "(f ---> 0) F"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1984
  assumes "0 < d"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1985
  shows "((\<lambda>x. f x powr d) ---> 0) F"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1986
proof (rule tendstoI)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1987
  fix e :: real assume "0 < e"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1988
  def Z \<equiv> "e powr (1 / d)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1989
  with `0 < e` have "0 < Z" by simp
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1990
  with assms have "eventually (\<lambda>x. 0 < f x \<and> dist (f x) 0 < Z) F"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1991
    by (intro eventually_conj tendstoD)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1992
  moreover
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1993
  from assms have "\<And>x. 0 < x \<and> dist x 0 < Z \<Longrightarrow> x powr d < Z powr d"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1994
    by (intro powr_less_mono2) (auto simp: dist_real_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1995
  with assms `0 < e` have "\<And>x. 0 < x \<and> dist x 0 < Z \<Longrightarrow> dist (x powr d) 0 < e"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1996
    unfolding dist_real_def Z_def by (auto simp: powr_powr)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1997
  ultimately
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1998
  show "eventually (\<lambda>x. dist (f x powr d) 0 < e) F" by (rule eventually_elim1)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  1999
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2000
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2001
lemma tendsto_neg_powr:
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2002
  assumes "s < 0" and "LIM x F. f x :> at_top"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2003
  shows "((\<lambda>x. f x powr s) ---> 0) F"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2004
proof (rule tendstoI)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2005
  fix e :: real assume "0 < e"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2006
  def Z \<equiv> "e powr (1 / s)"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2007
  from assms have "eventually (\<lambda>x. Z < f x) F"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2008
    by (simp add: filterlim_at_top_dense)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2009
  moreover
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2010
  from assms have "\<And>x. Z < x \<Longrightarrow> x powr s < Z powr s"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2011
    by (auto simp: Z_def intro!: powr_less_mono2_neg)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2012
  with assms `0 < e` have "\<And>x. Z < x \<Longrightarrow> dist (x powr s) 0 < e"
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2013
    by (simp add: powr_powr Z_def dist_real_def)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2014
  ultimately
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2015
  show "eventually (\<lambda>x. dist (f x powr s) 0 < e) F" by (rule eventually_elim1)
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2016
qed
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2017
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2018
subsection {* Sine and Cosine *}
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2019
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2020
definition sin_coeff :: "nat \<Rightarrow> real" where
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2021
  "sin_coeff = (\<lambda>n. if even n then 0 else -1 ^ ((n - Suc 0) div 2) / real (fact n))"
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2022
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2023
definition cos_coeff :: "nat \<Rightarrow> real" where
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2024
  "cos_coeff = (\<lambda>n. if even n then (-1 ^ (n div 2)) / real (fact n) else 0)"
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2025
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2026
definition sin :: "real \<Rightarrow> real" where
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2027
  "sin = (\<lambda>x. \<Sum>n. sin_coeff n * x ^ n)"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2028
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2029
definition cos :: "real \<Rightarrow> real" where
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2030
  "cos = (\<lambda>x. \<Sum>n. cos_coeff n * x ^ n)"
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2031
44319
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2032
lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0"
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2033
  unfolding sin_coeff_def by simp
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2034
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2035
lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1"
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2036
  unfolding cos_coeff_def by simp
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2037
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2038
lemma sin_coeff_Suc: "sin_coeff (Suc n) = cos_coeff n / real (Suc n)"
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2039
  unfolding cos_coeff_def sin_coeff_def
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2040
  by (simp del: mult_Suc)
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2041
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2042
lemma cos_coeff_Suc: "cos_coeff (Suc n) = - sin_coeff n / real (Suc n)"
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2043
  unfolding cos_coeff_def sin_coeff_def
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2044
  by (simp del: mult_Suc, auto simp add: odd_Suc_mult_two_ex)
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2045
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2046
lemma summable_sin: "summable (\<lambda>n. sin_coeff n * x ^ n)"
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2047
unfolding sin_coeff_def
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2048
apply (rule summable_comparison_test [OF _ summable_exp [where x="\<bar>x\<bar>"]])
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2049
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
  2050
done
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2051
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2052
lemma summable_cos: "summable (\<lambda>n. cos_coeff n * x ^ n)"
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2053
unfolding cos_coeff_def
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2054
apply (rule summable_comparison_test [OF _ summable_exp [where x="\<bar>x\<bar>"]])
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2055
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
  2056
done
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2057
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2058
lemma sin_converges: "(\<lambda>n. sin_coeff n * x ^ n) sums sin(x)"
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2059
unfolding sin_def by (rule summable_sin [THEN summable_sums])
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2060
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2061
lemma cos_converges: "(\<lambda>n. cos_coeff n * x ^ n) sums cos(x)"
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2062
unfolding cos_def by (rule summable_cos [THEN summable_sums])
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2063
44319
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2064
lemma diffs_sin_coeff: "diffs sin_coeff = cos_coeff"
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2065
  by (simp add: diffs_def sin_coeff_Suc real_of_nat_def del: of_nat_Suc)
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2066
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2067
lemma diffs_cos_coeff: "diffs cos_coeff = (\<lambda>n. - sin_coeff n)"
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2068
  by (simp add: diffs_def cos_coeff_Suc real_of_nat_def del: of_nat_Suc)
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2069
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2070
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
  2071
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2072
lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)"
44319
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2073
  unfolding sin_def cos_def
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2074
  apply (rule DERIV_cong, rule termdiffs [where K="1 + \<bar>x\<bar>"])
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2075
  apply (simp_all add: diffs_sin_coeff diffs_cos_coeff
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2076
    summable_minus summable_sin summable_cos)
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2077
  done
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2078
51527
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2079
declare DERIV_sin[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2080
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2081
lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)"
44319
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2082
  unfolding cos_def sin_def
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2083
  apply (rule DERIV_cong, rule termdiffs [where K="1 + \<bar>x\<bar>"])
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2084
  apply (simp_all add: diffs_sin_coeff diffs_cos_coeff diffs_minus
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2085
    summable_minus summable_sin summable_cos suminf_minus)
806e0390de53 move sin_coeff and cos_coeff lemmas to Transcendental.thy; simplify some proofs
huffman
parents: 44318
diff changeset
  2086
  done
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2087
51527
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2088
declare DERIV_cos[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
bd62e7ff103b move Ln.thy and Log.thy to Transcendental.thy
hoelzl
parents: 51482
diff changeset
  2089
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2090
lemma isCont_sin: "isCont sin x"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2091
  by (rule DERIV_sin [THEN DERIV_isCont])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2092
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2093
lemma isCont_cos: "isCont cos x"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2094
  by (rule DERIV_cos [THEN DERIV_isCont])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2095
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2096
lemma isCont_sin' [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. sin (f x)) a"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2097
  by (rule isCont_o2 [OF _ isCont_sin])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2098
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2099
lemma isCont_cos' [simp]: "isCont f a \<Longrightarrow> isCont (\<lambda>x. cos (f x)) a"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2100
  by (rule isCont_o2 [OF _ isCont_cos])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2101
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2102
lemma tendsto_sin [tendsto_intros]:
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2103
  "(f ---> a) F \<Longrightarrow> ((\<lambda>x. sin (f x)) ---> sin a) F"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2104
  by (rule isCont_tendsto_compose [OF isCont_sin])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2105
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2106
lemma tendsto_cos [tendsto_intros]:
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2107
  "(f ---> a) F \<Longrightarrow> ((\<lambda>x. cos (f x)) ---> cos a) F"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2108
  by (rule isCont_tendsto_compose [OF isCont_cos])
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2109
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2110
lemma continuous_sin [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2111
  "continuous F f \<Longrightarrow> continuous F (\<lambda>x. sin (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2112
  unfolding continuous_def by (rule tendsto_sin)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2113
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2114
lemma continuous_on_sin [continuous_on_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2115
  "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. sin (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2116
  unfolding continuous_on_def by (auto intro: tendsto_sin)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2117
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2118
lemma continuous_cos [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2119
  "continuous F f \<Longrightarrow> continuous F (\<lambda>x. cos (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2120
  unfolding continuous_def by (rule tendsto_cos)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2121
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2122
lemma continuous_on_cos [continuous_on_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2123
  "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. cos (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2124
  unfolding continuous_on_def by (auto intro: tendsto_cos)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2125
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2126
subsection {* Properties of Sine and Cosine *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2127
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2128
lemma sin_zero [simp]: "sin 0 = 0"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2129
  unfolding sin_def sin_coeff_def by (simp add: powser_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2130
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2131
lemma cos_zero [simp]: "cos 0 = 1"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2132
  unfolding cos_def cos_coeff_def by (simp add: powser_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2133
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2134
lemma sin_cos_squared_add [simp]: "(sin x)\<twosuperior> + (cos x)\<twosuperior> = 1"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2135
proof -
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2136
  have "\<forall>x. DERIV (\<lambda>x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :> 0"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2137
    by (auto intro!: DERIV_intros)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2138
  hence "(sin x)\<twosuperior> + (cos x)\<twosuperior> = (sin 0)\<twosuperior> + (cos 0)\<twosuperior>"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2139
    by (rule DERIV_isconst_all)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2140
  thus "(sin x)\<twosuperior> + (cos x)\<twosuperior> = 1" by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2141
qed
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2142
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2143
lemma sin_cos_squared_add2 [simp]: "(cos x)\<twosuperior> + (sin x)\<twosuperior> = 1"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2144
  by (subst add_commute, rule sin_cos_squared_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2145
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2146
lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2147
  using sin_cos_squared_add2 [unfolded power2_eq_square] .
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2148
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2149
lemma sin_squared_eq: "(sin x)\<twosuperior> = 1 - (cos x)\<twosuperior>"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2150
  unfolding eq_diff_eq by (rule sin_cos_squared_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2151
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2152
lemma cos_squared_eq: "(cos x)\<twosuperior> = 1 - (sin x)\<twosuperior>"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2153
  unfolding eq_diff_eq by (rule sin_cos_squared_add2)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2154
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
  2155
lemma abs_sin_le_one [simp]: "\<bar>sin x\<bar> \<le> 1"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2156
  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
  2157
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2158
lemma sin_ge_minus_one [simp]: "-1 \<le> sin x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2159
  using abs_sin_le_one [of x] unfolding abs_le_iff by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2160
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2161
lemma sin_le_one [simp]: "sin x \<le> 1"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2162
  using abs_sin_le_one [of x] unfolding abs_le_iff by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2163
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
  2164
lemma abs_cos_le_one [simp]: "\<bar>cos x\<bar> \<le> 1"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2165
  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
  2166
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2167
lemma cos_ge_minus_one [simp]: "-1 \<le> cos x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2168
  using abs_cos_le_one [of x] unfolding abs_le_iff by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2169
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2170
lemma cos_le_one [simp]: "cos x \<le> 1"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2171
  using abs_cos_le_one [of x] unfolding abs_le_iff by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2172
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2173
lemma DERIV_fun_pow: "DERIV g x :> m ==>
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2174
      DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2175
  by (auto intro!: DERIV_intros)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2176
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2177
lemma DERIV_fun_exp:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2178
     "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2179
  by (auto intro!: DERIV_intros)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2180
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2181
lemma DERIV_fun_sin:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2182
     "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2183
  by (auto intro!: DERIV_intros)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2184
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2185
lemma DERIV_fun_cos:
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2186
     "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2187
  by (auto intro!: DERIV_intros)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2188
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2189
lemma sin_cos_add_lemma:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2190
     "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 +
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2191
      (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2192
  (is "?f x = 0")
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2193
proof -
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2194
  have "\<forall>x. DERIV (\<lambda>x. ?f x) x :> 0"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2195
    by (auto intro!: DERIV_intros simp add: algebra_simps)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2196
  hence "?f x = ?f 0"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2197
    by (rule DERIV_isconst_all)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2198
  thus ?thesis by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2199
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2200
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2201
lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2202
  using sin_cos_add_lemma unfolding realpow_two_sum_zero_iff by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2203
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2204
lemma cos_add: "cos (x + y) = cos x * cos y - sin x * sin y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2205
  using sin_cos_add_lemma unfolding realpow_two_sum_zero_iff by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2206
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2207
lemma sin_cos_minus_lemma:
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2208
  "(sin(-x) + sin(x))\<twosuperior> + (cos(-x) - cos(x))\<twosuperior> = 0" (is "?f x = 0")
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2209
proof -
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2210
  have "\<forall>x. DERIV (\<lambda>x. ?f x) x :> 0"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2211
    by (auto intro!: DERIV_intros simp add: algebra_simps)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2212
  hence "?f x = ?f 0"
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2213
    by (rule DERIV_isconst_all)
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2214
  thus ?thesis by simp
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2215
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2216
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2217
lemma sin_minus [simp]: "sin (-x) = -sin(x)"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2218
  using sin_cos_minus_lemma [where x=x] by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2219
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2220
lemma cos_minus [simp]: "cos (-x) = cos(x)"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2221
  using sin_cos_minus_lemma [where x=x] by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2222
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2223
lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2224
  by (simp add: diff_minus sin_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2225
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2226
lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2227
  by (simp add: sin_diff mult_commute)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2228
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2229
lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2230
  by (simp add: diff_minus cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2231
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2232
lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x"
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2233
  by (simp add: cos_diff mult_commute)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2234
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2235
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
  2236
  using sin_add [where x=x and y=x] by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2237
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2238
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
  2239
  using cos_add [where x=x and y=x]
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  2240
  by (simp add: power2_eq_square)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2241
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2242
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2243
subsection {* The Constant Pi *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2244
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  2245
definition pi :: "real" where
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2246
  "pi = 2 * (THE x. 0 \<le> (x::real) & x \<le> 2 & cos x = 0)"
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2247
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2248
text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"};
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2249
   hence define pi.*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2250
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2251
lemma sin_paired:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2252
     "(%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
  2253
      sums  sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2254
proof -
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2255
  have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2. sin_coeff k * x ^ k) sums sin x"
44727
d45acd50a894 modify lemma sums_group, and shorten proofs that use it
huffman
parents: 44726
diff changeset
  2256
    by (rule sin_converges [THEN sums_group], simp)
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2257
  thus ?thesis unfolding One_nat_def sin_coeff_def by (simp add: mult_ac)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2258
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2259
44728
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2260
lemma sin_gt_zero:
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2261
  assumes "0 < x" and "x < 2" shows "0 < sin x"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2262
proof -
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2263
  let ?f = "\<lambda>n. \<Sum>k = n*2..<n*2+2. -1 ^ k / real (fact (2*k+1)) * x^(2*k+1)"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2264
  have pos: "\<forall>n. 0 < ?f n"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2265
  proof
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2266
    fix n :: nat
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2267
    let ?k2 = "real (Suc (Suc (4 * n)))"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2268
    let ?k3 = "real (Suc (Suc (Suc (4 * n))))"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2269
    have "x * x < ?k2 * ?k3"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2270
      using assms by (intro mult_strict_mono', simp_all)
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2271
    hence "x * x * x * x ^ (n * 4) < ?k2 * ?k3 * x * x ^ (n * 4)"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2272
      by (intro mult_strict_right_mono zero_less_power `0 < x`)
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2273
    thus "0 < ?f n"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2274
      by (simp del: mult_Suc,
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2275
        simp add: less_divide_eq mult_pos_pos field_simps del: mult_Suc)
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2276
  qed
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2277
  have sums: "?f sums sin x"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2278
    by (rule sin_paired [THEN sums_group], simp)
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2279
  show "0 < sin x"
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2280
    unfolding sums_unique [OF sums]
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2281
    using sums_summable [OF sums] pos
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2282
    by (rule suminf_gt_zero)
86f43cca4886 convert lemma sin_gt_zero to Isar style;
huffman
parents: 44727
diff changeset
  2283
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2284
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2285
lemma cos_double_less_one: "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2286
apply (cut_tac x = x in sin_gt_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2287
apply (auto simp add: cos_squared_eq cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2288
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2289
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2290
lemma cos_paired:
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  2291
     "(%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
  2292
proof -
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2293
  have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2. cos_coeff k * x ^ k) sums cos x"
44727
d45acd50a894 modify lemma sums_group, and shorten proofs that use it
huffman
parents: 44726
diff changeset
  2294
    by (rule cos_converges [THEN sums_group], simp)
31271
0237e5e40b71 add constants sin_coeff, cos_coeff
huffman
parents: 31148
diff changeset
  2295
  thus ?thesis unfolding cos_coeff_def by (simp add: mult_ac)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2296
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2297
36824
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2298
lemma real_mult_inverse_cancel:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2299
     "[|(0::real) < x; 0 < x1; x1 * y < x * u |]
36824
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2300
      ==> inverse x * y < inverse x1 * u"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2301
apply (rule_tac c=x in mult_less_imp_less_left)
36824
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2302
apply (auto simp add: mult_assoc [symmetric])
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2303
apply (simp (no_asm) add: mult_ac)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2304
apply (rule_tac c=x1 in mult_less_imp_less_right)
36824
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2305
apply (auto simp add: mult_ac)
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2306
done
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2307
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2308
lemma real_mult_inverse_cancel2:
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2309
     "[|(0::real) < x;0 < x1; x1 * y < x * u |] ==> y * inverse x < u * inverse x1"
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2310
apply (auto dest: real_mult_inverse_cancel simp add: mult_ac)
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2311
done
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2312
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2313
lemma realpow_num_eq_if:
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2314
  fixes m :: "'a::power"
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2315
  shows "m ^ n = (if n=0 then 1 else m * m ^ (n - 1))"
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2316
by (cases n, auto)
2e9a866141b8 move some theorems from RealPow.thy to Transcendental.thy
huffman
parents: 36777
diff changeset
  2317
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2318
lemma cos_two_less_zero [simp]: "cos (2) < 0"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2319
apply (cut_tac x = 2 in cos_paired)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2320
apply (drule sums_minus)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2321
apply (rule neg_less_iff_less [THEN iffD1])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2322
apply (frule sums_unique, auto)
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2323
apply (rule_tac y =
23177
3004310c95b1 replace (- 1) with -1
huffman
parents: 23176
diff changeset
  2324
 "\<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
  2325
       in order_less_trans)
32047
c141f139ce26 Changed fact_Suc_nat back to fact_Suc
avigad
parents: 32036
diff changeset
  2326
apply (simp (no_asm) add: fact_num_eq_if_nat realpow_num_eq_if del: fact_Suc)
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15546
diff changeset
  2327
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
  2328
apply (rule sumr_pos_lt_pair)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2329
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
  2330
unfolding One_nat_def
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2331
apply (simp (no_asm) add: divide_inverse real_0_less_add_iff mult_assoc [symmetric]
32047
c141f139ce26 Changed fact_Suc_nat back to fact_Suc
avigad
parents: 32036
diff changeset
  2332
            del: fact_Suc)
46240
933f35c4e126 factor-cancellation simprocs now call the full simplifier to prove that factors are non-zero
huffman
parents: 45915
diff changeset
  2333
apply (simp add: inverse_eq_divide less_divide_eq del: fact_Suc)
32047
c141f139ce26 Changed fact_Suc_nat back to fact_Suc
avigad
parents: 32036
diff changeset
  2334
apply (subst fact_Suc [of "Suc (Suc (Suc (Suc (Suc (Suc (Suc (4 * d)))))))"])
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15383
diff changeset
  2335
apply (simp only: real_of_nat_mult)
23007
e025695d9b0e use mult_strict_mono instead of real_mult_less_mono
huffman
parents: 22998
diff changeset
  2336
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
  2337
  apply (rule_tac [3] real_of_nat_ge_zero)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15383
diff changeset
  2338
 prefer 2 apply force
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2339
apply (rule real_of_nat_less_iff [THEN iffD2])
32036
8a9228872fbd Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents: 31881
diff changeset
  2340
apply (rule fact_less_mono_nat, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2341
done
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2342
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2343
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
  2344
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
  2345
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2346
lemma cos_is_zero: "EX! x. 0 \<le> x & x \<le> 2 & cos x = 0"
44730
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2347
proof (rule ex_ex1I)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2348
  show "\<exists>x. 0 \<le> x & x \<le> 2 & cos x = 0"
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2349
    by (rule IVT2, simp_all)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2350
next
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2351
  fix x y
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2352
  assume x: "0 \<le> x \<and> x \<le> 2 \<and> cos x = 0"
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2353
  assume y: "0 \<le> y \<and> y \<le> 2 \<and> cos y = 0"
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2354
  have [simp]: "\<forall>x. cos differentiable x"
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2355
    unfolding differentiable_def by (auto intro: DERIV_cos)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2356
  from x y show "x = y"
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2357
    apply (cut_tac less_linear [of x y], auto)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2358
    apply (drule_tac f = cos in Rolle)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2359
    apply (drule_tac [5] f = cos in Rolle)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2360
    apply (auto dest!: DERIV_cos [THEN DERIV_unique])
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2361
    apply (metis order_less_le_trans less_le sin_gt_zero)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2362
    apply (metis order_less_le_trans less_le sin_gt_zero)
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2363
    done
11a1290fd0ac convert lemma cos_is_zero to Isar-style
huffman
parents: 44728
diff changeset
  2364
qed
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31790
diff changeset
  2365
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2366
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
  2367
by (simp add: pi_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2368
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2369
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
  2370
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
  2371
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2372
lemma pi_half_gt_zero [simp]: "0 < pi / 2"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2373
apply (rule order_le_neq_trans)
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2374
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
  2375
apply (rule notI, drule arg_cong [where f=cos], simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2376
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2377
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2378
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
  2379
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
  2380
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2381
lemma pi_half_less_two [simp]: "pi / 2 < 2"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2382
apply (rule order_le_neq_trans)
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2383
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
  2384
apply (rule notI, drule arg_cong [where f=cos], simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2385
done
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2386
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2387
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
  2388
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
  2389
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2390
lemma pi_gt_zero [simp]: "0 < pi"
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2391
by (insert pi_half_gt_zero, simp)
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2392
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2393
lemma pi_ge_zero [simp]: "0 \<le> pi"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2394
by (rule pi_gt_zero [THEN order_less_imp_le])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2395
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2396
lemma pi_neq_zero [simp]: "pi \<noteq> 0"
22998
97e1f9c2cc46 avoid using redundant lemmas from RealDef.thy
huffman
parents: 22978
diff changeset
  2397
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
  2398
23053
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2399
lemma pi_not_less_zero [simp]: "\<not> pi < 0"
03fe1dafa418 define pi with THE instead of SOME; cleaned up
huffman
parents: 23052
diff changeset
  2400
by (simp add: linorder_not_less)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2401
29165
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  2402
lemma minus_pi_half_less_zero: "-(pi/2) < 0"
562f95f06244 cleaned up some proofs; removed redundant simp rules
huffman
parents: 29164
diff changeset
  2403
by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2404
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2405
lemma m2pi_less_pi: "- (2 * pi) < pi"
45308
2e84e5f0463b extend cancellation simproc patterns to cover terms like '- (2 * pi) < pi'
huffman
parents: 44756
diff changeset
  2406
by simp
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2407
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2408
lemma sin_pi_half [simp]: "sin(pi/2) = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2409
apply (cut_tac x = "pi/2" in sin_cos_squared_add2)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2410
apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two])
36970
fb3fdb4b585e remove simp attribute from square_eq_1_iff
huffman
parents: 36842
diff changeset
  2411
apply (simp add: power2_eq_1_iff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2412
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2413
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2414
lemma cos_pi [simp]: "cos pi = -1"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2415
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
  2416
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2417
lemma sin_pi [simp]: "sin pi = 0"
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2418
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
  2419
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2420
lemma sin_cos_eq: "sin x = cos (pi/2 - x)"
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  2421
by (simp add: cos_diff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2422
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2423
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
  2424
by (simp add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2425
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2426
lemma cos_sin_eq: "cos x = sin (pi/2 - x)"
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  2427
by (simp add: sin_diff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2428
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2429
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
  2430
by (simp add: sin_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2431
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2432
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
  2433
by (simp add: sin_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2434
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2435
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
  2436
by (simp add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2437
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2438
lemma sin_periodic [simp]: "sin (x + 2*pi) = sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2439
by (simp add: sin_add cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2440
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2441
lemma cos_periodic [simp]: "cos (x + 2*pi) = cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2442
by (simp add: cos_add cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2443
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2444
lemma cos_npi [simp]: "cos (real n * pi) = -1 ^ n"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15241
diff changeset
  2445
apply (induct "n")
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  2446
apply (auto simp add: real_of_nat_Suc distrib_right)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2447
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2448
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2449
lemma cos_npi2 [simp]: "cos (pi * real n) = -1 ^ n"
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2450
proof -
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2451
  have "cos (pi * real n) = cos (real n * pi)" by (simp only: mult_commute)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2452
  also have "... = -1 ^ n" by (rule cos_npi)
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2453
  finally show ?thesis .
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2454
qed
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  2455
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2456
lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0"
15251
bb6f072c8d10 converted some induct_tac to induct
paulson
parents: 15241
diff changeset
  2457
apply (induct "n")
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  2458
apply (auto simp add: real_of_nat_Suc distrib_right)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2459
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2460
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2461
lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2462
by (simp add: mult_commute [of pi])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2463
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2464
lemma cos_two_pi [simp]: "cos (2 * pi) = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2465
by (simp add: cos_double)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2466
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2467
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
  2468
by simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2469
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2470
lemma sin_gt_zero2: "[| 0 < x; x < pi/2 |] ==> 0 < sin x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2471
apply (rule sin_gt_zero, assumption)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2472
apply (rule order_less_trans, assumption)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2473
apply (rule pi_half_less_two)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2474
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2475
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2476
lemma sin_less_zero:
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2477
  assumes lb: "- pi/2 < x" and "x < 0" shows "sin x < 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2478
proof -
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2479
  have "0 < sin (- x)" using assms by (simp only: sin_gt_zero2)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2480
  thus ?thesis by simp
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2481
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2482
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2483
lemma pi_less_4: "pi < 4"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2484
by (cut_tac pi_half_less_two, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2485
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2486
lemma cos_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < cos x"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2487
apply (cut_tac pi_less_4)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2488
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
  2489
apply (cut_tac cos_is_zero, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2490
apply (rename_tac y z)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2491
apply (drule_tac x = y in spec)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2492
apply (drule_tac x = "pi/2" in spec, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2493
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2494
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2495
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
  2496
apply (rule_tac x = x and y = 0 in linorder_cases)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2497
apply (rule cos_minus [THEN subst])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2498
apply (rule cos_gt_zero)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2499
apply (auto intro: cos_gt_zero)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2500
done
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2501
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2502
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
  2503
apply (auto simp add: order_le_less cos_gt_zero_pi)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2504
apply (subgoal_tac "x = pi/2", auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2505
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2506
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2507
lemma sin_gt_zero_pi: "[| 0 < x; x < pi  |] ==> 0 < sin x"
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  2508
by (simp add: sin_cos_eq cos_gt_zero_pi)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2509
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2510
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
  2511
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
  2512
  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
  2513
  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
  2514
  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
  2515
    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
  2516
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2517
    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
  2518
    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
  2519
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2520
  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
  2521
  hence "0 < sin y" using sin_gt_zero by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2522
  moreover
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2523
  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
  2524
  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
  2525
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2526
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2527
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
  2528
by (auto simp add: order_le_less sin_gt_zero_pi)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2529
44745
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2530
text {* FIXME: This proof is almost identical to lemma @{text cos_is_zero}.
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2531
  It should be possible to factor out some of the common parts. *}
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2532
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2533
lemma cos_total: "[| -1 \<le> y; y \<le> 1 |] ==> EX! x. 0 \<le> x & x \<le> pi & (cos x = y)"
44745
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2534
proof (rule ex_ex1I)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2535
  assume y: "-1 \<le> y" "y \<le> 1"
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2536
  show "\<exists>x. 0 \<le> x & x \<le> pi & cos x = y"
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2537
    by (rule IVT2, simp_all add: y)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2538
next
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2539
  fix a b
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2540
  assume a: "0 \<le> a \<and> a \<le> pi \<and> cos a = y"
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2541
  assume b: "0 \<le> b \<and> b \<le> pi \<and> cos b = y"
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2542
  have [simp]: "\<forall>x. cos differentiable x"
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2543
    unfolding differentiable_def by (auto intro: DERIV_cos)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2544
  from a b show "a = b"
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2545
    apply (cut_tac less_linear [of a b], auto)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2546
    apply (drule_tac f = cos in Rolle)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2547
    apply (drule_tac [5] f = cos in Rolle)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2548
    apply (auto dest!: DERIV_cos [THEN DERIV_unique])
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2549
    apply (metis order_less_le_trans less_le sin_gt_zero_pi)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2550
    apply (metis order_less_le_trans less_le sin_gt_zero_pi)
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2551
    done
b068207a7400 convert lemma cos_total to Isar-style proof
huffman
parents: 44730
diff changeset
  2552
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2553
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2554
lemma sin_total:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2555
     "[| -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
  2556
apply (rule ccontr)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2557
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
  2558
apply (erule contrapos_np)
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  2559
apply simp
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2560
apply (cut_tac y="-y" in cos_total, simp) apply simp
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2561
apply (erule ex1E)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2562
apply (rule_tac a = "x - (pi/2)" in ex1I)
23286
huffman
parents: 23278
diff changeset
  2563
apply (simp (no_asm) add: add_assoc)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2564
apply (rotate_tac 3)
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  2565
apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2566
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2567
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2568
lemma reals_Archimedean4:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2569
     "[| 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
  2570
apply (auto dest!: reals_Archimedean3)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2571
apply (drule_tac x = x in spec, clarify)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2572
apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y")
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2573
 prefer 2 apply (erule LeastI)
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2574
apply (case_tac "LEAST m::nat. x < real m * y", simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2575
apply (subgoal_tac "~ x < real nat * y")
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2576
 prefer 2 apply (rule not_less_Least, simp, force)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2577
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2578
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2579
(* Pre Isabelle99-2 proof was simpler- numerals arithmetic
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2580
   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
  2581
lemma cos_zero_lemma:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2582
     "[| 0 \<le> x; cos x = 0 |] ==>
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2583
      \<exists>n::nat. ~even n & x = real n * (pi/2)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2584
apply (drule pi_gt_zero [THEN reals_Archimedean4], safe)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2585
apply (subgoal_tac "0 \<le> x - real n * pi &
15086
e6a2a98d5ef5 removal of more iff-rules from RealDef.thy
paulson
parents: 15085
diff changeset
  2586
                    (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
  2587
apply (auto simp add: algebra_simps real_of_nat_Suc)
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  2588
 prefer 2 apply (simp add: cos_diff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2589
apply (simp add: cos_diff)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2590
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
  2591
apply (rule_tac [2] cos_total, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2592
apply (drule_tac x = "x - real n * pi" in spec)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2593
apply (drule_tac x = "pi/2" in spec)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2594
apply (simp add: cos_diff)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2595
apply (rule_tac x = "Suc (2 * n)" in exI)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  2596
apply (simp add: real_of_nat_Suc algebra_simps, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2597
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2598
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2599
lemma sin_zero_lemma:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2600
     "[| 0 \<le> x; sin x = 0 |] ==>
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2601
      \<exists>n::nat. even n & x = real n * (pi/2)"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2602
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
  2603
 apply (clarify, rule_tac x = "n - 1" in exI)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  2604
 apply (force simp add: odd_Suc_mult_two_ex real_of_nat_Suc distrib_right)
15085
5693a977a767 removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents: 15081
diff changeset
  2605
apply (rule cos_zero_lemma)
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  2606
apply (simp_all add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2607
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2608
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2609
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2610
lemma cos_zero_iff:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2611
     "(cos x = 0) =
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2612
      ((\<exists>n::nat. ~even n & (x = real n * (pi/2))) |
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2613
       (\<exists>n::nat. ~even n & (x = -(real n * (pi/2)))))"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2614
apply (rule iffI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2615
apply (cut_tac linorder_linear [of 0 x], safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2616
apply (drule cos_zero_lemma, assumption+)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2617
apply (cut_tac x="-x" in cos_zero_lemma, simp, simp)
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2618
apply (force simp add: minus_equation_iff [of x])
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  2619
apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc distrib_right)
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2620
apply (auto simp add: cos_add)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2621
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2622
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2623
(* ditto: but to a lesser extent *)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2624
lemma sin_zero_iff:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2625
     "(sin x = 0) =
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2626
      ((\<exists>n::nat. even n & (x = real n * (pi/2))) |
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2627
       (\<exists>n::nat. even n & (x = -(real n * (pi/2)))))"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2628
apply (rule iffI)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2629
apply (cut_tac linorder_linear [of 0 x], safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2630
apply (drule sin_zero_lemma, assumption+)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2631
apply (cut_tac x="-x" in sin_zero_lemma, simp, simp, safe)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2632
apply (force simp add: minus_equation_iff [of x])
15539
333a88244569 comprehensive cleanup, replacing sumr by setsum
nipkow
parents: 15536
diff changeset
  2633
apply (auto simp add: even_mult_two_ex)
15077
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
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2636
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
  2637
  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
  2638
proof -
33549
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
  2639
  have "- (x - y) < 0" using assms 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
  2640
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2641
  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
  2642
  obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto
33549
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
  2643
  hence "0 < z" and "z < pi" using assms 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
  2644
  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
  2645
  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
  2646
  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
  2647
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2648
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2649
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
  2650
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
  2651
  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
  2652
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2653
  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
  2654
  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
  2655
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2656
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2657
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
  2658
  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
  2659
proof -
33549
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
  2660
  have "0 \<le> -x" and "-x < -y" and "-y \<le> pi" using assms 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
  2661
  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
  2662
  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
  2663
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2664
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2665
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
  2666
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
  2667
  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
  2668
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2669
  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
  2670
  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
  2671
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2672
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2673
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
  2674
proof -
33549
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
  2675
  have "0 \<le> y + pi / 2" and "y + pi / 2 \<le> x + pi / 2" and "x + pi /2 \<le> pi"
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
  2676
    using pi_ge_two and assms 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
  2677
  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
  2678
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2679
29164
0d49c5b55046 move sin and cos to their own subsection
huffman
parents: 29163
diff changeset
  2680
subsection {* Tangent *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2681
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2682
definition tan :: "real \<Rightarrow> real" where
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2683
  "tan = (\<lambda>x. sin x / cos x)"
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2684
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2685
lemma tan_zero [simp]: "tan 0 = 0"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2686
  by (simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2687
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2688
lemma tan_pi [simp]: "tan pi = 0"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2689
  by (simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2690
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2691
lemma tan_npi [simp]: "tan (real (n::nat) * pi) = 0"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2692
  by (simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2693
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2694
lemma tan_minus [simp]: "tan (-x) = - tan x"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2695
  by (simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2696
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2697
lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2698
  by (simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2699
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2700
lemma lemma_tan_add1:
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2701
  "\<lbrakk>cos x \<noteq> 0; cos y \<noteq> 0\<rbrakk> \<Longrightarrow> 1 - tan x * tan y = cos (x + y)/(cos x * cos y)"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2702
  by (simp add: tan_def cos_add field_simps)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2703
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2704
lemma add_tan_eq:
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2705
  "\<lbrakk>cos x \<noteq> 0; cos y \<noteq> 0\<rbrakk> \<Longrightarrow> tan x + tan y = sin(x + y)/(cos x * cos y)"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2706
  by (simp add: tan_def sin_add field_simps)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2707
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2708
lemma tan_add:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2709
     "[| 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
  2710
      ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2711
  by (simp add: add_tan_eq lemma_tan_add1, simp add: tan_def)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2712
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2713
lemma tan_double:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2714
     "[| cos x \<noteq> 0; cos (2 * x) \<noteq> 0 |]
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2715
      ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))"
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2716
  using tan_add [of x x] by (simp add: power2_eq_square)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2717
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2718
lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x"
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2719
by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi)
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2720
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2721
lemma tan_less_zero:
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2722
  assumes lb: "- pi/2 < x" and "x < 0" shows "tan x < 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2723
proof -
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2724
  have "0 < tan (- x)" using assms by (simp only: tan_gt_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2725
  thus ?thesis by simp
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2726
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2727
44756
efcd71fbaeec simplify proof of tan_half, removing unused assumptions
huffman
parents: 44755
diff changeset
  2728
lemma tan_half: "tan x = sin (2 * x) / (cos (2 * x) + 1)"
efcd71fbaeec simplify proof of tan_half, removing unused assumptions
huffman
parents: 44755
diff changeset
  2729
  unfolding tan_def sin_double cos_double sin_squared_eq
efcd71fbaeec simplify proof of tan_half, removing unused assumptions
huffman
parents: 44755
diff changeset
  2730
  by (simp add: power2_eq_square)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2731
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2732
lemma DERIV_tan [simp]: "cos x \<noteq> 0 \<Longrightarrow> DERIV tan x :> inverse ((cos x)\<twosuperior>)"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2733
  unfolding tan_def
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2734
  by (auto intro!: DERIV_intros, simp add: divide_inverse power2_eq_square)
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2735
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2736
lemma isCont_tan: "cos x \<noteq> 0 \<Longrightarrow> isCont tan x"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2737
  by (rule DERIV_tan [THEN DERIV_isCont])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2738
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2739
lemma isCont_tan' [simp]:
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2740
  "\<lbrakk>isCont f a; cos (f a) \<noteq> 0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. tan (f x)) a"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2741
  by (rule isCont_o2 [OF _ isCont_tan])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2742
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2743
lemma tendsto_tan [tendsto_intros]:
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2744
  "\<lbrakk>(f ---> a) F; cos a \<noteq> 0\<rbrakk> \<Longrightarrow> ((\<lambda>x. tan (f x)) ---> tan a) F"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2745
  by (rule isCont_tendsto_compose [OF isCont_tan])
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2746
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2747
lemma continuous_tan:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2748
  "continuous F f \<Longrightarrow> cos (f (Lim F (\<lambda>x. x))) \<noteq> 0 \<Longrightarrow> continuous F (\<lambda>x. tan (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2749
  unfolding continuous_def by (rule tendsto_tan)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2750
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2751
lemma isCont_tan'' [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2752
  "continuous (at x) f \<Longrightarrow> cos (f x) \<noteq> 0 \<Longrightarrow> continuous (at x) (\<lambda>x. tan (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2753
  unfolding continuous_at by (rule tendsto_tan)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2754
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2755
lemma continuous_within_tan [continuous_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2756
  "continuous (at x within s) f \<Longrightarrow> cos (f x) \<noteq> 0 \<Longrightarrow> continuous (at x within s) (\<lambda>x. tan (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2757
  unfolding continuous_within by (rule tendsto_tan)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2758
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2759
lemma continuous_on_tan [continuous_on_intros]:
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2760
  "continuous_on s f \<Longrightarrow> (\<forall>x\<in>s. cos (f x) \<noteq> 0) \<Longrightarrow> continuous_on s (\<lambda>x. tan (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2761
  unfolding continuous_on_def by (auto intro: tendsto_tan)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  2762
44311
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2763
lemma LIM_cos_div_sin: "(%x. cos(x)/sin(x)) -- pi/2 --> 0"
42c5cbf68052 Transcendental.thy: add tendsto_intros lemmas;
huffman
parents: 44308
diff changeset
  2764
  by (rule LIM_cong_limit, (rule tendsto_intros)+, simp_all)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2765
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2766
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
  2767
apply (cut_tac LIM_cos_div_sin)
31338
d41a8ba25b67 generalize constants from Lim.thy to class metric_space
huffman
parents: 31271
diff changeset
  2768
apply (simp only: LIM_eq)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2769
apply (drule_tac x = "inverse y" in spec, safe, force)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2770
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
  2771
apply (rule_tac x = "(pi/2) - e" in exI)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2772
apply (simp (no_asm_simp))
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2773
apply (drule_tac x = "(pi/2) - e" in spec)
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  2774
apply (auto simp add: tan_def sin_diff cos_diff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2775
apply (rule inverse_less_iff_less [THEN iffD1])
15079
2ef899e4526d conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents: 15077
diff changeset
  2776
apply (auto simp add: divide_inverse)
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
  2777
apply (rule mult_pos_pos)
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2778
apply (subgoal_tac [3] "0 < sin e & 0 < cos e")
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
  2779
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
  2780
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2781
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2782
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
  2783
apply (frule order_le_imp_less_or_eq, safe)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2784
 prefer 2 apply force
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2785
apply (drule lemma_tan_total, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2786
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
  2787
apply (auto intro!: DERIV_tan [THEN DERIV_isCont])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2788
apply (drule_tac y = xa in order_le_imp_less_or_eq)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2789
apply (auto dest: cos_gt_zero)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2790
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2791
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2792
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
  2793
apply (cut_tac linorder_linear [of 0 y], safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2794
apply (drule tan_total_pos)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2795
apply (cut_tac [2] y="-y" in tan_total_pos, safe)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2796
apply (rule_tac [3] x = "-x" in exI)
44710
9caf6883f1f4 remove redundant lemmas about LIMSEQ
huffman
parents: 44568
diff changeset
  2797
apply (auto del: exI intro!: exI)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2798
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2799
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2800
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
  2801
apply (cut_tac y = y in lemma_tan_total1, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2802
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
  2803
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
  2804
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
  2805
apply (rule_tac [4] Rolle)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2806
apply (rule_tac [2] Rolle)
44710
9caf6883f1f4 remove redundant lemmas about LIMSEQ
huffman
parents: 44568
diff changeset
  2807
apply (auto del: exI intro!: DERIV_tan DERIV_isCont exI
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2808
            simp add: differentiable_def)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2809
txt{*Now, simulate TRYALL*}
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2810
apply (rule_tac [!] DERIV_tan asm_rl)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2811
apply (auto dest!: DERIV_unique [OF _ DERIV_tan]
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2812
            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
  2813
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2814
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2815
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
  2816
  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
  2817
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2818
  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
  2819
  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
  2820
    fix x' :: real assume "y \<le> x' \<and> x' \<le> x"
33549
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
  2821
    hence "-(pi/2) < x'" and "x' < pi/2" using assms 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
  2822
    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
  2823
    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
  2824
    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
  2825
  qed
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2826
  from MVT2[OF `y < x` this]
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2827
  obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<twosuperior>)" by auto
33549
39f2855ce41b tuned proofs;
wenzelm
parents: 32960
diff changeset
  2828
  hence "- (pi / 2) < z" and "z < pi / 2" using assms 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
  2829
  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
  2830
  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
  2831
  have "0 < x - y" using `y < x` by auto
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
  2832
  from mult_pos_pos [OF this inv_pos]
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2833
  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
  2834
  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
  2835
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2836
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2837
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
  2838
  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
  2839
proof
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2840
  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
  2841
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2842
  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
  2843
  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
  2844
  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
  2845
    assume "\<not> y < x" hence "x \<le> y" by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2846
    hence "tan x \<le> tan y"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2847
    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
  2848
      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
  2849
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2850
      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
  2851
      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
  2852
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2853
    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
  2854
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2855
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2856
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2857
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
  2858
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2859
lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2860
  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
  2861
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2862
lemma tan_periodic_nat[simp]: fixes n :: nat shows "tan (x + real n * pi) = tan x"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2863
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
  2864
  case (Suc n)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  2865
  have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_eq_plus1 real_of_nat_add real_of_one distrib_right 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
  2866
  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
  2867
qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2868
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2869
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
  2870
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
  2871
  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
  2872
  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
  2873
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2874
  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
  2875
  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
  2876
  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
  2877
  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
  2878
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  2879
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46240
diff changeset
  2880
lemma tan_periodic_n[simp]: "tan (x + numeral n * pi) = tan x"
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46240
diff changeset
  2881
  using tan_periodic_int[of _ "numeral n" ] unfolding real_numeral .
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2882
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2883
subsection {* Inverse Trigonometric Functions *}
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2884
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2885
definition
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2886
  arcsin :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2887
  "arcsin y = (THE x. -(pi/2) \<le> x & x \<le> pi/2 & sin x = y)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2888
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2889
definition
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2890
  arccos :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2891
  "arccos y = (THE x. 0 \<le> x & x \<le> pi & cos x = y)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2892
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2893
definition
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2894
  arctan :: "real => real" where
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2895
  "arctan y = (THE x. -(pi/2) < x & x < pi/2 & tan x = y)"
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  2896
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  2897
lemma arcsin:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2898
     "[| -1 \<le> y; y \<le> 1 |]
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2899
      ==> -(pi/2) \<le> arcsin y &
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2900
           arcsin y \<le> pi/2 & sin(arcsin y) = y"
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2901
unfolding arcsin_def by (rule theI' [OF sin_total])
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2902
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2903
lemma arcsin_pi:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2904
     "[| -1 \<le> y; y \<le> 1 |]
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2905
      ==> -(pi/2) \<le> arcsin y & arcsin y \<le> pi & sin(arcsin y) = y"
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2906
apply (drule (1) arcsin)
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2907
apply (force intro: order_trans)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2908
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2909
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2910
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
  2911
by (blast dest: arcsin)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2912
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2913
lemma arcsin_bounded:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2914
     "[| -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
  2915
by (blast dest: arcsin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2916
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2917
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
  2918
by (blast dest: arcsin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2919
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2920
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
  2921
by (blast dest: arcsin)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2922
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2923
lemma arcsin_lt_bounded:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2924
     "[| -1 < y; y < 1 |] ==> -(pi/2) < arcsin y & arcsin y < pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2925
apply (frule order_less_imp_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2926
apply (frule_tac y = y in order_less_imp_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2927
apply (frule arcsin_bounded)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2928
apply (safe, simp)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2929
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
  2930
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
  2931
apply (drule_tac [!] f = sin in arg_cong, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2932
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2933
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2934
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
  2935
apply (unfold arcsin_def)
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2936
apply (rule the1_equality)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2937
apply (rule sin_total, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2938
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2939
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2940
lemma arccos:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2941
     "[| -1 \<le> y; y \<le> 1 |]
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2942
      ==> 0 \<le> arccos y & arccos y \<le> pi & cos(arccos y) = y"
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2943
unfolding arccos_def by (rule theI' [OF cos_total])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2944
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2945
lemma cos_arccos [simp]: "[| -1 \<le> y; y \<le> 1 |] ==> cos(arccos y) = y"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2946
by (blast dest: arccos)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2947
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2948
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
  2949
by (blast dest: arccos)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2950
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2951
lemma arccos_lbound: "[| -1 \<le> y; y \<le> 1 |] ==> 0 \<le> arccos y"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2952
by (blast dest: arccos)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2953
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2954
lemma arccos_ubound: "[| -1 \<le> y; y \<le> 1 |] ==> arccos y \<le> pi"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2955
by (blast dest: arccos)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2956
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2957
lemma arccos_lt_bounded:
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  2958
     "[| -1 < y; y < 1 |]
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2959
      ==> 0 < arccos y & arccos y < pi"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2960
apply (frule order_less_imp_le)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2961
apply (frule_tac y = y in order_less_imp_le)
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2962
apply (frule arccos_bounded, auto)
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2963
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
  2964
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
  2965
apply (drule_tac [!] f = cos in arg_cong, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2966
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2967
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2968
lemma arccos_cos: "[|0 \<le> x; x \<le> pi |] ==> arccos(cos x) = x"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2969
apply (simp add: arccos_def)
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2970
apply (auto intro!: the1_equality cos_total)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2971
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2972
22975
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2973
lemma arccos_cos2: "[|x \<le> 0; -pi \<le> x |] ==> arccos(cos x) = -x"
03085c441c14 spelling: rename arcos -> arccos
huffman
parents: 22969
diff changeset
  2974
apply (simp add: arccos_def)
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  2975
apply (auto intro!: the1_equality cos_total)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2976
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  2977
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2978
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
  2979
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
  2980
apply (rule power2_eq_imp_eq)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2981
apply (simp add: cos_squared_eq)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2982
apply (rule cos_ge_zero)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2983
apply (erule (1) arcsin_lbound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2984
apply (erule (1) arcsin_ubound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2985
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2986
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
  2987
apply (rule power_mono, simp, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2988
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2989
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2990
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
  2991
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
  2992
apply (rule power2_eq_imp_eq)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2993
apply (simp add: sin_squared_eq)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2994
apply (rule sin_ge_zero)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2995
apply (erule (1) arccos_lbound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2996
apply (erule (1) arccos_ubound)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2997
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  2998
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
  2999
apply (rule power_mono, simp, simp)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3000
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3001
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3002
lemma arctan [simp]:
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3003
     "- (pi/2) < arctan y  & arctan y < pi/2 & tan (arctan y) = y"
23011
3eae3140b4b2 use THE instead of SOME
huffman
parents: 23007
diff changeset
  3004
unfolding arctan_def by (rule theI' [OF tan_total])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3005
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3006
lemma tan_arctan: "tan(arctan y) = y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3007
by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3008
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3009
lemma arctan_bounded: "- (pi/2) < arctan y  & arctan y < pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3010
by (auto simp only: arctan)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3011
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3012
lemma arctan_lbound: "- (pi/2) < arctan y"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3013
by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3014
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3015
lemma arctan_ubound: "arctan y < pi/2"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3016
by (auto simp only: arctan)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3017
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3018
lemma arctan_unique:
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3019
  assumes "-(pi/2) < x" and "x < pi/2" and "tan x = y"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3020
  shows "arctan y = x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3021
  using assms arctan [of y] tan_total [of y] by (fast elim: ex1E)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3022
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3023
lemma arctan_tan:
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3024
      "[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x"
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3025
  by (rule arctan_unique, simp_all)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3026
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3027
lemma arctan_zero_zero [simp]: "arctan 0 = 0"
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3028
  by (rule arctan_unique, simp_all)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3029
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3030
lemma arctan_minus: "arctan (- x) = - arctan x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3031
  apply (rule arctan_unique)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3032
  apply (simp only: neg_less_iff_less arctan_ubound)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3033
  apply (metis minus_less_iff arctan_lbound)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3034
  apply simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3035
  done
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3036
44725
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3037
lemma cos_arctan_not_zero [simp]: "cos (arctan x) \<noteq> 0"
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3038
  by (intro less_imp_neq [symmetric] cos_gt_zero_pi
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3039
    arctan_lbound arctan_ubound)
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3040
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3041
lemma cos_arctan: "cos (arctan x) = 1 / sqrt (1 + x\<twosuperior>)"
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3042
proof (rule power2_eq_imp_eq)
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3043
  have "0 < 1 + x\<twosuperior>" by (simp add: add_pos_nonneg)
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3044
  show "0 \<le> 1 / sqrt (1 + x\<twosuperior>)" by simp
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3045
  show "0 \<le> cos (arctan x)"
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3046
    by (intro less_imp_le cos_gt_zero_pi arctan_lbound arctan_ubound)
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3047
  have "(cos (arctan x))\<twosuperior> * (1 + (tan (arctan x))\<twosuperior>) = 1"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  3048
    unfolding tan_def by (simp add: distrib_left power_divide)
44725
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3049
  thus "(cos (arctan x))\<twosuperior> = (1 / sqrt (1 + x\<twosuperior>))\<twosuperior>"
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3050
    using `0 < 1 + x\<twosuperior>` by (simp add: power_divide eq_divide_eq)
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3051
qed
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3052
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3053
lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<twosuperior>)"
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3054
  using add_pos_nonneg [OF zero_less_one zero_le_power2 [of x]]
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3055
  using tan_arctan [of x] unfolding tan_def cos_arctan
d3bf0e33c98a add lemmas cos_arctan and sin_arctan
huffman
parents: 44710
diff changeset
  3056
  by (simp add: eq_divide_eq)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3057
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3058
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
  3059
apply (rule power_inverse [THEN subst])
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3060
apply (rule_tac c1 = "(cos x)\<twosuperior>" in real_mult_right_cancel [THEN iffD1])
22960
114cf1906681 remove redundant lemmas
huffman
parents: 22956
diff changeset
  3061
apply (auto dest: field_power_not_zero
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  3062
        simp add: power_mult_distrib distrib_right power_divide tan_def
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 30082
diff changeset
  3063
                  mult_assoc power_inverse [symmetric])
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3064
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3065
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3066
lemma arctan_less_iff: "arctan x < arctan y \<longleftrightarrow> x < y"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3067
  by (metis tan_monotone' arctan_lbound arctan_ubound tan_arctan)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3068
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3069
lemma arctan_le_iff: "arctan x \<le> arctan y \<longleftrightarrow> x \<le> y"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3070
  by (simp only: not_less [symmetric] arctan_less_iff)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3071
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3072
lemma arctan_eq_iff: "arctan x = arctan y \<longleftrightarrow> x = y"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3073
  by (simp only: eq_iff [where 'a=real] arctan_le_iff)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3074
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3075
lemma zero_less_arctan_iff [simp]: "0 < arctan x \<longleftrightarrow> 0 < x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3076
  using arctan_less_iff [of 0 x] by simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3077
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3078
lemma arctan_less_zero_iff [simp]: "arctan x < 0 \<longleftrightarrow> x < 0"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3079
  using arctan_less_iff [of x 0] by simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3080
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3081
lemma zero_le_arctan_iff [simp]: "0 \<le> arctan x \<longleftrightarrow> 0 \<le> x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3082
  using arctan_le_iff [of 0 x] by simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3083
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3084
lemma arctan_le_zero_iff [simp]: "arctan x \<le> 0 \<longleftrightarrow> x \<le> 0"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3085
  using arctan_le_iff [of x 0] by simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3086
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3087
lemma arctan_eq_zero_iff [simp]: "arctan x = 0 \<longleftrightarrow> x = 0"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3088
  using arctan_eq_iff [of x 0] by simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3089
51482
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3090
lemma continuous_on_arcsin': "continuous_on {-1 .. 1} arcsin"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3091
proof -
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3092
  have "continuous_on (sin ` {- pi / 2 .. pi / 2}) arcsin"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3093
    by (rule continuous_on_inv) (auto intro: continuous_on_intros simp: arcsin_sin)
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3094
  also have "sin ` {- pi / 2 .. pi / 2} = {-1 .. 1}"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3095
  proof safe
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3096
    fix x :: real assume "x \<in> {-1..1}" then show "x \<in> sin ` {- pi / 2..pi / 2}"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3097
      using arcsin_lbound arcsin_ubound by (intro image_eqI[where x="arcsin x"]) auto
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3098
  qed simp
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3099
  finally show ?thesis .
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3100
qed
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3101
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3102
lemma continuous_on_arcsin [continuous_on_intros]:
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3103
  "continuous_on s f \<Longrightarrow> (\<forall>x\<in>s. -1 \<le> f x \<and> f x \<le> 1) \<Longrightarrow> continuous_on s (\<lambda>x. arcsin (f x))"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3104
  using continuous_on_compose[of s f, OF _ continuous_on_subset[OF  continuous_on_arcsin']]
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3105
  by (auto simp: comp_def subset_eq)
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3106
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3107
lemma isCont_arcsin: "-1 < x \<Longrightarrow> x < 1 \<Longrightarrow> isCont arcsin x"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3108
  using continuous_on_arcsin'[THEN continuous_on_subset, of "{ -1 <..< 1 }"]
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3109
  by (auto simp: continuous_on_eq_continuous_at subset_eq)
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3110
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3111
lemma continuous_on_arccos': "continuous_on {-1 .. 1} arccos"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3112
proof -
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3113
  have "continuous_on (cos ` {0 .. pi}) arccos"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3114
    by (rule continuous_on_inv) (auto intro: continuous_on_intros simp: arccos_cos)
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3115
  also have "cos ` {0 .. pi} = {-1 .. 1}"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3116
  proof safe
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3117
    fix x :: real assume "x \<in> {-1..1}" then show "x \<in> cos ` {0..pi}"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3118
      using arccos_lbound arccos_ubound by (intro image_eqI[where x="arccos x"]) auto
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3119
  qed simp
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3120
  finally show ?thesis .
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3121
qed
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3122
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3123
lemma continuous_on_arccos [continuous_on_intros]:
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3124
  "continuous_on s f \<Longrightarrow> (\<forall>x\<in>s. -1 \<le> f x \<and> f x \<le> 1) \<Longrightarrow> continuous_on s (\<lambda>x. arccos (f x))"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3125
  using continuous_on_compose[of s f, OF _ continuous_on_subset[OF  continuous_on_arccos']]
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3126
  by (auto simp: comp_def subset_eq)
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3127
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3128
lemma isCont_arccos: "-1 < x \<Longrightarrow> x < 1 \<Longrightarrow> isCont arccos x"
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3129
  using continuous_on_arccos'[THEN continuous_on_subset, of "{ -1 <..< 1 }"]
80efd8c49f52 arcsin and arccos are continuous on {0 .. 1} (including the endpoints)
hoelzl
parents: 51481
diff changeset
  3130
  by (auto simp: continuous_on_eq_continuous_at subset_eq)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3131
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3132
lemma isCont_arctan: "isCont arctan x"
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3133
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
  3134
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
  3135
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
  3136
apply (erule (1) isCont_inverse_function2 [where f=tan])
33667
958dc9f03611 A little rationalisation
paulson
parents: 33549
diff changeset
  3137
apply (metis arctan_tan order_le_less_trans order_less_le_trans)
36777
be5461582d0f avoid using real-specific versions of generic lemmas
huffman
parents: 36776
diff changeset
  3138
apply (metis cos_gt_zero_pi isCont_tan order_less_le_trans less_le)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3139
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3140
51478
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3141
lemma tendsto_arctan [tendsto_intros]: "(f ---> x) F \<Longrightarrow> ((\<lambda>x. arctan (f x)) ---> arctan x) F"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3142
  by (rule isCont_tendsto_compose [OF isCont_arctan])
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3143
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3144
lemma continuous_arctan [continuous_intros]: "continuous F f \<Longrightarrow> continuous F (\<lambda>x. arctan (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3145
  unfolding continuous_def by (rule tendsto_arctan)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3146
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3147
lemma continuous_on_arctan [continuous_on_intros]: "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. arctan (f x))"
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3148
  unfolding continuous_on_def by (auto intro: tendsto_arctan)
270b21f3ae0a move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
hoelzl
parents: 51477
diff changeset
  3149
  
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3150
lemma DERIV_arcsin:
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3151
  "\<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
  3152
apply (rule DERIV_inverse_function [where f=sin and a="-1" and b="1"])
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  3153
apply (rule DERIV_cong [OF DERIV_sin])
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3154
apply (simp add: cos_arcsin)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3155
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
  3156
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
  3157
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3158
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3159
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3160
apply (erule (1) isCont_arcsin)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3161
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3162
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3163
lemma DERIV_arccos:
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3164
  "\<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
  3165
apply (rule DERIV_inverse_function [where f=cos and a="-1" and b="1"])
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  3166
apply (rule DERIV_cong [OF DERIV_cos])
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3167
apply (simp add: sin_arccos)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3168
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
  3169
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
  3170
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3171
apply assumption
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3172
apply simp
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3173
apply (erule (1) isCont_arccos)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3174
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3175
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3176
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
  3177
apply (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"])
44308
d2a6f9af02f4 Transcendental.thy: remove several unused lemmas and simplify some proofs
huffman
parents: 44307
diff changeset
  3178
apply (rule DERIV_cong [OF DERIV_tan])
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3179
apply (rule cos_arctan_not_zero)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3180
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
  3181
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
  3182
apply (simp add: add_pos_nonneg)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3183
apply (simp, simp, simp, rule isCont_arctan)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3184
done
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3185
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31790
diff changeset
  3186
declare
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31790
diff changeset
  3187
  DERIV_arcsin[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31790
diff changeset
  3188
  DERIV_arccos[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31790
diff changeset
  3189
  DERIV_arctan[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31790
diff changeset
  3190
50346
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3191
lemma filterlim_tan_at_right: "filterlim tan at_bot (at_right (- pi/2))"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3192
  by (rule filterlim_at_bot_at_right[where Q="\<lambda>x. - pi/2 < x \<and> x < pi/2" and P="\<lambda>x. True" and g=arctan])
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3193
     (auto simp: le_less eventually_within_less dist_real_def simp del: less_divide_eq_numeral1
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3194
           intro!: tan_monotone exI[of _ "pi/2"])
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3195
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3196
lemma filterlim_tan_at_left: "filterlim tan at_top (at_left (pi/2))"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3197
  by (rule filterlim_at_top_at_left[where Q="\<lambda>x. - pi/2 < x \<and> x < pi/2" and P="\<lambda>x. True" and g=arctan])
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3198
     (auto simp: le_less eventually_within_less dist_real_def simp del: less_divide_eq_numeral1
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3199
           intro!: tan_monotone exI[of _ "pi/2"])
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3200
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3201
lemma tendsto_arctan_at_top: "(arctan ---> (pi/2)) at_top"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3202
proof (rule tendstoI)
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3203
  fix e :: real assume "0 < e"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3204
  def y \<equiv> "pi/2 - min (pi/2) e"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3205
  then have y: "0 \<le> y" "y < pi/2" "pi/2 \<le> e + y"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3206
    using `0 < e` by auto
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3207
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3208
  show "eventually (\<lambda>x. dist (arctan x) (pi / 2) < e) at_top"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3209
  proof (intro eventually_at_top_dense[THEN iffD2] exI allI impI)
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3210
    fix x assume "tan y < x"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3211
    then have "arctan (tan y) < arctan x"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3212
      by (simp add: arctan_less_iff)
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3213
    with y have "y < arctan x"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3214
      by (subst (asm) arctan_tan) simp_all
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3215
    with arctan_ubound[of x, arith] y `0 < e`
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3216
    show "dist (arctan x) (pi / 2) < e"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3217
      by (simp add: dist_real_def)
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3218
  qed
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3219
qed
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3220
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3221
lemma tendsto_arctan_at_bot: "(arctan ---> - (pi/2)) at_bot"
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3222
  unfolding filterlim_at_bot_mirror arctan_minus by (intro tendsto_minus tendsto_arctan_at_top)
a75c6429c3c3 add filterlim rules for eventually monotone bijective functions; mirror rules for at_top, at_bot; apply them to prove convergence of arctan at infinity and tan at pi/2
hoelzl
parents: 50326
diff changeset
  3223
23043
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  3224
subsection {* More Theorems about Sin and Cos *}
5dbfd67516a4 rearranged sections
huffman
parents: 23011
diff changeset
  3225
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3226
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
  3227
proof -
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3228
  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
  3229
  have nonneg: "0 \<le> ?c"
45308
2e84e5f0463b extend cancellation simproc patterns to cover terms like '- (2 * pi) < pi'
huffman
parents: 44756
diff changeset
  3230
    by (simp add: cos_ge_zero)
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3231
  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
  3232
    by simp
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3233
  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
  3234
    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
  3235
  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
  3236
    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
  3237
  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
  3238
    by (simp add: power_divide)
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3239
  thus ?thesis
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3240
    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
  3241
qed
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3242
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3243
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
  3244
proof -
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3245
  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
  3246
  have pos_c: "0 < ?c"
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3247
    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
  3248
  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
  3249
    by simp
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3250
  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
  3251
    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
  3252
  also have "\<dots> = ?c * (?c\<twosuperior> - 3 * ?s\<twosuperior>)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29171
diff changeset
  3253
    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
  3254
  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
  3255
    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
  3256
  thus ?thesis
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3257
    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
  3258
    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
  3259
qed
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3260
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3261
lemma sin_45: "sin (pi / 4) = sqrt 2 / 2"
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  3262
by (simp add: sin_cos_eq cos_45)
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3263
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3264
lemma sin_60: "sin (pi / 3) = sqrt 3 / 2"
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  3265
by (simp add: sin_cos_eq cos_30)
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3266
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3267
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
  3268
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
  3269
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
  3270
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
  3271
done
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3272
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3273
lemma sin_30: "sin (pi / 6) = 1 / 2"
45309
5885ec8eb6b0 removed ad-hoc simp rules sin_cos_eq[symmetric], minus_sin_cos_eq[symmetric], cos_sin_eq[symmetric]
huffman
parents: 45308
diff changeset
  3274
by (simp add: sin_cos_eq cos_60)
23052
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3275
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3276
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
  3277
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
  3278
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3279
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
  3280
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
  3281
0e36f0dbfa1c add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents: 23049
diff changeset
  3282
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
  3283
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
  3284
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  3285
lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n"
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  3286
proof -
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  3287
  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
  3288
    by (auto simp add: algebra_simps sin_add)
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  3289
  thus ?thesis
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  3290
    by (simp add: real_of_nat_Suc distrib_right add_divide_distrib
15383
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  3291
                  mult_commute [of pi])
c49e4225ef4f made proofs more robust
paulson
parents: 15251
diff changeset
  3292
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3293
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3294
lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3295
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
  3296
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3297
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
  3298
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
  3299
apply (subst cos_add, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3300
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3301
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3302
lemma sin_2npi [simp]: "sin (2 * real (n::nat) * pi) = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3303
by (auto simp add: mult_assoc)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3304
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3305
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
  3306
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
  3307
apply (subst sin_add, simp)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3308
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3309
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3310
lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0"
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 47489
diff changeset
  3311
by (simp only: cos_add sin_add real_of_nat_Suc distrib_right distrib_left add_divide_distrib, auto)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3312
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3313
lemma DERIV_cos_add [simp]: "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)"
31881
eba74a5790d2 use DERIV_intros
hoelzl
parents: 31880
diff changeset
  3314
  by (auto intro!: DERIV_intros)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3315
15081
32402f5624d1 abs notation
paulson
parents: 15079
diff changeset
  3316
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
  3317
by (auto simp add: sin_zero_iff even_mult_two_ex)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3318
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3319
lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0"
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3320
by (cut_tac x = x in sin_cos_squared_add3, auto)
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3321
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3322
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
  3323
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3324
lemma arctan_one: "arctan 1 = pi / 4"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3325
  by (rule arctan_unique, simp_all add: tan_45 m2pi_less_pi)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3326
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3327
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
  3328
  shows "\<exists> z. - (pi / 4) < z \<and> z < pi / 4 \<and> tan z = x"
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3329
proof
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3330
  show "- (pi / 4) < arctan x \<and> arctan x < pi / 4 \<and> tan (arctan x) = x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3331
    unfolding arctan_one [symmetric] arctan_minus [symmetric]
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3332
    unfolding arctan_less_iff using assms 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
  3333
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3334
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3335
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
  3336
  shows "arctan x + arctan y = arctan ((x + y) / (1 - x * y))"
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3337
proof (rule arctan_unique [symmetric])
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3338
  have "- (pi / 4) \<le> arctan x" and "- (pi / 4) < arctan y"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3339
    unfolding arctan_one [symmetric] arctan_minus [symmetric]
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3340
    unfolding arctan_le_iff arctan_less_iff using assms by auto
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3341
  from add_le_less_mono [OF this]
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3342
  show 1: "- (pi / 2) < arctan x + arctan y" by simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3343
  have "arctan x \<le> pi / 4" and "arctan y < pi / 4"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3344
    unfolding arctan_one [symmetric]
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3345
    unfolding arctan_le_iff arctan_less_iff using assms by auto
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3346
  from add_le_less_mono [OF this]
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3347
  show 2: "arctan x + arctan y < pi / 2" by simp
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3348
  show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3349
    using cos_gt_zero_pi [OF 1 2] by (simp add: tan_add)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3350
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3351
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3352
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
  3353
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3354
  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
  3355
  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
  3356
  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
  3357
  moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3358
  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
  3359
  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
  3360
  have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3361
  moreover
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3362
  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
  3363
  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
  3364
  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
  3365
  ultimately have "arctan 1 + arctan (1 / 239) = 4 * arctan (1 / 5)" by auto
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3366
  thus ?thesis unfolding arctan_one by algebra
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3367
qed
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3368
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3369
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
  3370
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3371
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
  3372
  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
  3373
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
  3374
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3375
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3376
  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
  3377
  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
  3378
  proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3379
    { 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
  3380
      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
  3381
      proof (rule mult_mono)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3382
        show "1 / real (Suc (Suc n * 2)) \<le> 1 / real (Suc (n * 2))" by (rule frac_le) simp_all
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3383
        show "0 \<le> 1 / real (Suc (n * 2))" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3384
        show "x ^ Suc (Suc n * 2) \<le> x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: `0 \<le> x` `x \<le> 1`)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3385
        show "0 \<le> x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: `0 \<le> x`)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3386
      qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3387
    } note mono = this
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3388
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3389
    show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3390
    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
  3391
      case True from mono[OF this `x \<le> 1`, THEN allI]
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31338
diff changeset
  3392
      show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI2)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3393
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3394
      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
  3395
      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
  3396
      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
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31338
diff changeset
  3397
      thus ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI1[OF allI])
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3398
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3399
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3400
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3401
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3402
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
  3403
  assumes "\<bar>x\<bar> \<le> 1" shows "(\<lambda> n. 1 / real (n*2+1) * x^(n*2+1)) ----> 0" (is "?a ----> 0")
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44319
diff changeset
  3404
proof (cases "x = 0") case True thus ?thesis unfolding One_nat_def by (auto simp add: tendsto_const)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3405
next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3406
  case False
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3407
  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
  3408
  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
  3409
  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
  3410
    case True hence "norm x < 1" by auto
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44319
diff changeset
  3411
    from tendsto_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
  3412
    have "(\<lambda>n. 1 / real (n + 1) * x ^ (n + 1)) ----> 0"
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31338
diff changeset
  3413
      unfolding inverse_eq_divide Suc_eq_plus1 by simp
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3414
    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
  3415
  next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3416
    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
  3417
    hence n_eq: "\<And> n. x ^ (n * 2 + 1) = x" unfolding One_nat_def by auto
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44319
diff changeset
  3418
    from tendsto_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] tendsto_const[of x]]
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31338
diff changeset
  3419
    show ?thesis unfolding n_eq Suc_eq_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
  3420
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3421
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3422
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3423
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
  3424
  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
  3425
  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
  3426
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3427
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
  3428
proof -
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37887
diff changeset
  3429
  from mult_left_mono[OF less_imp_le[OF `\<bar>x\<bar> < 1`] abs_ge_zero[of x]]
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3430
  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
  3431
  thus ?thesis using zero_le_power2 by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3432
qed
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3433
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3434
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
  3435
  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
  3436
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3437
  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
  3438
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3439
  { 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
  3440
  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
  3441
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3442
  { 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
  3443
    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
  3444
      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
  3445
    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
  3446
  } 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
  3447
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3448
  { 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
  3449
    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
  3450
    proof
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3451
      fix x :: real assume "f sums x"
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3452
      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
  3453
      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
  3454
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3455
      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
  3456
      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
  3457
      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
  3458
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3459
    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
  3460
  } 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
  3461
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3462
  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
  3463
    by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3464
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3465
  { fix x :: real
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3466
    have if_eq': "\<And> n. (if even n then -1 ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n =
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3467
      (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
  3468
      using n_even by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3469
    have idx_eq: "\<And> n. n * 2 + 1 = Suc (2 * n)" 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
  3470
    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
  3471
      by auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3472
  } 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
  3473
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3474
  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
  3475
  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
  3476
    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
  3477
    { 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
  3478
      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
  3479
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3480
      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
  3481
      show "summable (\<lambda> n. ?f n * real (Suc n) * x'^n)" unfolding if_eq
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3482
        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`])
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3483
    }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3484
  qed auto
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3485
  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
  3486
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3487
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3488
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
  3489
  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
  3490
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3491
  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
  3492
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3493
  { 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
  3494
    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
  3495
    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
  3496
    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
  3497
  } 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
  3498
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3499
  { 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
  3500
  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
  3501
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3502
  { 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
  3503
  proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3504
    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
  3505
    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
  3506
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3507
    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
  3508
    proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3509
      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
  3510
      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
  3511
      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
  3512
      proof (rule DERIV_isconst2[of "a" "b"])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3513
        show "a < b" and "a \<le> x" and "x \<le> b" using `a < b` `a \<le> x` `x \<le> b` by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3514
        have "\<forall> x. -r < x \<and> x < r \<longrightarrow> DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3515
        proof (rule allI, rule impI)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3516
          fix x assume "-r < x \<and> x < r" hence "\<bar>x\<bar> < r" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3517
          hence "\<bar>x\<bar> < 1" using `r < 1` by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3518
          have "\<bar> - (x^2) \<bar> < 1" using less_one_imp_sqr_less_one[OF `\<bar>x\<bar> < 1`] by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3519
          hence "(\<lambda> n. (- (x^2)) ^ n) sums (1 / (1 - (- (x^2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3520
          hence "(?c' x) sums (1 / (1 - (- (x^2))))" unfolding power_mult_distrib[symmetric] power_mult nat_mult_commute[of _ 2] by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3521
          hence suminf_c'_eq_geom: "inverse (1 + x^2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3522
          have "DERIV (\<lambda> x. suminf (?c x)) x :> (inverse (1 + x^2))" unfolding suminf_c'_eq_geom
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3523
            by (rule DERIV_arctan_suminf[OF `0 < r` `r < 1` `\<bar>x\<bar> < r`])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3524
          from DERIV_add_minus[OF this DERIV_arctan]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3525
          show "DERIV (\<lambda> x. suminf (?c x) - arctan x) x :> 0" unfolding diff_minus by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3526
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3527
        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
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3528
        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
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3529
        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
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3530
      qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3531
    qed
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3532
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3533
    have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0"
31790
05c92381363c corrected and unified thm names
nipkow
parents: 31338
diff changeset
  3534
      unfolding Suc_eq_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3535
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3536
    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
  3537
    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
  3538
      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
  3539
    next
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3540
      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
  3541
      have "suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>) = suminf (?c 0) - arctan 0"
35038
a1d93ce94235 more precise proofs
haftmann
parents: 35028
diff changeset
  3542
        by (rule suminf_eq_arctan_bounded[where x="0" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>", symmetric])
a1d93ce94235 more precise proofs
haftmann
parents: 35028
diff changeset
  3543
          (simp_all only: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>` neg_less_iff_less)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3544
      moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3545
      have "suminf (?c x) - arctan x = suminf (?c (-\<bar>x\<bar>)) - arctan (-\<bar>x\<bar>)"
35038
a1d93ce94235 more precise proofs
haftmann
parents: 35028
diff changeset
  3546
        by (rule suminf_eq_arctan_bounded[where x="x" and a="-\<bar>x\<bar>" and b="\<bar>x\<bar>"])
a1d93ce94235 more precise proofs
haftmann
parents: 35028
diff changeset
  3547
          (simp_all only: `\<bar>x\<bar> < r` `-\<bar>x\<bar> < \<bar>x\<bar>` neg_less_iff_less)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3548
      ultimately
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3549
      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
  3550
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3551
    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
  3552
  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
  3553
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3554
  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
  3555
  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
  3556
    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
  3557
  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
  3558
    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
  3559
    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
  3560
    { 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
  3561
      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
  3562
      moreover
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3563
      { fix x :: real assume "0 < x" and "x < 1" hence "\<bar>x\<bar> \<le> 1" and "\<bar>x\<bar> < 1" by auto
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3564
        from `0 < x` have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3565
        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]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3566
        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)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3567
        hence a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3568
        have "?diff x n \<le> ?a x n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3569
        proof (cases "even n")
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3570
          case True hence sgn_pos: "(-1)^n = (1::real)" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3571
          from `even n` obtain m where "2 * m = n" unfolding even_mult_two_ex by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3572
          from bounds[of m, unfolded this atLeastAtMost_iff]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3573
          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
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3574
          also have "\<dots> = ?c x n" unfolding One_nat_def by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3575
          also have "\<dots> = ?a x n" unfolding sgn_pos a_pos by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3576
          finally show ?thesis .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3577
        next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3578
          case False hence sgn_neg: "(-1)^n = (-1::real)" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3579
          from `odd n` obtain m where m_def: "2 * m + 1 = n" unfolding odd_Suc_mult_two_ex by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3580
          hence m_plus: "2 * (m + 1) = n + 1" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3581
          from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3582
          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
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3583
          also have "\<dots> = - ?c x n" unfolding One_nat_def by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3584
          also have "\<dots> = ?a x n" unfolding sgn_neg a_pos by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3585
          finally show ?thesis .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3586
        qed
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3587
        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
  3588
      }
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3589
      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
  3590
      moreover have "\<And>x. isCont (\<lambda> x. ?a x n - ?diff x n) x"
37887
2ae085b07f2f diff_minus subsumes diff_def
haftmann
parents: 36974
diff changeset
  3591
        unfolding diff_minus divide_inverse
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3592
        by (auto intro!: isCont_add isCont_rabs isCont_ident isCont_minus isCont_arctan isCont_inverse isCont_mult isCont_power isCont_const isCont_setsum)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3593
      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
  3594
      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
  3595
    }
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3596
    have "?a 1 ----> 0"
44568
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44319
diff changeset
  3597
      unfolding tendsto_rabs_zero_iff power_one divide_inverse One_nat_def
e6f291cb5810 discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents: 44319
diff changeset
  3598
      by (auto intro!: tendsto_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
  3599
    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
  3600
    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
  3601
      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
  3602
      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
  3603
      { fix n assume "N \<le> n" from `?diff 1 n \<le> ?a 1 n` N_I[OF this]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32047
diff changeset
  3604
        have "norm (?diff 1 n - 0) < r" 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
  3605
      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
  3606
    qed
44710
9caf6883f1f4 remove redundant lemmas about LIMSEQ
huffman
parents: 44568
diff changeset
  3607
    from this [unfolded tendsto_rabs_zero_iff, THEN tendsto_add [OF _ tendsto_const], of "- arctan 1", THEN tendsto_minus]
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3608
    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
  3609
    hence "arctan 1 = (\<Sum> i. ?c 1 i)" by (rule sums_unique)
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3610
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3611
    show ?thesis
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3612
    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
  3613
      assume "x \<noteq> 1" hence "x = -1" using `\<bar>x\<bar> = 1` by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3614
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3615
      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
  3616
      have "- (2 * pi) < 0" using pi_gt_zero by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3617
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3618
      have c_minus_minus: "\<And> i. ?c (- 1) i = - ?c 1 i" unfolding One_nat_def by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3619
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3620
      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
  3621
      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
  3622
      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
  3623
      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
  3624
      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
  3625
      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
  3626
      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
  3627
    qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3628
  qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3629
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3630
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3631
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
  3632
  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
  3633
proof -
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3634
  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
  3635
  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
  3636
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3637
  have divide_nonzero_divide: "\<And> A B C :: real. C \<noteq> 0 \<Longrightarrow> A / B = (A / C) / (B / C)" by auto
41970
47d6e13d1710 generalize infinite sums
hoelzl
parents: 41550
diff changeset
  3638
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3639
  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
  3640
  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
  3641
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3642
  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
  3643
  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
  3644
  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
  3645
  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
  3646
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3647
  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
  3648
  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
  3649
  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
  3650
  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
  3651
  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
  3652
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3653
  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
  3654
  also have "\<dots> = 2 * (arctan (tan (y/2)))" using arctan_tan[OF low2 high2] by auto
44756
efcd71fbaeec simplify proof of tan_half, removing unused assumptions
huffman
parents: 44755
diff changeset
  3655
  also have "\<dots> = 2 * (arctan (sin y / (cos y + 1)))" unfolding tan_half 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
  3656
  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
  3657
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3658
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3659
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
  3660
  shows "arctan x < arctan y"
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3661
  using assms by (simp only: arctan_less_iff)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3662
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3663
lemma arctan_monotone': assumes "x \<le> y" shows "arctan x \<le> arctan y"
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3664
  using assms by (simp only: arctan_le_iff)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3665
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3666
lemma arctan_inverse:
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3667
  assumes "x \<noteq> 0" shows "arctan (1 / x) = sgn x * pi / 2 - arctan x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3668
proof (rule arctan_unique)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3669
  show "- (pi / 2) < sgn x * pi / 2 - arctan x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3670
    using arctan_bounded [of x] assms
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3671
    unfolding sgn_real_def
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3672
    apply (auto simp add: algebra_simps)
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3673
    apply (drule zero_less_arctan_iff [THEN iffD2])
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3674
    apply arith
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3675
    done
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3676
  show "sgn x * pi / 2 - arctan x < pi / 2"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3677
    using arctan_bounded [of "- x"] assms
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3678
    unfolding sgn_real_def arctan_minus
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3679
    by auto
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3680
  show "tan (sgn x * pi / 2 - arctan x) = 1 / x"
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3681
    unfolding tan_inverse [of "arctan x", unfolded tan_arctan]
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3682
    unfolding sgn_real_def
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3683
    by (simp add: tan_def cos_arctan sin_arctan sin_diff cos_diff)
29803
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3684
qed
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3685
c56a5571f60a Added derivation lemmas for power series and theorems for the pi, arcus tangens and logarithm series
hoelzl
parents: 29695
diff changeset
  3686
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
  3687
proof -
44746
9e4f7d3b5376 add lemmas about arctan;
huffman
parents: 44745
diff changeset
  3688
  have "pi / 4 = arctan 1" using arctan_one 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
  3689
  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
  3690
  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
  3691
qed
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3692
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3693
subsection {* Existence of Polar Coordinates *}
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3694
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3695
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
  3696
apply (rule power2_le_imp_le [OF _ zero_le_one])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35213
diff changeset
  3697
apply (simp add: power_divide divide_le_eq not_sum_power2_lt_zero)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3698
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3699
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3700
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
  3701
by (simp add: abs_le_iff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3702
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3703
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
  3704
by (simp add: sin_arccos abs_le_iff)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3705
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3706
lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one]
15228
4d332d10fa3d revised simprules for division
paulson
parents: 15140
diff changeset
  3707
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3708
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
  3709
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  3710
lemma polar_ex1:
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3711
     "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
  3712
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
  3713
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
  3714
apply (simp add: cos_arccos_lemma1)
23045
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3715
apply (simp add: sin_arccos_lemma1)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3716
apply (simp add: power_divide)
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3717
apply (simp add: real_sqrt_mult [symmetric])
95e04f335940 add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents: 23043
diff changeset
  3718
apply (simp add: right_diff_distrib)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3719
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3720
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
  3721
lemma polar_ex2:
22978
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3722
     "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
  3723
apply (insert polar_ex1 [where x=x and y="-y"], simp, clarify)
33667
958dc9f03611 A little rationalisation
paulson
parents: 33549
diff changeset
  3724
apply (metis cos_minus minus_minus minus_mult_right sin_minus)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3725
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3726
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3727
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
  3728
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
  3729
apply (erule polar_ex1)
1cd8cc21a7c3 clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents: 22977
diff changeset
  3730
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
  3731
apply (erule polar_ex2)
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3732
done
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15013
diff changeset
  3733
30082
43c5b7bfc791 make more proofs work whether or not One_nat_def is a simp rule
huffman
parents: 29803
diff changeset
  3734
end