author | huffman |
Thu, 31 May 2007 23:02:16 +0200 | |
changeset 23177 | 3004310c95b1 |
parent 23176 | 40a760921d94 |
child 23241 | 5f12b40a95bf |
permissions | -rw-r--r-- |
12196 | 1 |
(* Title : Transcendental.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 1998,1999 University of Cambridge |
|
13958
c1c67582c9b5
New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
12196
diff
changeset
|
4 |
1999,2001 University of Edinburgh |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
5 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2004 |
12196 | 6 |
*) |
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 | 10 |
theory Transcendental |
22654
c2b6b5a9e136
new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents:
22653
diff
changeset
|
11 |
imports NthRoot Fact Series EvenOdd Deriv |
15131 | 12 |
begin |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
13 |
|
23043 | 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: |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
17 |
fixes y :: "'a::recpower" |
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) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
22 |
thus ?thesis by (simp add: power_Suc power_commutes) |
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: |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
26 |
fixes y :: "'a::{recpower,comm_semiring_0}" shows |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
27 |
"(\<Sum>p=0..<Suc n. (x ^ p) * y ^ (Suc n - p)) = |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
28 |
y * (\<Sum>p=0..<Suc n. (x ^ p) * y ^ (n - p))" |
19279 | 29 |
by (auto simp add: setsum_right_distrib lemma_realpow_diff mult_ac |
16641
fce796ad9c2b
Simplified some proofs (thanks to strong_setsum_cong).
berghofe
parents:
15561
diff
changeset
|
30 |
simp del: setsum_op_ivl_Suc cong: strong_setsum_cong) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
31 |
|
15229 | 32 |
lemma lemma_realpow_diff_sumr2: |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
33 |
fixes y :: "'a::{recpower,comm_ring}" shows |
15229 | 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))" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
36 |
apply (induct "n", simp add: power_Suc) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
37 |
apply (simp add: power_Suc del: setsum_op_ivl_Suc) |
15561 | 38 |
apply (subst setsum_op_ivl_Suc) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
39 |
apply (subst lemma_realpow_diff_sumr) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
40 |
apply (simp add: right_distrib del: setsum_op_ivl_Suc) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
41 |
apply (subst mult_left_commute [where a="x - y"]) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
42 |
apply (erule subst) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
43 |
apply (simp add: power_Suc ring_eq_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 | 46 |
lemma lemma_realpow_rev_sumr: |
15539 | 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: |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
59 |
fixes x z :: "'a::{real_normed_field,banach,recpower}" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
60 |
assumes 1: "summable (\<lambda>n. f n * x ^ n)" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
61 |
assumes 2: "norm z < norm x" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
62 |
shows "summable (\<lambda>n. norm (f n * z ^ n))" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
63 |
proof - |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
64 |
from 2 have x_neq_0: "x \<noteq> 0" by clarsimp |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
65 |
from 1 have "(\<lambda>n. f n * x ^ n) ----> 0" |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
66 |
by (rule summable_LIMSEQ_zero) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
67 |
hence "convergent (\<lambda>n. f n * x ^ n)" |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
68 |
by (rule convergentI) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
69 |
hence "Cauchy (\<lambda>n. f n * x ^ n)" |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
70 |
by (simp add: Cauchy_convergent_iff) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
71 |
hence "Bseq (\<lambda>n. f n * x ^ n)" |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
72 |
by (rule Cauchy_Bseq) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
73 |
then obtain K where 3: "0 < K" and 4: "\<forall>n. norm (f n * x ^ n) \<le> K" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
74 |
by (simp add: Bseq_def, safe) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
75 |
have "\<exists>N. \<forall>n\<ge>N. norm (norm (f n * z ^ n)) \<le> |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
76 |
K * norm (z ^ n) * inverse (norm (x ^ n))" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
77 |
proof (intro exI allI impI) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
78 |
fix n::nat assume "0 \<le> n" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
79 |
have "norm (norm (f n * z ^ n)) * norm (x ^ n) = |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
80 |
norm (f n * x ^ n) * norm (z ^ n)" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
81 |
by (simp add: norm_mult abs_mult) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
82 |
also have "\<dots> \<le> K * norm (z ^ n)" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
83 |
by (simp only: mult_right_mono 4 norm_ge_zero) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
84 |
also have "\<dots> = K * norm (z ^ n) * (inverse (norm (x ^ n)) * norm (x ^ n))" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
85 |
by (simp add: x_neq_0) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
86 |
also have "\<dots> = K * norm (z ^ n) * inverse (norm (x ^ n)) * norm (x ^ n)" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
87 |
by (simp only: mult_assoc) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
88 |
finally show "norm (norm (f n * z ^ n)) \<le> |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
89 |
K * norm (z ^ n) * inverse (norm (x ^ n))" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
90 |
by (simp add: mult_le_cancel_right x_neq_0) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
91 |
qed |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
92 |
moreover have "summable (\<lambda>n. K * norm (z ^ n) * inverse (norm (x ^ n)))" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
93 |
proof - |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
94 |
from 2 have "norm (norm (z * inverse x)) < 1" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
95 |
using x_neq_0 |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
96 |
by (simp add: nonzero_norm_divide divide_inverse [symmetric]) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
97 |
hence "summable (\<lambda>n. norm (z * inverse x) ^ n)" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
98 |
by (rule summable_geometric) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
99 |
hence "summable (\<lambda>n. K * norm (z * inverse x) ^ n)" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
100 |
by (rule summable_mult) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
101 |
thus "summable (\<lambda>n. K * norm (z ^ n) * inverse (norm (x ^ n)))" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
102 |
using x_neq_0 |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
103 |
by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
104 |
power_inverse norm_power mult_assoc) |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
105 |
qed |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
106 |
ultimately show "summable (\<lambda>n. norm (f n * z ^ n))" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
107 |
by (rule summable_comparison_test) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
108 |
qed |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
109 |
|
15229 | 110 |
lemma powser_inside: |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
111 |
fixes f :: "nat \<Rightarrow> 'a::{real_normed_field,banach,recpower}" shows |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
112 |
"[| summable (%n. f(n) * (x ^ n)); norm z < norm x |] |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
113 |
==> summable (%n. f(n) * (z ^ n))" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
114 |
by (rule powser_insidea [THEN summable_norm_cancel]) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
115 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
116 |
|
23043 | 117 |
subsection{*Term-by-Term Differentiability of Power Series*} |
118 |
||
119 |
definition |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
120 |
diffs :: "(nat => 'a::ring_1) => nat => 'a" where |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
121 |
"diffs c = (%n. of_nat (Suc n) * c(Suc n))" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
122 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
123 |
text{*Lemma about distributing negation over it*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
124 |
lemma diffs_minus: "diffs (%n. - c n) = (%n. - diffs c n)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
125 |
by (simp add: diffs_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
126 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
127 |
text{*Show that we can shift the terms down one*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
128 |
lemma lemma_diffs: |
15539 | 129 |
"(\<Sum>n=0..<n. (diffs c)(n) * (x ^ n)) = |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
130 |
(\<Sum>n=0..<n. of_nat n * c(n) * (x ^ (n - Suc 0))) + |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
131 |
(of_nat n * c(n) * x ^ (n - Suc 0))" |
15251 | 132 |
apply (induct "n") |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
133 |
apply (auto simp add: mult_assoc add_assoc [symmetric] diffs_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
134 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
135 |
|
15229 | 136 |
lemma lemma_diffs2: |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
137 |
"(\<Sum>n=0..<n. of_nat n * c(n) * (x ^ (n - Suc 0))) = |
15539 | 138 |
(\<Sum>n=0..<n. (diffs c)(n) * (x ^ n)) - |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
139 |
(of_nat n * c(n) * x ^ (n - Suc 0))" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
140 |
by (auto simp add: lemma_diffs) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
141 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
142 |
|
15229 | 143 |
lemma diffs_equiv: |
144 |
"summable (%n. (diffs c)(n) * (x ^ n)) ==> |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
145 |
(%n. of_nat n * c(n) * (x ^ (n - Suc 0))) sums |
15546 | 146 |
(\<Sum>n. (diffs c)(n) * (x ^ n))" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
147 |
apply (subgoal_tac " (%n. of_nat n * c (n) * (x ^ (n - Suc 0))) ----> 0") |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
148 |
apply (rule_tac [2] LIMSEQ_imp_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
149 |
apply (drule summable_sums) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
150 |
apply (auto simp add: sums_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
151 |
apply (drule_tac X="(\<lambda>n. \<Sum>n = 0..<n. diffs c n * x ^ n)" in LIMSEQ_diff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
152 |
apply (auto simp add: lemma_diffs2 [symmetric] diffs_def [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
153 |
apply (simp add: diffs_def summable_LIMSEQ_zero) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
154 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
155 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
156 |
lemma lemma_termdiff1: |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
157 |
fixes z :: "'a :: {recpower,comm_ring}" shows |
15539 | 158 |
"(\<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
|
159 |
(\<Sum>p=0..<m. (z ^ p) * (((z + h) ^ (m - p)) - (z ^ (m - p))))" |
16641
fce796ad9c2b
Simplified some proofs (thanks to strong_setsum_cong).
berghofe
parents:
15561
diff
changeset
|
160 |
by (auto simp add: right_distrib diff_minus power_add [symmetric] mult_ac |
fce796ad9c2b
Simplified some proofs (thanks to strong_setsum_cong).
berghofe
parents:
15561
diff
changeset
|
161 |
cong: strong_setsum_cong) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
162 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
163 |
lemma less_add_one: "m < n ==> (\<exists>d. n = m + d + Suc 0)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
164 |
by (simp add: less_iff_Suc_add) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
165 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
166 |
lemma sumdiff: "a + b - (c + d) = a - c + b - (d::real)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
167 |
by arith |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
168 |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
169 |
lemma sumr_diff_mult_const2: |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
170 |
"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
|
171 |
by (simp add: setsum_subtractf) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
172 |
|
15229 | 173 |
lemma lemma_termdiff2: |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
174 |
fixes h :: "'a :: {recpower,field}" |
20860 | 175 |
assumes h: "h \<noteq> 0" shows |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
176 |
"((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) = |
20860 | 177 |
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
|
178 |
(z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
179 |
apply (subgoal_tac "h * ?lhs = h * ?rhs", simp add: h) |
20860 | 180 |
apply (simp add: right_diff_distrib diff_divide_distrib h) |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
181 |
apply (simp only: times_divide_eq_left [symmetric]) |
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
182 |
apply (simp add: divide_self [OF h]) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
183 |
apply (simp add: mult_assoc [symmetric]) |
20860 | 184 |
apply (cases "n", simp) |
185 |
apply (simp add: lemma_realpow_diff_sumr2 h |
|
186 |
right_diff_distrib [symmetric] mult_assoc |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
187 |
del: realpow_Suc setsum_op_ivl_Suc of_nat_Suc) |
20860 | 188 |
apply (subst lemma_realpow_rev_sumr) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
189 |
apply (subst sumr_diff_mult_const2) |
20860 | 190 |
apply simp |
191 |
apply (simp only: lemma_termdiff1 setsum_right_distrib) |
|
192 |
apply (rule setsum_cong [OF refl]) |
|
15539 | 193 |
apply (simp add: diff_minus [symmetric] less_iff_Suc_add) |
20860 | 194 |
apply (clarify) |
195 |
apply (simp add: setsum_right_distrib lemma_realpow_diff_sumr2 mult_ac |
|
196 |
del: setsum_op_ivl_Suc realpow_Suc) |
|
197 |
apply (subst mult_assoc [symmetric], subst power_add [symmetric]) |
|
198 |
apply (simp add: mult_ac) |
|
199 |
done |
|
200 |
||
201 |
lemma real_setsum_nat_ivl_bounded2: |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
202 |
fixes K :: "'a::ordered_semidom" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
203 |
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
|
204 |
assumes K: "0 \<le> K" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
205 |
shows "setsum f {0..<n-k} \<le> of_nat n * K" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
206 |
apply (rule order_trans [OF setsum_mono]) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
207 |
apply (rule f, simp) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
208 |
apply (simp add: mult_right_mono K) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
209 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
210 |
|
15229 | 211 |
lemma lemma_termdiff3: |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
212 |
fixes h z :: "'a::{real_normed_field,recpower}" |
20860 | 213 |
assumes 1: "h \<noteq> 0" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
214 |
assumes 2: "norm z \<le> K" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
215 |
assumes 3: "norm (z + h) \<le> K" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
216 |
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
|
217 |
\<le> of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" |
20860 | 218 |
proof - |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
219 |
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
|
220 |
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
|
221 |
(z + h) ^ q * z ^ (n - 2 - q)) * norm h" |
20860 | 222 |
apply (subst lemma_termdiff2 [OF 1]) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
223 |
apply (subst norm_mult) |
20860 | 224 |
apply (rule mult_commute) |
225 |
done |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
226 |
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
|
227 |
proof (rule mult_right_mono [OF _ norm_ge_zero]) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
228 |
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
|
229 |
have le_Kn: "\<And>i j n. i + j = n \<Longrightarrow> norm ((z + h) ^ i * z ^ j) \<le> K ^ n" |
20860 | 230 |
apply (erule subst) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
231 |
apply (simp only: norm_mult norm_power power_add) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
232 |
apply (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) |
20860 | 233 |
done |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
234 |
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
|
235 |
(z + h) ^ q * z ^ (n - 2 - q)) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
236 |
\<le> of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))" |
20860 | 237 |
apply (intro |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
238 |
order_trans [OF norm_setsum] |
20860 | 239 |
real_setsum_nat_ivl_bounded2 |
240 |
mult_nonneg_nonneg |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
241 |
zero_le_imp_of_nat |
20860 | 242 |
zero_le_power K) |
243 |
apply (rule le_Kn, simp) |
|
244 |
done |
|
245 |
qed |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
246 |
also have "\<dots> = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" |
20860 | 247 |
by (simp only: mult_assoc) |
248 |
finally show ?thesis . |
|
249 |
qed |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
250 |
|
20860 | 251 |
lemma lemma_termdiff4: |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
252 |
fixes f :: "'a::{real_normed_field,recpower} \<Rightarrow> |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
253 |
'b::real_normed_vector" |
20860 | 254 |
assumes k: "0 < (k::real)" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
255 |
assumes le: "\<And>h. \<lbrakk>h \<noteq> 0; norm h < k\<rbrakk> \<Longrightarrow> norm (f h) \<le> K * norm h" |
20860 | 256 |
shows "f -- 0 --> 0" |
257 |
proof (simp add: LIM_def, safe) |
|
258 |
fix r::real assume r: "0 < r" |
|
259 |
have zero_le_K: "0 \<le> K" |
|
260 |
apply (cut_tac k) |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
261 |
apply (cut_tac h="of_real (k/2)" in le, simp) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
262 |
apply (simp del: of_real_divide) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
263 |
apply (drule order_trans [OF norm_ge_zero]) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
264 |
apply (simp add: zero_le_mult_iff) |
20860 | 265 |
done |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
266 |
show "\<exists>s. 0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < s \<longrightarrow> norm (f x) < r)" |
20860 | 267 |
proof (cases) |
268 |
assume "K = 0" |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
269 |
with k r le have "0 < k \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < k \<longrightarrow> norm (f x) < r)" |
20860 | 270 |
by simp |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
271 |
thus "\<exists>s. 0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < s \<longrightarrow> norm (f x) < r)" .. |
20860 | 272 |
next |
273 |
assume K_neq_zero: "K \<noteq> 0" |
|
274 |
with zero_le_K have K: "0 < K" by simp |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
275 |
show "\<exists>s. 0 < s \<and> (\<forall>x. x \<noteq> 0 \<and> norm x < s \<longrightarrow> norm (f x) < r)" |
20860 | 276 |
proof (rule exI, safe) |
277 |
from k r K show "0 < min k (r * inverse K / 2)" |
|
278 |
by (simp add: mult_pos_pos positive_imp_inverse_positive) |
|
279 |
next |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
280 |
fix x::'a |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
281 |
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
|
282 |
from x2 have x3: "norm x < k" and x4: "norm x < r * inverse K / 2" |
20860 | 283 |
by simp_all |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
284 |
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
|
285 |
also from x4 K have "K * norm x < K * (r * inverse K / 2)" |
20860 | 286 |
by (rule mult_strict_left_mono) |
287 |
also have "\<dots> = r / 2" |
|
288 |
using K_neq_zero by simp |
|
289 |
also have "r / 2 < r" |
|
290 |
using r by simp |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
291 |
finally show "norm (f x) < r" . |
20860 | 292 |
qed |
293 |
qed |
|
294 |
qed |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
295 |
|
15229 | 296 |
lemma lemma_termdiff5: |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
297 |
fixes g :: "'a::{recpower,real_normed_field} \<Rightarrow> |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
298 |
nat \<Rightarrow> 'b::banach" |
20860 | 299 |
assumes k: "0 < (k::real)" |
300 |
assumes f: "summable f" |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
301 |
assumes le: "\<And>h n. \<lbrakk>h \<noteq> 0; norm h < k\<rbrakk> \<Longrightarrow> norm (g h n) \<le> f n * norm h" |
20860 | 302 |
shows "(\<lambda>h. suminf (g h)) -- 0 --> 0" |
303 |
proof (rule lemma_termdiff4 [OF k]) |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
304 |
fix h::'a assume "h \<noteq> 0" and "norm h < k" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
305 |
hence A: "\<forall>n. norm (g h n) \<le> f n * norm h" |
20860 | 306 |
by (simp add: le) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
307 |
hence "\<exists>N. \<forall>n\<ge>N. norm (norm (g h n)) \<le> f n * norm h" |
20860 | 308 |
by simp |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
309 |
moreover from f have B: "summable (\<lambda>n. f n * norm h)" |
20860 | 310 |
by (rule summable_mult2) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
311 |
ultimately have C: "summable (\<lambda>n. norm (g h n))" |
20860 | 312 |
by (rule summable_comparison_test) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
313 |
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
|
314 |
by (rule summable_norm) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
315 |
also from A C B have "(\<Sum>n. norm (g h n)) \<le> (\<Sum>n. f n * norm h)" |
20860 | 316 |
by (rule summable_le) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
317 |
also from f have "(\<Sum>n. f n * norm h) = suminf f * norm h" |
20860 | 318 |
by (rule suminf_mult2 [symmetric]) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
319 |
finally show "norm (suminf (g h)) \<le> suminf f * norm h" . |
20860 | 320 |
qed |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
321 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
322 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
323 |
text{* FIXME: Long proofs*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
324 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
325 |
lemma termdiffs_aux: |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
326 |
fixes x :: "'a::{recpower,real_normed_field,banach}" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
327 |
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
|
328 |
assumes 2: "norm x < norm K" |
20860 | 329 |
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
|
330 |
- of_nat n * x ^ (n - Suc 0))) -- 0 --> 0" |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
331 |
proof - |
20860 | 332 |
from dense [OF 2] |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
333 |
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
|
334 |
from norm_ge_zero r1 have r: "0 < r" |
20860 | 335 |
by (rule order_le_less_trans) |
336 |
hence r_neq_0: "r \<noteq> 0" by simp |
|
337 |
show ?thesis |
|
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
338 |
proof (rule lemma_termdiff5) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
339 |
show "0 < r - norm x" using r1 by simp |
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
340 |
next |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
341 |
from r r2 have "norm (of_real r::'a) < norm K" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
342 |
by simp |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
343 |
with 1 have "summable (\<lambda>n. norm (diffs (diffs c) n * (of_real r ^ n)))" |
20860 | 344 |
by (rule powser_insidea) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
345 |
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
|
346 |
using r |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
347 |
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
|
348 |
hence "summable (\<lambda>n. of_nat n * diffs (\<lambda>n. norm (c n)) n * r ^ (n - Suc 0))" |
20860 | 349 |
by (rule diffs_equiv [THEN sums_summable]) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
350 |
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
|
351 |
= (\<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
|
352 |
apply (rule ext) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
353 |
apply (simp add: diffs_def) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
354 |
apply (case_tac n, simp_all add: r_neq_0) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
355 |
done |
20860 | 356 |
finally have "summable |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
357 |
(\<lambda>n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))" |
20860 | 358 |
by (rule diffs_equiv [THEN sums_summable]) |
359 |
also have |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
360 |
"(\<lambda>n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * |
20860 | 361 |
r ^ (n - Suc 0)) = |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
362 |
(\<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
|
363 |
apply (rule ext) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
364 |
apply (case_tac "n", simp) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
365 |
apply (case_tac "nat", simp) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
366 |
apply (simp add: r_neq_0) |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
367 |
done |
20860 | 368 |
finally show |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
369 |
"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
|
370 |
next |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
371 |
fix h::'a and n::nat |
20860 | 372 |
assume h: "h \<noteq> 0" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
373 |
assume "norm h < r - norm x" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
374 |
hence "norm x + norm h < r" by simp |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
375 |
with norm_triangle_ineq have xh: "norm (x + h) < r" |
20860 | 376 |
by (rule order_le_less_trans) |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
377 |
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
|
378 |
\<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
|
379 |
apply (simp only: norm_mult mult_assoc) |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
380 |
apply (rule mult_left_mono [OF _ norm_ge_zero]) |
20860 | 381 |
apply (simp (no_asm) add: mult_assoc [symmetric]) |
382 |
apply (rule lemma_termdiff3) |
|
383 |
apply (rule h) |
|
384 |
apply (rule r1 [THEN order_less_imp_le]) |
|
385 |
apply (rule xh [THEN order_less_imp_le]) |
|
386 |
done |
|
20849
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
387 |
qed |
389cd9c8cfe1
rewrite proofs of powser_insidea and termdiffs_aux
huffman
parents:
20692
diff
changeset
|
388 |
qed |
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
19765
diff
changeset
|
389 |
|
20860 | 390 |
lemma termdiffs: |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
391 |
fixes K x :: "'a::{recpower,real_normed_field,banach}" |
20860 | 392 |
assumes 1: "summable (\<lambda>n. c n * K ^ n)" |
393 |
assumes 2: "summable (\<lambda>n. (diffs c) n * K ^ n)" |
|
394 |
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
|
395 |
assumes 4: "norm x < norm K" |
20860 | 396 |
shows "DERIV (\<lambda>x. \<Sum>n. c n * x ^ n) x :> (\<Sum>n. (diffs c) n * x ^ n)" |
397 |
proof (simp add: deriv_def, rule LIM_zero_cancel) |
|
398 |
show "(\<lambda>h. (suminf (\<lambda>n. c n * (x + h) ^ n) - suminf (\<lambda>n. c n * x ^ n)) / h |
|
399 |
- suminf (\<lambda>n. diffs c n * x ^ n)) -- 0 --> 0" |
|
400 |
proof (rule LIM_equal2) |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
401 |
show "0 < norm K - norm x" by (simp add: less_diff_eq 4) |
20860 | 402 |
next |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
403 |
fix h :: 'a |
20860 | 404 |
assume "h \<noteq> 0" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
405 |
assume "norm (h - 0) < norm K - norm x" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
406 |
hence "norm x + norm h < norm K" by simp |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
407 |
hence 5: "norm (x + h) < norm K" |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
408 |
by (rule norm_triangle_ineq [THEN order_le_less_trans]) |
20860 | 409 |
have A: "summable (\<lambda>n. c n * x ^ n)" |
410 |
by (rule powser_inside [OF 1 4]) |
|
411 |
have B: "summable (\<lambda>n. c n * (x + h) ^ n)" |
|
412 |
by (rule powser_inside [OF 1 5]) |
|
413 |
have C: "summable (\<lambda>n. diffs c n * x ^ n)" |
|
414 |
by (rule powser_inside [OF 2 4]) |
|
415 |
show "((\<Sum>n. c n * (x + h) ^ n) - (\<Sum>n. c n * x ^ n)) / h |
|
416 |
- (\<Sum>n. diffs c n * x ^ n) = |
|
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
417 |
(\<Sum>n. c n * (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)))" |
20860 | 418 |
apply (subst sums_unique [OF diffs_equiv [OF C]]) |
419 |
apply (subst suminf_diff [OF B A]) |
|
420 |
apply (subst suminf_divide [symmetric]) |
|
421 |
apply (rule summable_diff [OF B A]) |
|
422 |
apply (subst suminf_diff) |
|
423 |
apply (rule summable_divide) |
|
424 |
apply (rule summable_diff [OF B A]) |
|
425 |
apply (rule sums_summable [OF diffs_equiv [OF C]]) |
|
426 |
apply (rule_tac f="suminf" in arg_cong) |
|
427 |
apply (rule ext) |
|
428 |
apply (simp add: ring_eq_simps) |
|
429 |
done |
|
430 |
next |
|
431 |
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
|
432 |
of_nat n * x ^ (n - Suc 0))) -- 0 --> 0" |
20860 | 433 |
by (rule termdiffs_aux [OF 3 4]) |
434 |
qed |
|
435 |
qed |
|
436 |
||
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
437 |
|
23043 | 438 |
subsection{*Exponential Function*} |
439 |
||
440 |
definition |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
441 |
exp :: "'a \<Rightarrow> 'a::{recpower,real_normed_field,banach}" where |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
442 |
"exp x = (\<Sum>n. x ^ n /# real (fact n))" |
23043 | 443 |
|
444 |
definition |
|
445 |
sin :: "real => real" where |
|
446 |
"sin x = (\<Sum>n. (if even(n) then 0 else |
|
23177 | 447 |
(-1 ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)" |
23043 | 448 |
|
449 |
definition |
|
450 |
cos :: "real => real" where |
|
23177 | 451 |
"cos x = (\<Sum>n. (if even(n) then (-1 ^ (n div 2))/(real (fact n)) |
23043 | 452 |
else 0) * x ^ n)" |
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
453 |
|
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
454 |
lemma summable_exp_generic: |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
455 |
fixes x :: "'a::{real_normed_algebra_1,recpower,banach}" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
456 |
defines S_def: "S \<equiv> \<lambda>n. x ^ n /# real (fact n)" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
457 |
shows "summable S" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
458 |
proof - |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
459 |
have S_Suc: "\<And>n. S (Suc n) = (x * S n) /# real (Suc n)" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
460 |
unfolding S_def by (simp add: power_Suc del: mult_Suc) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
461 |
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
|
462 |
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
|
463 |
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
|
464 |
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
|
465 |
from r1 show ?thesis |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
466 |
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
|
467 |
fix n :: nat |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
468 |
assume n: "N \<le> n" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
469 |
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
|
470 |
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
|
471 |
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
|
472 |
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
|
473 |
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
|
474 |
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
|
475 |
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
|
476 |
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
|
477 |
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
|
478 |
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
|
479 |
thus "norm (S (Suc n)) \<le> r * norm (S n)" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
480 |
by (simp add: S_Suc norm_scaleR inverse_eq_divide) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
481 |
qed |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
482 |
qed |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
483 |
|
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
484 |
lemma summable_norm_exp: |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
485 |
fixes x :: "'a::{real_normed_algebra_1,recpower,banach}" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
486 |
shows "summable (\<lambda>n. norm (x ^ n /# real (fact n)))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
487 |
proof (rule summable_norm_comparison_test [OF exI, rule_format]) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
488 |
show "summable (\<lambda>n. norm x ^ n /# real (fact n))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
489 |
by (rule summable_exp_generic) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
490 |
next |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
491 |
fix n show "norm (x ^ n /# real (fact n)) \<le> norm x ^ n /# real (fact n)" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
492 |
by (simp add: norm_scaleR norm_power_ineq) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
493 |
qed |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
494 |
|
23043 | 495 |
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
|
496 |
by (insert summable_exp_generic [where x=x], simp) |
23043 | 497 |
|
498 |
lemma summable_sin: |
|
499 |
"summable (%n. |
|
500 |
(if even n then 0 |
|
23177 | 501 |
else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * |
23043 | 502 |
x ^ n)" |
503 |
apply (rule_tac g = "(%n. inverse (real (fact n)) * \<bar>x\<bar> ^ n)" in summable_comparison_test) |
|
504 |
apply (rule_tac [2] summable_exp) |
|
505 |
apply (rule_tac x = 0 in exI) |
|
506 |
apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) |
|
507 |
done |
|
508 |
||
509 |
lemma summable_cos: |
|
510 |
"summable (%n. |
|
511 |
(if even n then |
|
23177 | 512 |
-1 ^ (n div 2)/(real (fact n)) else 0) * x ^ n)" |
23043 | 513 |
apply (rule_tac g = "(%n. inverse (real (fact n)) * \<bar>x\<bar> ^ n)" in summable_comparison_test) |
514 |
apply (rule_tac [2] summable_exp) |
|
515 |
apply (rule_tac x = 0 in exI) |
|
516 |
apply (auto simp add: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) |
|
517 |
done |
|
518 |
||
519 |
lemma lemma_STAR_sin [simp]: |
|
520 |
"(if even n then 0 |
|
23177 | 521 |
else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0" |
23043 | 522 |
by (induct "n", auto) |
523 |
||
524 |
lemma lemma_STAR_cos [simp]: |
|
525 |
"0 < n --> |
|
23177 | 526 |
-1 ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" |
23043 | 527 |
by (induct "n", auto) |
528 |
||
529 |
lemma lemma_STAR_cos1 [simp]: |
|
530 |
"0 < n --> |
|
531 |
(-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0" |
|
532 |
by (induct "n", auto) |
|
533 |
||
534 |
lemma lemma_STAR_cos2 [simp]: |
|
23177 | 535 |
"(\<Sum>n=1..<n. if even n then -1 ^ (n div 2)/(real (fact n)) * 0 ^ n |
23043 | 536 |
else 0) = 0" |
537 |
apply (induct "n") |
|
538 |
apply (case_tac [2] "n", auto) |
|
539 |
done |
|
540 |
||
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
541 |
lemma exp_converges: "(\<lambda>n. x ^ n /# real (fact n)) sums exp x" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
542 |
unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) |
23043 | 543 |
|
544 |
lemma sin_converges: |
|
545 |
"(%n. (if even n then 0 |
|
23177 | 546 |
else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * |
23043 | 547 |
x ^ n) sums sin(x)" |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
548 |
unfolding sin_def by (rule summable_sin [THEN summable_sums]) |
23043 | 549 |
|
550 |
lemma cos_converges: |
|
551 |
"(%n. (if even n then |
|
23177 | 552 |
-1 ^ (n div 2)/(real (fact n)) |
23043 | 553 |
else 0) * x ^ n) sums cos(x)" |
23112
2bc882fbe51c
remove division_by_zero requirement from termdiffs lemmas; cleaned up some proofs
huffman
parents:
23097
diff
changeset
|
554 |
unfolding cos_def by (rule summable_cos [THEN summable_sums]) |
23043 | 555 |
|
556 |
||
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
557 |
subsection{*Formal Derivatives of Exp, Sin, and Cos Series*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
558 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
559 |
lemma exp_fdiffs: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
560 |
"diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))" |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
561 |
by (simp add: diffs_def mult_assoc [symmetric] real_of_nat_def |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
562 |
del: mult_Suc of_nat_Suc) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
563 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
564 |
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
|
565 |
by (simp add: diffs_def) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
566 |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
567 |
lemma sin_fdiffs: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
568 |
"diffs(%n. if even n then 0 |
23177 | 569 |
else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
570 |
= (%n. if even n then |
23177 | 571 |
-1 ^ (n div 2)/(real (fact n)) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
572 |
else 0)" |
15229 | 573 |
by (auto intro!: ext |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
574 |
simp add: diffs_def divide_inverse real_of_nat_def |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
575 |
simp del: mult_Suc of_nat_Suc) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
576 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
577 |
lemma sin_fdiffs2: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
578 |
"diffs(%n. if even n then 0 |
23177 | 579 |
else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) n |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
580 |
= (if even n then |
23177 | 581 |
-1 ^ (n div 2)/(real (fact n)) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
582 |
else 0)" |
23176 | 583 |
by (simp only: sin_fdiffs) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
584 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
585 |
lemma cos_fdiffs: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
586 |
"diffs(%n. if even n then |
23177 | 587 |
-1 ^ (n div 2)/(real (fact n)) else 0) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
588 |
= (%n. - (if even n then 0 |
23177 | 589 |
else -1 ^ ((n - Suc 0)div 2)/(real (fact n))))" |
15229 | 590 |
by (auto intro!: ext |
23082
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
591 |
simp add: diffs_def divide_inverse odd_Suc_mult_two_ex real_of_nat_def |
ffef77eed382
generalize powerseries and termdiffs lemmas using axclasses
huffman
parents:
23069
diff
changeset
|
592 |
simp del: mult_Suc of_nat_Suc) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
593 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
594 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
595 |
lemma cos_fdiffs2: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
596 |
"diffs(%n. if even n then |
23177 | 597 |
-1 ^ (n div 2)/(real (fact n)) else 0) n |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
598 |
= - (if even n then 0 |
23177 | 599 |
else -1 ^ ((n - Suc 0)div 2)/(real (fact n)))" |
23176 | 600 |
by (simp only: cos_fdiffs) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
601 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
602 |
text{*Now at last we can get the derivatives of exp, sin and cos*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
603 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
604 |
lemma lemma_sin_minus: |
15546 | 605 |
"- sin x = (\<Sum>n. - ((if even n then 0 |
23177 | 606 |
else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
607 |
by (auto intro!: sums_unique sums_minus sin_converges) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
608 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
609 |
lemma lemma_exp_ext: "exp = (\<lambda>x. \<Sum>n. x ^ n /# real (fact n))" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
610 |
by (auto intro!: ext simp add: exp_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
611 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
612 |
lemma DERIV_exp [simp]: "DERIV exp x :> exp(x)" |
15229 | 613 |
apply (simp add: exp_def) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
614 |
apply (subst lemma_exp_ext) |
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
615 |
apply (subgoal_tac "DERIV (\<lambda>u. \<Sum>n. of_real (inverse (real (fact n))) * u ^ n) x :> (\<Sum>n. diffs (\<lambda>n. of_real (inverse (real (fact n)))) n * x ^ n)") |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
616 |
apply (rule_tac [2] K = "of_real (1 + norm x)" in termdiffs) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
617 |
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
|
618 |
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
|
619 |
apply (simp del: of_real_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
620 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
621 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
622 |
lemma lemma_sin_ext: |
15546 | 623 |
"sin = (%x. \<Sum>n. |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
624 |
(if even n then 0 |
23177 | 625 |
else -1 ^ ((n - Suc 0) div 2)/(real (fact n))) * |
15546 | 626 |
x ^ n)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
627 |
by (auto intro!: ext simp add: sin_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
628 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
629 |
lemma lemma_cos_ext: |
15546 | 630 |
"cos = (%x. \<Sum>n. |
23177 | 631 |
(if even n then -1 ^ (n div 2)/(real (fact n)) else 0) * |
15546 | 632 |
x ^ n)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
633 |
by (auto intro!: ext simp add: cos_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
634 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
635 |
lemma DERIV_sin [simp]: "DERIV sin x :> cos(x)" |
15229 | 636 |
apply (simp add: cos_def) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
637 |
apply (subst lemma_sin_ext) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
638 |
apply (auto simp add: sin_fdiffs2 [symmetric]) |
15229 | 639 |
apply (rule_tac K = "1 + \<bar>x\<bar>" in termdiffs) |
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
19765
diff
changeset
|
640 |
apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
641 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
642 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
643 |
lemma DERIV_cos [simp]: "DERIV cos x :> -sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
644 |
apply (subst lemma_cos_ext) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
645 |
apply (auto simp add: lemma_sin_minus cos_fdiffs2 [symmetric] minus_mult_left) |
15229 | 646 |
apply (rule_tac K = "1 + \<bar>x\<bar>" in termdiffs) |
20217
25b068a99d2b
linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents:
19765
diff
changeset
|
647 |
apply (auto intro: sin_converges cos_converges sums_summable intro!: sums_minus [THEN sums_summable] simp add: cos_fdiffs sin_fdiffs diffs_minus) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
648 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
649 |
|
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
650 |
lemma isCont_exp [simp]: "isCont exp x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
651 |
by (rule DERIV_exp [THEN DERIV_isCont]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
652 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
653 |
lemma isCont_sin [simp]: "isCont sin x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
654 |
by (rule DERIV_sin [THEN DERIV_isCont]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
655 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
656 |
lemma isCont_cos [simp]: "isCont cos x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
657 |
by (rule DERIV_cos [THEN DERIV_isCont]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
658 |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
659 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
660 |
subsection{*Properties of the Exponential Function*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
661 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
662 |
lemma exp_zero [simp]: "exp 0 = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
663 |
proof - |
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
664 |
have "(\<Sum>n = 0..<1. (0::'a) ^ n /# real (fact n)) = |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
665 |
(\<Sum>n. 0 ^ n /# real (fact n))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
666 |
by (rule sums_unique [OF series_zero], simp add: power_0_left) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
667 |
thus ?thesis by (simp add: exp_def) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
668 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
669 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
670 |
lemma setsum_head2: |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
671 |
"m \<le> n \<Longrightarrow> setsum f {m..n} = f m + setsum f {Suc m..n}" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
672 |
by (simp add: setsum_head atLeastSucAtMost_greaterThanAtMost) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
673 |
|
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
674 |
lemma setsum_cl_ivl_Suc2: |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
675 |
"(\<Sum>i=m..Suc n. f i) = (if Suc n < m then 0 else f m + (\<Sum>i=m..n. f (Suc i)))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
676 |
by (simp add: setsum_head2 setsum_shift_bounds_cl_Suc_ivl |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
677 |
del: setsum_cl_ivl_Suc) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
678 |
|
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
679 |
lemma exp_series_add: |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
680 |
fixes x y :: "'a::{real_field,recpower}" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
681 |
defines S_def: "S \<equiv> \<lambda>x n. x ^ n /# real (fact n)" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
682 |
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
|
683 |
proof (induct n) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
684 |
case 0 |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
685 |
show ?case |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
686 |
unfolding S_def by simp |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
687 |
next |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
688 |
case (Suc n) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
689 |
have S_Suc: "\<And>x n. S x (Suc n) = (x * S x n) /# real (Suc n)" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
690 |
unfolding S_def by (simp add: power_Suc del: mult_Suc) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
691 |
hence times_S: "\<And>x n. x * S x n = real (Suc n) *# S x (Suc n)" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
692 |
by simp |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
693 |
|
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
694 |
have "real (Suc n) *# S (x + y) (Suc n) = (x + y) * S (x + y) n" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
695 |
by (simp only: times_S) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
696 |
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
|
697 |
by (simp only: Suc) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
698 |
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
|
699 |
+ 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
|
700 |
by (rule left_distrib) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
701 |
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
|
702 |
+ (\<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
|
703 |
by (simp only: setsum_right_distrib mult_ac) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
704 |
also have "\<dots> = (\<Sum>i=0..n. real (Suc i) *# (S x (Suc i) * S y (n-i))) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
705 |
+ (\<Sum>i=0..n. real (Suc n-i) *# (S x i * S y (Suc n-i)))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
706 |
by (simp add: times_S Suc_diff_le) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
707 |
also have "(\<Sum>i=0..n. real (Suc i) *# (S x (Suc i) * S y (n-i))) = |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
708 |
(\<Sum>i=0..Suc n. real i *# (S x i * S y (Suc n-i)))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
709 |
by (subst setsum_cl_ivl_Suc2, simp) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
710 |
also have "(\<Sum>i=0..n. real (Suc n-i) *# (S x i * S y (Suc n-i))) = |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
711 |
(\<Sum>i=0..Suc n. real (Suc n-i) *# (S x i * S y (Suc n-i)))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
712 |
by (subst setsum_cl_ivl_Suc, simp) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
713 |
also have "(\<Sum>i=0..Suc n. real i *# (S x i * S y (Suc n-i))) + |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
714 |
(\<Sum>i=0..Suc n. real (Suc n-i) *# (S x i * S y (Suc n-i))) = |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
715 |
(\<Sum>i=0..Suc n. real (Suc n) *# (S x i * S y (Suc n-i)))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
716 |
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
|
717 |
real_of_nat_add [symmetric], simp) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
718 |
also have "\<dots> = real (Suc n) *# (\<Sum>i=0..Suc n. S x i * S y (Suc n-i))" |
23127 | 719 |
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
|
720 |
finally show |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
721 |
"S (x + y) (Suc n) = (\<Sum>i=0..Suc n. S x i * S y (Suc n - i))" |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
722 |
by (simp add: scaleR_cancel_left del: setsum_cl_ivl_Suc) |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
723 |
qed |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
724 |
|
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
725 |
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
|
726 |
unfolding exp_def |
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
727 |
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
|
728 |
|
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
729 |
lemma exp_ge_add_one_self_aux: "0 \<le> (x::real) ==> (1 + x) \<le> exp(x)" |
22998 | 730 |
apply (drule order_le_imp_less_or_eq, auto) |
15229 | 731 |
apply (simp add: exp_def) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
732 |
apply (rule real_le_trans) |
15229 | 733 |
apply (rule_tac [2] n = 2 and f = "(%n. inverse (real (fact n)) * x ^ n)" in series_pos_le) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
734 |
apply (auto intro: summable_exp simp add: numeral_2_eq_2 zero_le_power zero_le_mult_iff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
735 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
736 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
737 |
lemma exp_gt_one [simp]: "0 < (x::real) ==> 1 < exp x" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
738 |
apply (rule order_less_le_trans) |
17014
ad5ceb90877d
renamed exp_ge_add_one_self to exp_ge_add_one_self_aux
avigad
parents:
16924
diff
changeset
|
739 |
apply (rule_tac [2] exp_ge_add_one_self_aux, auto) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
740 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
741 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
742 |
lemma DERIV_exp_add_const: "DERIV (%x. exp (x + y)) x :> exp(x + y)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
743 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
744 |
have "DERIV (exp \<circ> (\<lambda>x. x + y)) x :> exp (x + y) * (1+0)" |
23069
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents:
23066
diff
changeset
|
745 |
by (fast intro: DERIV_chain DERIV_add DERIV_exp DERIV_ident DERIV_const) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
746 |
thus ?thesis by (simp add: o_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
747 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
748 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
749 |
lemma DERIV_exp_minus [simp]: "DERIV (%x. exp (-x)) x :> - exp(-x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
750 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
751 |
have "DERIV (exp \<circ> uminus) x :> exp (- x) * - 1" |
23069
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents:
23066
diff
changeset
|
752 |
by (fast intro: DERIV_chain DERIV_minus DERIV_exp DERIV_ident) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
753 |
thus ?thesis by (simp add: o_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
754 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
755 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
756 |
lemma DERIV_exp_exp_zero [simp]: "DERIV (%x. exp (x + y) * exp (- x)) x :> 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
757 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
758 |
have "DERIV (\<lambda>x. exp (x + y) * exp (- x)) x |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
759 |
:> exp (x + y) * exp (- x) + - exp (- x) * exp (x + y)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
760 |
by (fast intro: DERIV_exp_add_const DERIV_exp_minus DERIV_mult) |
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
761 |
thus ?thesis by (simp add: mult_commute) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
762 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
763 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
764 |
lemma exp_add_mult_minus [simp]: "exp(x + y)*exp(-x) = exp(y::real)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
765 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
766 |
have "\<forall>x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0" by simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
767 |
hence "exp (x + y) * exp (- x) = exp (0 + y) * exp (- 0)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
768 |
by (rule DERIV_isconst_all) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
769 |
thus ?thesis by simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
770 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
771 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
772 |
lemma exp_mult_minus [simp]: "exp x * exp(-x) = 1" |
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
773 |
by (simp add: exp_add [symmetric]) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
774 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
775 |
lemma exp_mult_minus2 [simp]: "exp(-x)*exp(x) = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
776 |
by (simp add: mult_commute) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
777 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
778 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
779 |
lemma exp_minus: "exp(-x) = inverse(exp(x))" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
780 |
by (auto intro: inverse_unique [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
781 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
782 |
text{*Proof: because every exponential can be seen as a square.*} |
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
783 |
lemma exp_ge_zero [simp]: "0 \<le> exp (x::real)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
784 |
apply (rule_tac t = x in real_sum_of_halves [THEN subst]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
785 |
apply (subst exp_add, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
786 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
787 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
788 |
lemma exp_not_eq_zero [simp]: "exp x \<noteq> 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
789 |
apply (cut_tac x = x in exp_mult_minus2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
790 |
apply (auto simp del: exp_mult_minus2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
791 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
792 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
793 |
lemma exp_gt_zero [simp]: "0 < exp (x::real)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
794 |
by (simp add: order_less_le) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
795 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
796 |
lemma inv_exp_gt_zero [simp]: "0 < inverse(exp x::real)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
797 |
by (auto intro: positive_imp_inverse_positive) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
798 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
799 |
lemma abs_exp_cancel [simp]: "\<bar>exp x::real\<bar> = exp x" |
15229 | 800 |
by auto |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
801 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
802 |
lemma exp_real_of_nat_mult: "exp(real n * x) = exp(x) ^ n" |
15251 | 803 |
apply (induct "n") |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
804 |
apply (auto simp add: real_of_nat_Suc right_distrib exp_add mult_commute) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
805 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
806 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
807 |
lemma exp_diff: "exp(x - y) = exp(x)/(exp y)" |
15229 | 808 |
apply (simp add: diff_minus divide_inverse) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
809 |
apply (simp (no_asm) add: exp_add exp_minus) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
810 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
811 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
812 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
813 |
lemma exp_less_mono: |
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
814 |
fixes x y :: real |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
815 |
assumes xy: "x < y" shows "exp x < exp y" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
816 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
817 |
have "1 < exp (y + - x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
818 |
by (rule real_less_sum_gt_zero [THEN exp_gt_one]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
819 |
hence "exp x * inverse (exp x) < exp y * inverse (exp x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
820 |
by (auto simp add: exp_add exp_minus) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
821 |
thus ?thesis |
15539 | 822 |
by (simp add: divide_inverse [symmetric] pos_less_divide_eq |
15228 | 823 |
del: divide_self_if) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
824 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
825 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
826 |
lemma exp_less_cancel: "exp (x::real) < exp y ==> x < y" |
15228 | 827 |
apply (simp add: linorder_not_le [symmetric]) |
828 |
apply (auto simp add: order_le_less exp_less_mono) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
829 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
830 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
831 |
lemma exp_less_cancel_iff [iff]: "(exp(x::real) < exp(y)) = (x < y)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
832 |
by (auto intro: exp_less_mono exp_less_cancel) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
833 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
834 |
lemma exp_le_cancel_iff [iff]: "(exp(x::real) \<le> exp(y)) = (x \<le> y)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
835 |
by (auto simp add: linorder_not_less [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
836 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
837 |
lemma exp_inj_iff [iff]: "(exp (x::real) = exp y) = (x = y)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
838 |
by (simp add: order_eq_iff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
839 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
840 |
lemma lemma_exp_total: "1 \<le> y ==> \<exists>x. 0 \<le> x & x \<le> y - 1 & exp(x::real) = y" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
841 |
apply (rule IVT) |
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
842 |
apply (auto intro: isCont_exp simp add: le_diff_eq) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
843 |
apply (subgoal_tac "1 + (y - 1) \<le> exp (y - 1)") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
844 |
apply simp |
17014
ad5ceb90877d
renamed exp_ge_add_one_self to exp_ge_add_one_self_aux
avigad
parents:
16924
diff
changeset
|
845 |
apply (rule exp_ge_add_one_self_aux, simp) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
846 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
847 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
848 |
lemma exp_total: "0 < (y::real) ==> \<exists>x. exp x = y" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
849 |
apply (rule_tac x = 1 and y = y in linorder_cases) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
850 |
apply (drule order_less_imp_le [THEN lemma_exp_total]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
851 |
apply (rule_tac [2] x = 0 in exI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
852 |
apply (frule_tac [3] real_inverse_gt_one) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
853 |
apply (drule_tac [4] order_less_imp_le [THEN lemma_exp_total], auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
854 |
apply (rule_tac x = "-x" in exI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
855 |
apply (simp add: exp_minus) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
856 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
857 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
858 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
859 |
subsection{*Properties of the Logarithmic Function*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
860 |
|
23043 | 861 |
definition |
862 |
ln :: "real => real" where |
|
863 |
"ln x = (THE u. exp u = x)" |
|
864 |
||
865 |
lemma ln_exp [simp]: "ln (exp x) = x" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
866 |
by (simp add: ln_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
867 |
|
22654
c2b6b5a9e136
new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents:
22653
diff
changeset
|
868 |
lemma exp_ln [simp]: "0 < x \<Longrightarrow> exp (ln x) = x" |
c2b6b5a9e136
new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents:
22653
diff
changeset
|
869 |
by (auto dest: exp_total) |
c2b6b5a9e136
new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents:
22653
diff
changeset
|
870 |
|
23043 | 871 |
lemma exp_ln_iff [simp]: "(exp (ln x) = x) = (0 < x)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
872 |
apply (auto dest: exp_total) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
873 |
apply (erule subst, simp) |
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 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
876 |
lemma ln_mult: "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
877 |
apply (rule exp_inj_iff [THEN iffD1]) |
22654
c2b6b5a9e136
new simp rule exp_ln; new standard proof of DERIV_exp_ln_one; changed imports
huffman
parents:
22653
diff
changeset
|
878 |
apply (simp add: exp_add exp_ln mult_pos_pos) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
879 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
880 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
881 |
lemma ln_inj_iff[simp]: "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
882 |
apply (simp only: exp_ln_iff [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
883 |
apply (erule subst)+ |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
884 |
apply simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
885 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
886 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
887 |
lemma ln_one[simp]: "ln 1 = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
888 |
by (rule exp_inj_iff [THEN iffD1], auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
889 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
890 |
lemma ln_inverse: "0 < x ==> ln(inverse x) = - ln x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
891 |
apply (rule_tac a1 = "ln x" in add_left_cancel [THEN iffD1]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
892 |
apply (auto simp add: positive_imp_inverse_positive ln_mult [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
893 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
894 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
895 |
lemma ln_div: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
896 |
"[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y" |
15229 | 897 |
apply (simp add: divide_inverse) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
898 |
apply (auto simp add: positive_imp_inverse_positive ln_mult ln_inverse) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
899 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
900 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
901 |
lemma ln_less_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
902 |
apply (simp only: exp_ln_iff [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
903 |
apply (erule subst)+ |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
904 |
apply simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
905 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
906 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
907 |
lemma ln_le_cancel_iff[simp]: "[| 0 < x; 0 < y|] ==> (ln x \<le> ln y) = (x \<le> y)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
908 |
by (auto simp add: linorder_not_less [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
909 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
910 |
lemma ln_realpow: "0 < x ==> ln(x ^ n) = real n * ln(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
911 |
by (auto dest!: exp_total simp add: exp_real_of_nat_mult [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
912 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
913 |
lemma ln_add_one_self_le_self [simp]: "0 \<le> x ==> ln(1 + x) \<le> x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
914 |
apply (rule ln_exp [THEN subst]) |
17014
ad5ceb90877d
renamed exp_ge_add_one_self to exp_ge_add_one_self_aux
avigad
parents:
16924
diff
changeset
|
915 |
apply (rule ln_le_cancel_iff [THEN iffD2]) |
ad5ceb90877d
renamed exp_ge_add_one_self to exp_ge_add_one_self_aux
avigad
parents:
16924
diff
changeset
|
916 |
apply (auto simp add: exp_ge_add_one_self_aux) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
917 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
918 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
919 |
lemma ln_less_self [simp]: "0 < x ==> ln x < x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
920 |
apply (rule order_less_le_trans) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
921 |
apply (rule_tac [2] ln_add_one_self_le_self) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
922 |
apply (rule ln_less_cancel_iff [THEN iffD2], auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
923 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
924 |
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
925 |
lemma ln_ge_zero [simp]: |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
926 |
assumes x: "1 \<le> x" shows "0 \<le> ln x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
927 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
928 |
have "0 < x" using x by arith |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
929 |
hence "exp 0 \<le> exp (ln x)" |
22915 | 930 |
by (simp add: x) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
931 |
thus ?thesis by (simp only: exp_le_cancel_iff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
932 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
933 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
934 |
lemma ln_ge_zero_imp_ge_one: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
935 |
assumes ln: "0 \<le> ln x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
936 |
and x: "0 < x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
937 |
shows "1 \<le> x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
938 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
939 |
from ln have "ln 1 \<le> ln x" by simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
940 |
thus ?thesis by (simp add: x del: ln_one) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
941 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
942 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
943 |
lemma ln_ge_zero_iff [simp]: "0 < x ==> (0 \<le> ln x) = (1 \<le> x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
944 |
by (blast intro: ln_ge_zero ln_ge_zero_imp_ge_one) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
945 |
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
946 |
lemma ln_less_zero_iff [simp]: "0 < x ==> (ln x < 0) = (x < 1)" |
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
947 |
by (insert ln_ge_zero_iff [of x], arith) |
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
948 |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
949 |
lemma ln_gt_zero: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
950 |
assumes x: "1 < x" shows "0 < ln x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
951 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
952 |
have "0 < x" using x by arith |
22915 | 953 |
hence "exp 0 < exp (ln x)" by (simp add: x) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
954 |
thus ?thesis by (simp only: exp_less_cancel_iff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
955 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
956 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
957 |
lemma ln_gt_zero_imp_gt_one: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
958 |
assumes ln: "0 < ln x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
959 |
and x: "0 < x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
960 |
shows "1 < x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
961 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
962 |
from ln have "ln 1 < ln x" by simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
963 |
thus ?thesis by (simp add: x del: ln_one) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
964 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
965 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
966 |
lemma ln_gt_zero_iff [simp]: "0 < x ==> (0 < ln x) = (1 < x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
967 |
by (blast intro: ln_gt_zero ln_gt_zero_imp_gt_one) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
968 |
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
969 |
lemma ln_eq_zero_iff [simp]: "0 < x ==> (ln x = 0) = (x = 1)" |
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
970 |
by (insert ln_less_zero_iff [of x] ln_gt_zero_iff [of x], arith) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
971 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
972 |
lemma ln_less_zero: "[| 0 < x; x < 1 |] ==> ln x < 0" |
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
973 |
by simp |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
974 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
975 |
lemma exp_ln_eq: "exp u = x ==> ln x = u" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
976 |
by auto |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
977 |
|
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
978 |
lemma isCont_ln: "0 < x \<Longrightarrow> isCont ln x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
979 |
apply (subgoal_tac "isCont ln (exp (ln x))", simp) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
980 |
apply (rule isCont_inverse_function [where f=exp], simp_all) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
981 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
982 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
983 |
lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
984 |
by simp (* TODO: put in Deriv.thy *) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
985 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
986 |
lemma DERIV_ln: "0 < x \<Longrightarrow> DERIV ln x :> inverse x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
987 |
apply (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
988 |
apply (erule lemma_DERIV_subst [OF DERIV_exp exp_ln]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
989 |
apply (simp_all add: abs_if isCont_ln) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
990 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
991 |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
992 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
993 |
subsection{*Basic Properties of the Trigonometric Functions*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
994 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
995 |
lemma sin_zero [simp]: "sin 0 = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
996 |
by (auto intro!: sums_unique [symmetric] LIMSEQ_const |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
997 |
simp add: sin_def sums_def simp del: power_0_left) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
998 |
|
15539 | 999 |
lemma lemma_series_zero2: |
1000 |
"(\<forall>m. n \<le> m --> f m = 0) --> f sums setsum f {0..<n}" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1001 |
by (auto intro: series_zero) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1002 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1003 |
lemma cos_zero [simp]: "cos 0 = 1" |
15229 | 1004 |
apply (simp add: cos_def) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1005 |
apply (rule sums_unique [symmetric]) |
23177 | 1006 |
apply (cut_tac n = 1 and f = "(%n. (if even n then -1 ^ (n div 2) / (real (fact n)) else 0) * 0 ^ n)" in lemma_series_zero2) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1007 |
apply auto |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1008 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1009 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1010 |
lemma DERIV_sin_sin_mult [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1011 |
"DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1012 |
by (rule DERIV_mult, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1013 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1014 |
lemma DERIV_sin_sin_mult2 [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1015 |
"DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1016 |
apply (cut_tac x = x in DERIV_sin_sin_mult) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1017 |
apply (auto simp add: mult_assoc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1018 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1019 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1020 |
lemma DERIV_sin_realpow2 [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1021 |
"DERIV (%x. (sin x)\<twosuperior>) x :> cos(x) * sin(x) + cos(x) * sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1022 |
by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1023 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1024 |
lemma DERIV_sin_realpow2a [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1025 |
"DERIV (%x. (sin x)\<twosuperior>) x :> 2 * cos(x) * sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1026 |
by (auto simp add: numeral_2_eq_2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1027 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1028 |
lemma DERIV_cos_cos_mult [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1029 |
"DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1030 |
by (rule DERIV_mult, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1031 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1032 |
lemma DERIV_cos_cos_mult2 [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1033 |
"DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1034 |
apply (cut_tac x = x in DERIV_cos_cos_mult) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1035 |
apply (auto simp add: mult_ac) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1036 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1037 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1038 |
lemma DERIV_cos_realpow2 [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1039 |
"DERIV (%x. (cos x)\<twosuperior>) x :> -sin(x) * cos(x) + -sin(x) * cos(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1040 |
by (auto simp add: numeral_2_eq_2 real_mult_assoc [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1041 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1042 |
lemma DERIV_cos_realpow2a [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1043 |
"DERIV (%x. (cos x)\<twosuperior>) x :> -2 * cos(x) * sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1044 |
by (auto simp add: numeral_2_eq_2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1045 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1046 |
lemma lemma_DERIV_subst: "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1047 |
by auto |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1048 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1049 |
lemma DERIV_cos_realpow2b: "DERIV (%x. (cos x)\<twosuperior>) x :> -(2 * cos(x) * sin(x))" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1050 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1051 |
apply (rule DERIV_cos_realpow2a, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1052 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1053 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1054 |
(* most useful *) |
15229 | 1055 |
lemma DERIV_cos_cos_mult3 [simp]: |
1056 |
"DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1057 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1058 |
apply (rule DERIV_cos_cos_mult2, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1059 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1060 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1061 |
lemma DERIV_sin_circle_all: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1062 |
"\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :> |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1063 |
(2*cos(x)*sin(x) - 2*cos(x)*sin(x))" |
15229 | 1064 |
apply (simp only: diff_minus, safe) |
1065 |
apply (rule DERIV_add) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1066 |
apply (auto simp add: numeral_2_eq_2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1067 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1068 |
|
15229 | 1069 |
lemma DERIV_sin_circle_all_zero [simp]: |
1070 |
"\<forall>x. DERIV (%x. (sin x)\<twosuperior> + (cos x)\<twosuperior>) x :> 0" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1071 |
by (cut_tac DERIV_sin_circle_all, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1072 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1073 |
lemma sin_cos_squared_add [simp]: "((sin x)\<twosuperior>) + ((cos x)\<twosuperior>) = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1074 |
apply (cut_tac x = x and y = 0 in DERIV_sin_circle_all_zero [THEN DERIV_isconst_all]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1075 |
apply (auto simp add: numeral_2_eq_2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1076 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1077 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1078 |
lemma sin_cos_squared_add2 [simp]: "((cos x)\<twosuperior>) + ((sin x)\<twosuperior>) = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1079 |
apply (subst real_add_commute) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1080 |
apply (simp (no_asm) del: realpow_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1081 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1082 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1083 |
lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1084 |
apply (cut_tac x = x in sin_cos_squared_add2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1085 |
apply (auto simp add: numeral_2_eq_2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1086 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1087 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1088 |
lemma sin_squared_eq: "(sin x)\<twosuperior> = 1 - (cos x)\<twosuperior>" |
15229 | 1089 |
apply (rule_tac a1 = "(cos x)\<twosuperior>" in add_right_cancel [THEN iffD1]) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1090 |
apply (simp del: realpow_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1091 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1092 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1093 |
lemma cos_squared_eq: "(cos x)\<twosuperior> = 1 - (sin x)\<twosuperior>" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1094 |
apply (rule_tac a1 = "(sin x)\<twosuperior>" in add_right_cancel [THEN iffD1]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1095 |
apply (simp del: realpow_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1096 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1097 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1098 |
lemma real_gt_one_ge_zero_add_less: "[| 1 < x; 0 \<le> y |] ==> 1 < x + (y::real)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1099 |
by arith |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1100 |
|
15081 | 1101 |
lemma abs_sin_le_one [simp]: "\<bar>sin x\<bar> \<le> 1" |
23097 | 1102 |
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
|
1103 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1104 |
lemma sin_ge_minus_one [simp]: "-1 \<le> sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1105 |
apply (insert abs_sin_le_one [of x]) |
22998 | 1106 |
apply (simp add: abs_le_iff del: abs_sin_le_one) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1107 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1108 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1109 |
lemma sin_le_one [simp]: "sin x \<le> 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1110 |
apply (insert abs_sin_le_one [of x]) |
22998 | 1111 |
apply (simp add: abs_le_iff del: abs_sin_le_one) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1112 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1113 |
|
15081 | 1114 |
lemma abs_cos_le_one [simp]: "\<bar>cos x\<bar> \<le> 1" |
23097 | 1115 |
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
|
1116 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1117 |
lemma cos_ge_minus_one [simp]: "-1 \<le> cos x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1118 |
apply (insert abs_cos_le_one [of x]) |
22998 | 1119 |
apply (simp add: abs_le_iff del: abs_cos_le_one) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1120 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1121 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1122 |
lemma cos_le_one [simp]: "cos x \<le> 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1123 |
apply (insert abs_cos_le_one [of x]) |
22998 | 1124 |
apply (simp add: abs_le_iff del: abs_cos_le_one) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1125 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1126 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1127 |
lemma DERIV_fun_pow: "DERIV g x :> m ==> |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1128 |
DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1129 |
apply (rule lemma_DERIV_subst) |
15229 | 1130 |
apply (rule_tac f = "(%x. x ^ n)" in DERIV_chain2) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1131 |
apply (rule DERIV_pow, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1132 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1133 |
|
15229 | 1134 |
lemma DERIV_fun_exp: |
1135 |
"DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1136 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1137 |
apply (rule_tac f = exp in DERIV_chain2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1138 |
apply (rule DERIV_exp, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1139 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1140 |
|
15229 | 1141 |
lemma DERIV_fun_sin: |
1142 |
"DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1143 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1144 |
apply (rule_tac f = sin in DERIV_chain2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1145 |
apply (rule DERIV_sin, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1146 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1147 |
|
15229 | 1148 |
lemma DERIV_fun_cos: |
1149 |
"DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1150 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1151 |
apply (rule_tac f = cos in DERIV_chain2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1152 |
apply (rule DERIV_cos, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1153 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1154 |
|
23069
cdfff0241c12
rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents:
23066
diff
changeset
|
1155 |
lemmas DERIV_intros = DERIV_ident DERIV_const DERIV_cos DERIV_cmult |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1156 |
DERIV_sin DERIV_exp DERIV_inverse DERIV_pow |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1157 |
DERIV_add DERIV_diff DERIV_mult DERIV_minus |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1158 |
DERIV_inverse_fun DERIV_quotient DERIV_fun_pow |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1159 |
DERIV_fun_exp DERIV_fun_sin DERIV_fun_cos |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1160 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1161 |
(* lemma *) |
15229 | 1162 |
lemma lemma_DERIV_sin_cos_add: |
1163 |
"\<forall>x. |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1164 |
DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1165 |
(cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1166 |
apply (safe, rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1167 |
apply (best intro!: DERIV_intros intro: DERIV_chain2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1168 |
--{*replaces the old @{text DERIV_tac}*} |
15229 | 1169 |
apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1170 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1171 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1172 |
lemma sin_cos_add [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1173 |
"(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1174 |
(cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1175 |
apply (cut_tac y = 0 and x = x and y7 = y |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1176 |
in lemma_DERIV_sin_cos_add [THEN DERIV_isconst_all]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1177 |
apply (auto simp add: numeral_2_eq_2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1178 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1179 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1180 |
lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1181 |
apply (cut_tac x = x and y = y in sin_cos_add) |
22969 | 1182 |
apply (simp del: sin_cos_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1183 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1184 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1185 |
lemma cos_add: "cos (x + y) = cos x * cos y - sin x * sin y" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1186 |
apply (cut_tac x = x and y = y in sin_cos_add) |
22969 | 1187 |
apply (simp del: sin_cos_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1188 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1189 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1190 |
lemma lemma_DERIV_sin_cos_minus: |
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1191 |
"\<forall>x. DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1192 |
apply (safe, rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1193 |
apply (best intro!: DERIV_intros intro: DERIV_chain2) |
15229 | 1194 |
apply (auto simp add: diff_minus left_distrib right_distrib mult_ac add_ac) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1195 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1196 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1197 |
lemma sin_cos_minus [simp]: |
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1198 |
"(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0" |
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1199 |
apply (cut_tac y = 0 and x = x |
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1200 |
in lemma_DERIV_sin_cos_minus [THEN DERIV_isconst_all]) |
22969 | 1201 |
apply simp |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1202 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1203 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1204 |
lemma sin_minus [simp]: "sin (-x) = -sin(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1205 |
apply (cut_tac x = x in sin_cos_minus) |
22969 | 1206 |
apply (simp del: sin_cos_minus) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1207 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1208 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1209 |
lemma cos_minus [simp]: "cos (-x) = cos(x)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1210 |
apply (cut_tac x = x in sin_cos_minus) |
22969 | 1211 |
apply (simp del: sin_cos_minus) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1212 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1213 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1214 |
lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" |
22969 | 1215 |
by (simp add: diff_minus sin_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1216 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1217 |
lemma sin_diff2: "sin (x - y) = cos y * sin x - sin y * cos x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1218 |
by (simp add: sin_diff mult_commute) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1219 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1220 |
lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" |
22969 | 1221 |
by (simp add: diff_minus cos_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1222 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1223 |
lemma cos_diff2: "cos (x - y) = cos y * cos x + sin y * sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1224 |
by (simp add: cos_diff mult_commute) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1225 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1226 |
lemma sin_double [simp]: "sin(2 * x) = 2* sin x * cos x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1227 |
by (cut_tac x = x and y = x in sin_add, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1228 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1229 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1230 |
lemma cos_double: "cos(2* x) = ((cos x)\<twosuperior>) - ((sin x)\<twosuperior>)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1231 |
apply (cut_tac x = x and y = x in cos_add) |
22969 | 1232 |
apply (simp add: power2_eq_square) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1233 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1234 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1235 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1236 |
subsection{*The Constant Pi*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1237 |
|
23043 | 1238 |
definition |
1239 |
pi :: "real" where |
|
23053 | 1240 |
"pi = 2 * (THE x. 0 \<le> (x::real) & x \<le> 2 & cos x = 0)" |
23043 | 1241 |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1242 |
text{*Show that there's a least positive @{term x} with @{term "cos(x) = 0"}; |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1243 |
hence define pi.*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1244 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1245 |
lemma sin_paired: |
23177 | 1246 |
"(%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
|
1247 |
sums sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1248 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1249 |
have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2. |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1250 |
(if even k then 0 |
23177 | 1251 |
else -1 ^ ((k - Suc 0) div 2) / real (fact k)) * |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1252 |
x ^ k) |
23176 | 1253 |
sums sin x" |
1254 |
unfolding sin_def |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1255 |
by (rule sin_converges [THEN sums_summable, THEN sums_group], simp) |
23176 | 1256 |
thus ?thesis by (simp add: mult_ac) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1257 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1258 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1259 |
lemma sin_gt_zero: "[|0 < x; x < 2 |] ==> 0 < sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1260 |
apply (subgoal_tac |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1261 |
"(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2. |
23177 | 1262 |
-1 ^ k / real (fact (2 * k + 1)) * x ^ (2 * k + 1)) |
1263 |
sums (\<Sum>n. -1 ^ n / real (fact (2 * n + 1)) * x ^ (2 * n + 1))") |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1264 |
prefer 2 |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1265 |
apply (rule sin_paired [THEN sums_summable, THEN sums_group], simp) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1266 |
apply (rotate_tac 2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1267 |
apply (drule sin_paired [THEN sums_unique, THEN ssubst]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1268 |
apply (auto simp del: fact_Suc realpow_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1269 |
apply (frule sums_unique) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1270 |
apply (auto simp del: fact_Suc realpow_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1271 |
apply (rule_tac n1 = 0 in series_pos_less [THEN [2] order_le_less_trans]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1272 |
apply (auto simp del: fact_Suc realpow_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1273 |
apply (erule sums_summable) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1274 |
apply (case_tac "m=0") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1275 |
apply (simp (no_asm_simp)) |
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
1276 |
apply (subgoal_tac "6 * (x * (x * x) / real (Suc (Suc (Suc (Suc (Suc (Suc 0))))))) < 6 * x") |
15539 | 1277 |
apply (simp only: mult_less_cancel_left, simp) |
1278 |
apply (simp (no_asm_simp) add: numeral_2_eq_2 [symmetric] mult_assoc [symmetric]) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1279 |
apply (subgoal_tac "x*x < 2*3", simp) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1280 |
apply (rule mult_strict_mono) |
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1281 |
apply (auto simp add: real_0_less_add_iff real_of_nat_Suc simp del: fact_Suc) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1282 |
apply (subst fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1283 |
apply (subst fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1284 |
apply (subst fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1285 |
apply (subst fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1286 |
apply (subst real_of_nat_mult) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1287 |
apply (subst real_of_nat_mult) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1288 |
apply (subst real_of_nat_mult) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1289 |
apply (subst real_of_nat_mult) |
15539 | 1290 |
apply (simp (no_asm) add: divide_inverse del: fact_Suc) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1291 |
apply (auto simp add: mult_assoc [symmetric] simp del: fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1292 |
apply (rule_tac c="real (Suc (Suc (4*m)))" in mult_less_imp_less_right) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1293 |
apply (auto simp add: mult_assoc simp del: fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1294 |
apply (rule_tac c="real (Suc (Suc (Suc (4*m))))" in mult_less_imp_less_right) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1295 |
apply (auto simp add: mult_assoc mult_less_cancel_left simp del: fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1296 |
apply (subgoal_tac "x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * x)") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1297 |
apply (erule ssubst)+ |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1298 |
apply (auto simp del: fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1299 |
apply (subgoal_tac "0 < x ^ (4 * m) ") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1300 |
prefer 2 apply (simp only: zero_less_power) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1301 |
apply (simp (no_asm_simp) add: mult_less_cancel_left) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1302 |
apply (rule mult_strict_mono) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1303 |
apply (simp_all (no_asm_simp)) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1304 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1305 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1306 |
lemma sin_gt_zero1: "[|0 < x; x < 2 |] ==> 0 < sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1307 |
by (auto intro: sin_gt_zero) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1308 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1309 |
lemma cos_double_less_one: "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1310 |
apply (cut_tac x = x in sin_gt_zero1) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1311 |
apply (auto simp add: cos_squared_eq cos_double) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1312 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1313 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1314 |
lemma cos_paired: |
23177 | 1315 |
"(%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
|
1316 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1317 |
have "(\<lambda>n. \<Sum>k = n * 2..<n * 2 + 2. |
23177 | 1318 |
(if even k then -1 ^ (k div 2) / real (fact k) else 0) * |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1319 |
x ^ k) |
23176 | 1320 |
sums cos x" |
1321 |
unfolding cos_def |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1322 |
by (rule cos_converges [THEN sums_summable, THEN sums_group], simp) |
23176 | 1323 |
thus ?thesis by (simp add: mult_ac) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1324 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1325 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1326 |
declare zero_less_power [simp] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1327 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1328 |
lemma fact_lemma: "real (n::nat) * 4 = real (4 * n)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1329 |
by simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1330 |
|
23053 | 1331 |
lemma cos_two_less_zero [simp]: "cos (2) < 0" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1332 |
apply (cut_tac x = 2 in cos_paired) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1333 |
apply (drule sums_minus) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1334 |
apply (rule neg_less_iff_less [THEN iffD1]) |
15539 | 1335 |
apply (frule sums_unique, auto) |
1336 |
apply (rule_tac y = |
|
23177 | 1337 |
"\<Sum>n=0..< Suc(Suc(Suc 0)). - (-1 ^ n / (real(fact (2*n))) * 2 ^ (2*n))" |
15481 | 1338 |
in order_less_trans) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1339 |
apply (simp (no_asm) add: fact_num_eq_if realpow_num_eq_if del: fact_Suc realpow_Suc) |
15561 | 1340 |
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
|
1341 |
apply (rule sumr_pos_lt_pair) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1342 |
apply (erule sums_summable, safe) |
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1343 |
apply (simp (no_asm) add: divide_inverse real_0_less_add_iff mult_assoc [symmetric] |
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1344 |
del: fact_Suc) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1345 |
apply (rule real_mult_inverse_cancel2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1346 |
apply (rule real_of_nat_fact_gt_zero)+ |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1347 |
apply (simp (no_asm) add: mult_assoc [symmetric] del: fact_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1348 |
apply (subst fact_lemma) |
15481 | 1349 |
apply (subst fact_Suc [of "Suc (Suc (Suc (Suc (Suc (Suc (Suc (4 * d)))))))"]) |
1350 |
apply (simp only: real_of_nat_mult) |
|
23007
e025695d9b0e
use mult_strict_mono instead of real_mult_less_mono
huffman
parents:
22998
diff
changeset
|
1351 |
apply (rule mult_strict_mono, force) |
e025695d9b0e
use mult_strict_mono instead of real_mult_less_mono
huffman
parents:
22998
diff
changeset
|
1352 |
apply (rule_tac [3] real_of_nat_fact_ge_zero) |
15481 | 1353 |
prefer 2 apply force |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1354 |
apply (rule real_of_nat_less_iff [THEN iffD2]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1355 |
apply (rule fact_less_mono, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1356 |
done |
23053 | 1357 |
|
1358 |
lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] |
|
1359 |
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
|
1360 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1361 |
lemma cos_is_zero: "EX! x. 0 \<le> x & x \<le> 2 & cos x = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1362 |
apply (subgoal_tac "\<exists>x. 0 \<le> x & x \<le> 2 & cos x = 0") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1363 |
apply (rule_tac [2] IVT2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1364 |
apply (auto intro: DERIV_isCont DERIV_cos) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1365 |
apply (cut_tac x = xa and y = y in linorder_less_linear) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1366 |
apply (rule ccontr) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1367 |
apply (subgoal_tac " (\<forall>x. cos differentiable x) & (\<forall>x. isCont cos x) ") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1368 |
apply (auto intro: DERIV_cos DERIV_isCont simp add: differentiable_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1369 |
apply (drule_tac f = cos in Rolle) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1370 |
apply (drule_tac [5] f = cos in Rolle) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1371 |
apply (auto dest!: DERIV_cos [THEN DERIV_unique] simp add: differentiable_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1372 |
apply (drule_tac y1 = xa in order_le_less_trans [THEN sin_gt_zero]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1373 |
apply (assumption, rule_tac y=y in order_less_le_trans, simp_all) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1374 |
apply (drule_tac y1 = y in order_le_less_trans [THEN sin_gt_zero], assumption, simp_all) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1375 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1376 |
|
23053 | 1377 |
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
|
1378 |
by (simp add: pi_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1379 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1380 |
lemma cos_pi_half [simp]: "cos (pi / 2) = 0" |
23053 | 1381 |
by (simp add: pi_half cos_is_zero [THEN theI']) |
1382 |
||
1383 |
lemma pi_half_gt_zero [simp]: "0 < pi / 2" |
|
1384 |
apply (rule order_le_neq_trans) |
|
1385 |
apply (simp add: pi_half cos_is_zero [THEN theI']) |
|
1386 |
apply (rule notI, drule arg_cong [where f=cos], simp) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1387 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1388 |
|
23053 | 1389 |
lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric] |
1390 |
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
|
1391 |
|
23053 | 1392 |
lemma pi_half_less_two [simp]: "pi / 2 < 2" |
1393 |
apply (rule order_le_neq_trans) |
|
1394 |
apply (simp add: pi_half cos_is_zero [THEN theI']) |
|
1395 |
apply (rule notI, drule arg_cong [where f=cos], simp) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1396 |
done |
23053 | 1397 |
|
1398 |
lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq] |
|
1399 |
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
|
1400 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1401 |
lemma pi_gt_zero [simp]: "0 < pi" |
23053 | 1402 |
by (insert pi_half_gt_zero, simp) |
1403 |
||
1404 |
lemma pi_ge_zero [simp]: "0 \<le> pi" |
|
1405 |
by (rule pi_gt_zero [THEN order_less_imp_le]) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1406 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1407 |
lemma pi_neq_zero [simp]: "pi \<noteq> 0" |
22998 | 1408 |
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
|
1409 |
|
23053 | 1410 |
lemma pi_not_less_zero [simp]: "\<not> pi < 0" |
1411 |
by (simp add: linorder_not_less) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1412 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1413 |
lemma minus_pi_half_less_zero [simp]: "-(pi/2) < 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1414 |
by auto |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1415 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1416 |
lemma sin_pi_half [simp]: "sin(pi/2) = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1417 |
apply (cut_tac x = "pi/2" in sin_cos_squared_add2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1418 |
apply (cut_tac sin_gt_zero [OF pi_half_gt_zero pi_half_less_two]) |
23053 | 1419 |
apply (simp add: power2_eq_square) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1420 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1421 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1422 |
lemma cos_pi [simp]: "cos pi = -1" |
15539 | 1423 |
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
|
1424 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1425 |
lemma sin_pi [simp]: "sin pi = 0" |
15539 | 1426 |
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
|
1427 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1428 |
lemma sin_cos_eq: "sin x = cos (pi/2 - x)" |
15229 | 1429 |
by (simp add: diff_minus cos_add) |
23053 | 1430 |
declare sin_cos_eq [symmetric, simp] |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1431 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1432 |
lemma minus_sin_cos_eq: "-sin x = cos (x + pi/2)" |
15229 | 1433 |
by (simp add: cos_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1434 |
declare minus_sin_cos_eq [symmetric, simp] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1435 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1436 |
lemma cos_sin_eq: "cos x = sin (pi/2 - x)" |
15229 | 1437 |
by (simp add: diff_minus sin_add) |
23053 | 1438 |
declare cos_sin_eq [symmetric, simp] |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1439 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1440 |
lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" |
15229 | 1441 |
by (simp add: sin_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1442 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1443 |
lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" |
15229 | 1444 |
by (simp add: sin_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1445 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1446 |
lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" |
15229 | 1447 |
by (simp add: cos_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1448 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1449 |
lemma sin_periodic [simp]: "sin (x + 2*pi) = sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1450 |
by (simp add: sin_add cos_double) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1451 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1452 |
lemma cos_periodic [simp]: "cos (x + 2*pi) = cos x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1453 |
by (simp add: cos_add cos_double) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1454 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1455 |
lemma cos_npi [simp]: "cos (real n * pi) = -1 ^ n" |
15251 | 1456 |
apply (induct "n") |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1457 |
apply (auto simp add: real_of_nat_Suc left_distrib) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1458 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1459 |
|
15383 | 1460 |
lemma cos_npi2 [simp]: "cos (pi * real n) = -1 ^ n" |
1461 |
proof - |
|
1462 |
have "cos (pi * real n) = cos (real n * pi)" by (simp only: mult_commute) |
|
1463 |
also have "... = -1 ^ n" by (rule cos_npi) |
|
1464 |
finally show ?thesis . |
|
1465 |
qed |
|
1466 |
||
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1467 |
lemma sin_npi [simp]: "sin (real (n::nat) * pi) = 0" |
15251 | 1468 |
apply (induct "n") |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1469 |
apply (auto simp add: real_of_nat_Suc left_distrib) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1470 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1471 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1472 |
lemma sin_npi2 [simp]: "sin (pi * real (n::nat)) = 0" |
15383 | 1473 |
by (simp add: mult_commute [of pi]) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1474 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1475 |
lemma cos_two_pi [simp]: "cos (2 * pi) = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1476 |
by (simp add: cos_double) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1477 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1478 |
lemma sin_two_pi [simp]: "sin (2 * pi) = 0" |
15229 | 1479 |
by simp |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1480 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1481 |
lemma sin_gt_zero2: "[| 0 < x; x < pi/2 |] ==> 0 < sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1482 |
apply (rule sin_gt_zero, assumption) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1483 |
apply (rule order_less_trans, assumption) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1484 |
apply (rule pi_half_less_two) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1485 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1486 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1487 |
lemma sin_less_zero: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1488 |
assumes lb: "- pi/2 < x" and "x < 0" shows "sin x < 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1489 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1490 |
have "0 < sin (- x)" using prems by (simp only: sin_gt_zero2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1491 |
thus ?thesis by simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1492 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1493 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1494 |
lemma pi_less_4: "pi < 4" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1495 |
by (cut_tac pi_half_less_two, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1496 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1497 |
lemma cos_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < cos x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1498 |
apply (cut_tac pi_less_4) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1499 |
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
|
1500 |
apply (cut_tac cos_is_zero, safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1501 |
apply (rename_tac y z) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1502 |
apply (drule_tac x = y in spec) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1503 |
apply (drule_tac x = "pi/2" in spec, simp) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1504 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1505 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1506 |
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
|
1507 |
apply (rule_tac x = x and y = 0 in linorder_cases) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1508 |
apply (rule cos_minus [THEN subst]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1509 |
apply (rule cos_gt_zero) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1510 |
apply (auto intro: cos_gt_zero) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1511 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1512 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1513 |
lemma cos_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
|
1514 |
apply (auto simp add: order_le_less cos_gt_zero_pi) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1515 |
apply (subgoal_tac "x = pi/2", auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1516 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1517 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1518 |
lemma sin_gt_zero_pi: "[| 0 < x; x < pi |] ==> 0 < sin x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1519 |
apply (subst sin_cos_eq) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1520 |
apply (rotate_tac 1) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1521 |
apply (drule real_sum_of_halves [THEN ssubst]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1522 |
apply (auto intro!: cos_gt_zero_pi simp del: sin_cos_eq [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1523 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1524 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1525 |
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
|
1526 |
by (auto simp add: order_le_less sin_gt_zero_pi) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1527 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1528 |
lemma cos_total: "[| -1 \<le> y; y \<le> 1 |] ==> EX! x. 0 \<le> x & x \<le> pi & (cos x = y)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1529 |
apply (subgoal_tac "\<exists>x. 0 \<le> x & x \<le> pi & cos x = y") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1530 |
apply (rule_tac [2] IVT2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1531 |
apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1532 |
apply (cut_tac x = xa and y = y in linorder_less_linear) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1533 |
apply (rule ccontr, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1534 |
apply (drule_tac f = cos in Rolle) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1535 |
apply (drule_tac [5] f = cos in Rolle) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1536 |
apply (auto intro: order_less_imp_le DERIV_isCont DERIV_cos |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1537 |
dest!: DERIV_cos [THEN DERIV_unique] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1538 |
simp add: differentiable_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1539 |
apply (auto dest: sin_gt_zero_pi [OF order_le_less_trans order_less_le_trans]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1540 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1541 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1542 |
lemma sin_total: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1543 |
"[| -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
|
1544 |
apply (rule ccontr) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1545 |
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 | 1546 |
apply (erule contrapos_np) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1547 |
apply (simp del: minus_sin_cos_eq [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1548 |
apply (cut_tac y="-y" in cos_total, simp) apply simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1549 |
apply (erule ex1E) |
15229 | 1550 |
apply (rule_tac a = "x - (pi/2)" in ex1I) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1551 |
apply (simp (no_asm) add: real_add_assoc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1552 |
apply (rotate_tac 3) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1553 |
apply (drule_tac x = "xa + pi/2" in spec, safe, simp_all) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1554 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1555 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1556 |
lemma reals_Archimedean4: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1557 |
"[| 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
|
1558 |
apply (auto dest!: reals_Archimedean3) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1559 |
apply (drule_tac x = x in spec, clarify) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1560 |
apply (subgoal_tac "x < real(LEAST m::nat. x < real m * y) * y") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1561 |
prefer 2 apply (erule LeastI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1562 |
apply (case_tac "LEAST m::nat. x < real m * y", simp) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1563 |
apply (subgoal_tac "~ x < real nat * y") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1564 |
prefer 2 apply (rule not_less_Least, simp, force) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1565 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1566 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1567 |
(* Pre Isabelle99-2 proof was simpler- numerals arithmetic |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1568 |
now causes some unwanted re-arrangements of literals! *) |
15229 | 1569 |
lemma cos_zero_lemma: |
1570 |
"[| 0 \<le> x; cos x = 0 |] ==> |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1571 |
\<exists>n::nat. ~even n & x = real n * (pi/2)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1572 |
apply (drule pi_gt_zero [THEN reals_Archimedean4], safe) |
15086 | 1573 |
apply (subgoal_tac "0 \<le> x - real n * pi & |
1574 |
(x - real n * pi) \<le> pi & (cos (x - real n * pi) = 0) ") |
|
1575 |
apply (auto simp add: compare_rls) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1576 |
prefer 3 apply (simp add: cos_diff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1577 |
prefer 2 apply (simp add: real_of_nat_Suc left_distrib) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1578 |
apply (simp add: cos_diff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1579 |
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
|
1580 |
apply (rule_tac [2] cos_total, safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1581 |
apply (drule_tac x = "x - real n * pi" in spec) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1582 |
apply (drule_tac x = "pi/2" in spec) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1583 |
apply (simp add: cos_diff) |
15229 | 1584 |
apply (rule_tac x = "Suc (2 * n)" in exI) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1585 |
apply (simp add: real_of_nat_Suc left_distrib, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1586 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1587 |
|
15229 | 1588 |
lemma sin_zero_lemma: |
1589 |
"[| 0 \<le> x; sin x = 0 |] ==> |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1590 |
\<exists>n::nat. even n & x = real n * (pi/2)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1591 |
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
|
1592 |
apply (clarify, rule_tac x = "n - 1" in exI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1593 |
apply (force simp add: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) |
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1594 |
apply (rule cos_zero_lemma) |
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
1595 |
apply (simp_all add: add_increasing) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1596 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1597 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1598 |
|
15229 | 1599 |
lemma cos_zero_iff: |
1600 |
"(cos x = 0) = |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1601 |
((\<exists>n::nat. ~even n & (x = real n * (pi/2))) | |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1602 |
(\<exists>n::nat. ~even n & (x = -(real n * (pi/2)))))" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1603 |
apply (rule iffI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1604 |
apply (cut_tac linorder_linear [of 0 x], safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1605 |
apply (drule cos_zero_lemma, assumption+) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1606 |
apply (cut_tac x="-x" in cos_zero_lemma, simp, simp) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1607 |
apply (force simp add: minus_equation_iff [of x]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1608 |
apply (auto simp only: odd_Suc_mult_two_ex real_of_nat_Suc left_distrib) |
15539 | 1609 |
apply (auto simp add: cos_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1610 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1611 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1612 |
(* ditto: but to a lesser extent *) |
15229 | 1613 |
lemma sin_zero_iff: |
1614 |
"(sin x = 0) = |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1615 |
((\<exists>n::nat. even n & (x = real n * (pi/2))) | |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1616 |
(\<exists>n::nat. even n & (x = -(real n * (pi/2)))))" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1617 |
apply (rule iffI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1618 |
apply (cut_tac linorder_linear [of 0 x], safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1619 |
apply (drule sin_zero_lemma, assumption+) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1620 |
apply (cut_tac x="-x" in sin_zero_lemma, simp, simp, safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1621 |
apply (force simp add: minus_equation_iff [of x]) |
15539 | 1622 |
apply (auto simp add: even_mult_two_ex) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1623 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1624 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1625 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1626 |
subsection{*Tangent*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1627 |
|
23043 | 1628 |
definition |
1629 |
tan :: "real => real" where |
|
1630 |
"tan x = (sin x)/(cos x)" |
|
1631 |
||
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1632 |
lemma tan_zero [simp]: "tan 0 = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1633 |
by (simp add: tan_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1634 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1635 |
lemma tan_pi [simp]: "tan pi = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1636 |
by (simp add: tan_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1637 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1638 |
lemma tan_npi [simp]: "tan (real (n::nat) * pi) = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1639 |
by (simp add: tan_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1640 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1641 |
lemma tan_minus [simp]: "tan (-x) = - tan x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1642 |
by (simp add: tan_def minus_mult_left) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1643 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1644 |
lemma tan_periodic [simp]: "tan (x + 2*pi) = tan x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1645 |
by (simp add: tan_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1646 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1647 |
lemma lemma_tan_add1: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1648 |
"[| cos x \<noteq> 0; cos y \<noteq> 0 |] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1649 |
==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)" |
15229 | 1650 |
apply (simp add: tan_def divide_inverse) |
1651 |
apply (auto simp del: inverse_mult_distrib |
|
1652 |
simp add: inverse_mult_distrib [symmetric] mult_ac) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1653 |
apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) |
15229 | 1654 |
apply (auto simp del: inverse_mult_distrib |
1655 |
simp add: mult_assoc left_diff_distrib cos_add) |
|
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15229
diff
changeset
|
1656 |
done |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1657 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1658 |
lemma add_tan_eq: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1659 |
"[| cos x \<noteq> 0; cos y \<noteq> 0 |] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1660 |
==> tan x + tan y = sin(x + y)/(cos x * cos y)" |
15229 | 1661 |
apply (simp add: tan_def) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1662 |
apply (rule_tac c1 = "cos x * cos y" in real_mult_right_cancel [THEN subst]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1663 |
apply (auto simp add: mult_assoc left_distrib) |
15539 | 1664 |
apply (simp add: sin_add) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1665 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1666 |
|
15229 | 1667 |
lemma tan_add: |
1668 |
"[| 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
|
1669 |
==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1670 |
apply (simp (no_asm_simp) add: add_tan_eq lemma_tan_add1) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1671 |
apply (simp add: tan_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1672 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1673 |
|
15229 | 1674 |
lemma tan_double: |
1675 |
"[| cos x \<noteq> 0; cos (2 * x) \<noteq> 0 |] |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1676 |
==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1677 |
apply (insert tan_add [of x x]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1678 |
apply (simp add: mult_2 [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1679 |
apply (auto simp add: numeral_2_eq_2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1680 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1681 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1682 |
lemma tan_gt_zero: "[| 0 < x; x < pi/2 |] ==> 0 < tan x" |
15229 | 1683 |
by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1684 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1685 |
lemma tan_less_zero: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1686 |
assumes lb: "- pi/2 < x" and "x < 0" shows "tan x < 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1687 |
proof - |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1688 |
have "0 < tan (- x)" using prems by (simp only: tan_gt_zero) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1689 |
thus ?thesis by simp |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1690 |
qed |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1691 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1692 |
lemma lemma_DERIV_tan: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1693 |
"cos x \<noteq> 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse((cos x)\<twosuperior>)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1694 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1695 |
apply (best intro!: DERIV_intros intro: DERIV_chain2) |
15079
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents:
15077
diff
changeset
|
1696 |
apply (auto simp add: divide_inverse numeral_2_eq_2) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1697 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1698 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1699 |
lemma DERIV_tan [simp]: "cos x \<noteq> 0 ==> DERIV tan x :> inverse((cos x)\<twosuperior>)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1700 |
by (auto dest: lemma_DERIV_tan simp add: tan_def [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1701 |
|
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1702 |
lemma isCont_tan [simp]: "cos x \<noteq> 0 ==> isCont tan x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1703 |
by (rule DERIV_tan [THEN DERIV_isCont]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1704 |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1705 |
lemma LIM_cos_div_sin [simp]: "(%x. cos(x)/sin(x)) -- pi/2 --> 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1706 |
apply (subgoal_tac "(\<lambda>x. cos x * inverse (sin x)) -- pi * inverse 2 --> 0*1") |
15229 | 1707 |
apply (simp add: divide_inverse [symmetric]) |
22613 | 1708 |
apply (rule LIM_mult) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1709 |
apply (rule_tac [2] inverse_1 [THEN subst]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1710 |
apply (rule_tac [2] LIM_inverse) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1711 |
apply (simp_all add: divide_inverse [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1712 |
apply (simp_all only: isCont_def [symmetric] cos_pi_half [symmetric] sin_pi_half [symmetric]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1713 |
apply (blast intro!: DERIV_isCont DERIV_sin DERIV_cos)+ |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1714 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1715 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1716 |
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
|
1717 |
apply (cut_tac LIM_cos_div_sin) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1718 |
apply (simp only: LIM_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1719 |
apply (drule_tac x = "inverse y" in spec, safe, force) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1720 |
apply (drule_tac ?d1.0 = s in pi_half_gt_zero [THEN [2] real_lbound_gt_zero], safe) |
15229 | 1721 |
apply (rule_tac x = "(pi/2) - e" in exI) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1722 |
apply (simp (no_asm_simp)) |
15229 | 1723 |
apply (drule_tac x = "(pi/2) - e" in spec) |
1724 |
apply (auto simp add: tan_def) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1725 |
apply (rule inverse_less_iff_less [THEN iffD1]) |
15079
2ef899e4526d
conversion of Hyperreal/MacLaurin_lemmas to Isar script
paulson
parents:
15077
diff
changeset
|
1726 |
apply (auto simp add: divide_inverse) |
15229 | 1727 |
apply (rule real_mult_order) |
1728 |
apply (subgoal_tac [3] "0 < sin e & 0 < cos e") |
|
1729 |
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
|
1730 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1731 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1732 |
lemma tan_total_pos: "0 \<le> y ==> \<exists>x. 0 \<le> x & x < pi/2 & tan x = y" |
22998 | 1733 |
apply (frule order_le_imp_less_or_eq, safe) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1734 |
prefer 2 apply force |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1735 |
apply (drule lemma_tan_total, safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1736 |
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
|
1737 |
apply (auto intro!: DERIV_tan [THEN DERIV_isCont]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1738 |
apply (drule_tac y = xa in order_le_imp_less_or_eq) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1739 |
apply (auto dest: cos_gt_zero) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1740 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1741 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1742 |
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
|
1743 |
apply (cut_tac linorder_linear [of 0 y], safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1744 |
apply (drule tan_total_pos) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1745 |
apply (cut_tac [2] y="-y" in tan_total_pos, safe) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1746 |
apply (rule_tac [3] x = "-x" in exI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1747 |
apply (auto intro!: exI) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1748 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1749 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1750 |
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
|
1751 |
apply (cut_tac y = y in lemma_tan_total1, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1752 |
apply (cut_tac x = xa and y = y in linorder_less_linear, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1753 |
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
|
1754 |
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
|
1755 |
apply (rule_tac [4] Rolle) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1756 |
apply (rule_tac [2] Rolle) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1757 |
apply (auto intro!: DERIV_tan DERIV_isCont exI |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1758 |
simp add: differentiable_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1759 |
txt{*Now, simulate TRYALL*} |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1760 |
apply (rule_tac [!] DERIV_tan asm_rl) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1761 |
apply (auto dest!: DERIV_unique [OF _ DERIV_tan] |
22998 | 1762 |
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
|
1763 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1764 |
|
23043 | 1765 |
|
1766 |
subsection {* Inverse Trigonometric Functions *} |
|
1767 |
||
1768 |
definition |
|
1769 |
arcsin :: "real => real" where |
|
1770 |
"arcsin y = (THE x. -(pi/2) \<le> x & x \<le> pi/2 & sin x = y)" |
|
1771 |
||
1772 |
definition |
|
1773 |
arccos :: "real => real" where |
|
1774 |
"arccos y = (THE x. 0 \<le> x & x \<le> pi & cos x = y)" |
|
1775 |
||
1776 |
definition |
|
1777 |
arctan :: "real => real" where |
|
1778 |
"arctan y = (THE x. -(pi/2) < x & x < pi/2 & tan x = y)" |
|
1779 |
||
15229 | 1780 |
lemma arcsin: |
1781 |
"[| -1 \<le> y; y \<le> 1 |] |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1782 |
==> -(pi/2) \<le> arcsin y & |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1783 |
arcsin y \<le> pi/2 & sin(arcsin y) = y" |
23011 | 1784 |
unfolding arcsin_def by (rule theI' [OF sin_total]) |
1785 |
||
1786 |
lemma arcsin_pi: |
|
1787 |
"[| -1 \<le> y; y \<le> 1 |] |
|
1788 |
==> -(pi/2) \<le> arcsin y & arcsin y \<le> pi & sin(arcsin y) = y" |
|
1789 |
apply (drule (1) arcsin) |
|
1790 |
apply (force intro: order_trans) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1791 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1792 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1793 |
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
|
1794 |
by (blast dest: arcsin) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1795 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1796 |
lemma arcsin_bounded: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1797 |
"[| -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
|
1798 |
by (blast dest: arcsin) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1799 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1800 |
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
|
1801 |
by (blast dest: arcsin) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1802 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1803 |
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
|
1804 |
by (blast dest: arcsin) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1805 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1806 |
lemma arcsin_lt_bounded: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1807 |
"[| -1 < y; y < 1 |] ==> -(pi/2) < arcsin y & arcsin y < pi/2" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1808 |
apply (frule order_less_imp_le) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1809 |
apply (frule_tac y = y in order_less_imp_le) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1810 |
apply (frule arcsin_bounded) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1811 |
apply (safe, simp) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1812 |
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
|
1813 |
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
|
1814 |
apply (drule_tac [!] f = sin in arg_cong, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1815 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1816 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1817 |
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
|
1818 |
apply (unfold arcsin_def) |
23011 | 1819 |
apply (rule the1_equality) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1820 |
apply (rule sin_total, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1821 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1822 |
|
22975 | 1823 |
lemma arccos: |
15229 | 1824 |
"[| -1 \<le> y; y \<le> 1 |] |
22975 | 1825 |
==> 0 \<le> arccos y & arccos y \<le> pi & cos(arccos y) = y" |
23011 | 1826 |
unfolding arccos_def by (rule theI' [OF cos_total]) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1827 |
|
22975 | 1828 |
lemma cos_arccos [simp]: "[| -1 \<le> y; y \<le> 1 |] ==> cos(arccos y) = y" |
1829 |
by (blast dest: arccos) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1830 |
|
22975 | 1831 |
lemma arccos_bounded: "[| -1 \<le> y; y \<le> 1 |] ==> 0 \<le> arccos y & arccos y \<le> pi" |
1832 |
by (blast dest: arccos) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1833 |
|
22975 | 1834 |
lemma arccos_lbound: "[| -1 \<le> y; y \<le> 1 |] ==> 0 \<le> arccos y" |
1835 |
by (blast dest: arccos) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1836 |
|
22975 | 1837 |
lemma arccos_ubound: "[| -1 \<le> y; y \<le> 1 |] ==> arccos y \<le> pi" |
1838 |
by (blast dest: arccos) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1839 |
|
22975 | 1840 |
lemma arccos_lt_bounded: |
15229 | 1841 |
"[| -1 < y; y < 1 |] |
22975 | 1842 |
==> 0 < arccos y & arccos y < pi" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1843 |
apply (frule order_less_imp_le) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1844 |
apply (frule_tac y = y in order_less_imp_le) |
22975 | 1845 |
apply (frule arccos_bounded, auto) |
1846 |
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
|
1847 |
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
|
1848 |
apply (drule_tac [!] f = cos in arg_cong, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1849 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1850 |
|
22975 | 1851 |
lemma arccos_cos: "[|0 \<le> x; x \<le> pi |] ==> arccos(cos x) = x" |
1852 |
apply (simp add: arccos_def) |
|
23011 | 1853 |
apply (auto intro!: the1_equality cos_total) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1854 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1855 |
|
22975 | 1856 |
lemma arccos_cos2: "[|x \<le> 0; -pi \<le> x |] ==> arccos(cos x) = -x" |
1857 |
apply (simp add: arccos_def) |
|
23011 | 1858 |
apply (auto intro!: the1_equality cos_total) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1859 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1860 |
|
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1861 |
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
|
1862 |
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
|
1863 |
apply (rule power2_eq_imp_eq) |
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1864 |
apply (simp add: cos_squared_eq) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1865 |
apply (rule cos_ge_zero) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1866 |
apply (erule (1) arcsin_lbound) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1867 |
apply (erule (1) arcsin_ubound) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1868 |
apply simp |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1869 |
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
|
1870 |
apply (rule power_mono, simp, simp) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1871 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1872 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1873 |
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
|
1874 |
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
|
1875 |
apply (rule power2_eq_imp_eq) |
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1876 |
apply (simp add: sin_squared_eq) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1877 |
apply (rule sin_ge_zero) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1878 |
apply (erule (1) arccos_lbound) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1879 |
apply (erule (1) arccos_ubound) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1880 |
apply simp |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1881 |
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
|
1882 |
apply (rule power_mono, simp, simp) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1883 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1884 |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1885 |
lemma arctan [simp]: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1886 |
"- (pi/2) < arctan y & arctan y < pi/2 & tan (arctan y) = y" |
23011 | 1887 |
unfolding arctan_def by (rule theI' [OF tan_total]) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1888 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1889 |
lemma tan_arctan: "tan(arctan y) = y" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1890 |
by auto |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1891 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1892 |
lemma arctan_bounded: "- (pi/2) < arctan y & arctan y < pi/2" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1893 |
by (auto simp only: arctan) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1894 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1895 |
lemma arctan_lbound: "- (pi/2) < arctan y" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1896 |
by auto |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1897 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1898 |
lemma arctan_ubound: "arctan y < pi/2" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1899 |
by (auto simp only: arctan) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1900 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1901 |
lemma arctan_tan: |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1902 |
"[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1903 |
apply (unfold arctan_def) |
23011 | 1904 |
apply (rule the1_equality) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1905 |
apply (rule tan_total, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1906 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1907 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1908 |
lemma arctan_zero_zero [simp]: "arctan 0 = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1909 |
by (insert arctan_tan [of 0], simp) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1910 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1911 |
lemma cos_arctan_not_zero [simp]: "cos(arctan x) \<noteq> 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1912 |
apply (auto simp add: cos_zero_iff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1913 |
apply (case_tac "n") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1914 |
apply (case_tac [3] "n") |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1915 |
apply (cut_tac [2] y = x in arctan_ubound) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1916 |
apply (cut_tac [4] y = x in arctan_lbound) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1917 |
apply (auto simp add: real_of_nat_Suc left_distrib mult_less_0_iff) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1918 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1919 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1920 |
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
|
1921 |
apply (rule power_inverse [THEN subst]) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1922 |
apply (rule_tac c1 = "(cos x)\<twosuperior>" in real_mult_right_cancel [THEN iffD1]) |
22960 | 1923 |
apply (auto dest: field_power_not_zero |
20516 | 1924 |
simp add: power_mult_distrib left_distrib power_divide tan_def |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1925 |
mult_assoc power_inverse [symmetric] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1926 |
simp del: realpow_Suc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1927 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
1928 |
|
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1929 |
lemma isCont_inverse_function2: |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1930 |
fixes f g :: "real \<Rightarrow> real" shows |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1931 |
"\<lbrakk>a < x; x < b; |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1932 |
\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> g (f z) = z; |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1933 |
\<forall>z. a \<le> z \<and> z \<le> b \<longrightarrow> isCont f z\<rbrakk> |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1934 |
\<Longrightarrow> isCont g (f x)" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1935 |
apply (rule isCont_inverse_function |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1936 |
[where f=f and d="min (x - a) (b - x)"]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1937 |
apply (simp_all add: abs_le_iff) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1938 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1939 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1940 |
lemma isCont_arcsin: "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> isCont arcsin x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1941 |
apply (subgoal_tac "isCont arcsin (sin (arcsin x))", simp) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1942 |
apply (rule isCont_inverse_function2 [where f=sin]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1943 |
apply (erule (1) arcsin_lt_bounded [THEN conjunct1]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1944 |
apply (erule (1) arcsin_lt_bounded [THEN conjunct2]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1945 |
apply (fast intro: arcsin_sin, simp) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1946 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1947 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1948 |
lemma isCont_arccos: "\<lbrakk>-1 < x; x < 1\<rbrakk> \<Longrightarrow> isCont arccos x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1949 |
apply (subgoal_tac "isCont arccos (cos (arccos x))", simp) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1950 |
apply (rule isCont_inverse_function2 [where f=cos]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1951 |
apply (erule (1) arccos_lt_bounded [THEN conjunct1]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1952 |
apply (erule (1) arccos_lt_bounded [THEN conjunct2]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1953 |
apply (fast intro: arccos_cos, simp) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1954 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1955 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1956 |
lemma isCont_arctan: "isCont arctan x" |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1957 |
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
|
1958 |
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
|
1959 |
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
|
1960 |
apply (erule (1) isCont_inverse_function2 [where f=tan]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1961 |
apply (clarify, rule arctan_tan) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1962 |
apply (erule (1) order_less_le_trans) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1963 |
apply (erule (1) order_le_less_trans) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1964 |
apply (clarify, rule isCont_tan) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1965 |
apply (rule less_imp_neq [symmetric]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1966 |
apply (rule cos_gt_zero_pi) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1967 |
apply (erule (1) order_less_le_trans) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1968 |
apply (erule (1) order_le_less_trans) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1969 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1970 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1971 |
lemma DERIV_arcsin: |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1972 |
"\<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
|
1973 |
apply (rule DERIV_inverse_function [where f=sin and a="-1" and b="1"]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1974 |
apply (rule lemma_DERIV_subst [OF DERIV_sin]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1975 |
apply (simp add: cos_arcsin) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1976 |
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
|
1977 |
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
|
1978 |
apply assumption |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1979 |
apply assumption |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1980 |
apply simp |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1981 |
apply (erule (1) isCont_arcsin) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1982 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1983 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1984 |
lemma DERIV_arccos: |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1985 |
"\<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
|
1986 |
apply (rule DERIV_inverse_function [where f=cos and a="-1" and b="1"]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1987 |
apply (rule lemma_DERIV_subst [OF DERIV_cos]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1988 |
apply (simp add: sin_arccos) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1989 |
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
|
1990 |
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
|
1991 |
apply assumption |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1992 |
apply assumption |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1993 |
apply simp |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1994 |
apply (erule (1) isCont_arccos) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1995 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1996 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1997 |
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
|
1998 |
apply (rule DERIV_inverse_function [where f=tan and a="x - 1" and b="x + 1"]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
1999 |
apply (rule lemma_DERIV_subst [OF DERIV_tan]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2000 |
apply (rule cos_arctan_not_zero) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2001 |
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
|
2002 |
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
|
2003 |
apply (simp add: add_pos_nonneg) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2004 |
apply (simp, simp, simp, rule isCont_arctan) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2005 |
done |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2006 |
|
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2007 |
|
23043 | 2008 |
subsection {* More Theorems about Sin and Cos *} |
2009 |
||
23052
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2010 |
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
|
2011 |
proof - |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2012 |
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
|
2013 |
have nonneg: "0 \<le> ?c" |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2014 |
by (rule cos_ge_zero, rule order_trans [where y=0], simp_all) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2015 |
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
|
2016 |
by simp |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2017 |
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
|
2018 |
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
|
2019 |
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
|
2020 |
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
|
2021 |
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
|
2022 |
by (simp add: power_divide) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2023 |
thus ?thesis |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2024 |
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
|
2025 |
qed |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2026 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2027 |
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
|
2028 |
proof - |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2029 |
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
|
2030 |
have pos_c: "0 < ?c" |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2031 |
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
|
2032 |
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
|
2033 |
by simp |
23052
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2034 |
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
|
2035 |
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
|
2036 |
also have "\<dots> = ?c * (?c\<twosuperior> - 3 * ?s\<twosuperior>)" |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2037 |
by (simp add: ring_eq_simps power2_eq_square) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2038 |
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
|
2039 |
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
|
2040 |
thus ?thesis |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2041 |
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
|
2042 |
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
|
2043 |
qed |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2044 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2045 |
lemma sin_45: "sin (pi / 4) = sqrt 2 / 2" |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2046 |
proof - |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2047 |
have "sin (pi / 4) = cos (pi / 2 - pi / 4)" by (rule sin_cos_eq) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2048 |
also have "pi / 2 - pi / 4 = pi / 4" by simp |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2049 |
also have "cos (pi / 4) = sqrt 2 / 2" by (rule cos_45) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2050 |
finally show ?thesis . |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2051 |
qed |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2052 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2053 |
lemma sin_60: "sin (pi / 3) = sqrt 3 / 2" |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2054 |
proof - |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2055 |
have "sin (pi / 3) = cos (pi / 2 - pi / 3)" by (rule sin_cos_eq) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2056 |
also have "pi / 2 - pi / 3 = pi / 6" by simp |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2057 |
also have "cos (pi / 6) = sqrt 3 / 2" by (rule cos_30) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2058 |
finally show ?thesis . |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2059 |
qed |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2060 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2061 |
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
|
2062 |
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
|
2063 |
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
|
2064 |
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
|
2065 |
done |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2066 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2067 |
lemma sin_30: "sin (pi / 6) = 1 / 2" |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2068 |
proof - |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2069 |
have "sin (pi / 6) = cos (pi / 2 - pi / 6)" by (rule sin_cos_eq) |
23066
26a9157b620a
new field_combine_numerals simproc, which uses fractions as coefficients
huffman
parents:
23053
diff
changeset
|
2070 |
also have "pi / 2 - pi / 6 = pi / 3" by simp |
23052
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2071 |
also have "cos (pi / 3) = 1 / 2" by (rule cos_60) |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2072 |
finally show ?thesis . |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2073 |
qed |
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2074 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2075 |
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
|
2076 |
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
|
2077 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2078 |
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
|
2079 |
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
|
2080 |
|
0e36f0dbfa1c
add lemmas for sin,cos,tan of 30,45,60 degrees; cleaned up
huffman
parents:
23049
diff
changeset
|
2081 |
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
|
2082 |
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
|
2083 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
2084 |
text{*NEEDED??*} |
15229 | 2085 |
lemma [simp]: |
2086 |
"sin (x + 1 / 2 * real (Suc m) * pi) = |
|
2087 |
cos (x + 1 / 2 * real (m) * pi)" |
|
2088 |
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib, auto) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2089 |
|
15085
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
paulson
parents:
15081
diff
changeset
|
2090 |
text{*NEEDED??*} |
15229 | 2091 |
lemma [simp]: |
2092 |
"sin (x + real (Suc m) * pi / 2) = |
|
2093 |
cos (x + real (m) * pi / 2)" |
|
2094 |
by (simp only: cos_add sin_add real_of_nat_Suc add_divide_distrib left_distrib, auto) |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2095 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2096 |
lemma DERIV_sin_add [simp]: "DERIV (%x. sin (x + k)) xa :> cos (xa + k)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2097 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2098 |
apply (rule_tac f = sin and g = "%x. x + k" in DERIV_chain2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2099 |
apply (best intro!: DERIV_intros intro: DERIV_chain2)+ |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2100 |
apply (simp (no_asm)) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2101 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2102 |
|
15383 | 2103 |
lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi / 2) = (-1) ^ n" |
2104 |
proof - |
|
2105 |
have "sin ((real n + 1/2) * pi) = cos (real n * pi)" |
|
2106 |
by (auto simp add: right_distrib sin_add left_distrib mult_ac) |
|
2107 |
thus ?thesis |
|
2108 |
by (simp add: real_of_nat_Suc left_distrib add_divide_distrib |
|
2109 |
mult_commute [of pi]) |
|
2110 |
qed |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2111 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2112 |
lemma cos_2npi [simp]: "cos (2 * real (n::nat) * pi) = 1" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2113 |
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
|
2114 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2115 |
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
|
2116 |
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
|
2117 |
apply (subst cos_add, simp) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2118 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2119 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2120 |
lemma sin_2npi [simp]: "sin (2 * real (n::nat) * pi) = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2121 |
by (auto simp add: mult_assoc) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2122 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2123 |
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
|
2124 |
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
|
2125 |
apply (subst sin_add, simp) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2126 |
done |
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 |
(*NEEDED??*) |
15229 | 2129 |
lemma [simp]: |
2130 |
"cos(x + 1 / 2 * real(Suc m) * pi) = -sin (x + 1 / 2 * real m * pi)" |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2131 |
apply (simp only: cos_add sin_add real_of_nat_Suc right_distrib left_distrib minus_mult_right, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2132 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2133 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2134 |
(*NEEDED??*) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2135 |
lemma [simp]: "cos (x + real(Suc m) * pi / 2) = -sin (x + real m * pi / 2)" |
15229 | 2136 |
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib add_divide_distrib, auto) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2137 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2138 |
lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" |
15229 | 2139 |
by (simp only: cos_add sin_add real_of_nat_Suc left_distrib right_distrib add_divide_distrib, auto) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2140 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2141 |
lemma DERIV_cos_add [simp]: "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2142 |
apply (rule lemma_DERIV_subst) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2143 |
apply (rule_tac f = cos and g = "%x. x + k" in DERIV_chain2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2144 |
apply (best intro!: DERIV_intros intro: DERIV_chain2)+ |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2145 |
apply (simp (no_asm)) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2146 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2147 |
|
15081 | 2148 |
lemma sin_zero_abs_cos_one: "sin x = 0 ==> \<bar>cos x\<bar> = 1" |
15539 | 2149 |
by (auto simp add: sin_zero_iff even_mult_two_ex) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2150 |
|
23115
4615b2078592
generalized exp to work over any complete field; new proof of exp_add
huffman
parents:
23112
diff
changeset
|
2151 |
lemma exp_eq_one_iff [simp]: "(exp (x::real) = 1) = (x = 0)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2152 |
apply auto |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2153 |
apply (drule_tac f = ln in arg_cong, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2154 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2155 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2156 |
lemma cos_one_sin_zero: "cos x = 1 ==> sin x = 0" |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2157 |
by (cut_tac x = x in sin_cos_squared_add3, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2158 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2159 |
|
22978
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2160 |
subsection {* Existence of Polar Coordinates *} |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2161 |
|
22978
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2162 |
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
|
2163 |
apply (rule power2_le_imp_le [OF _ zero_le_one]) |
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2164 |
apply (simp add: abs_divide power_divide divide_le_eq not_sum_power2_lt_zero) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2165 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2166 |
|
22978
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2167 |
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
|
2168 |
by (simp add: abs_le_iff) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2169 |
|
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2170 |
lemma 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
|
2171 |
by (simp add: sin_arccos abs_le_iff) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2172 |
|
22978
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2173 |
lemmas cos_arccos_lemma1 = cos_arccos_abs [OF cos_x_y_le_one] |
15228 | 2174 |
|
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2175 |
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
|
2176 |
|
15229 | 2177 |
lemma polar_ex1: |
22978
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2178 |
"0 < y ==> \<exists>r a. x = r * cos a & y = r * sin a" |
15229 | 2179 |
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
|
2180 |
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
|
2181 |
apply (simp add: cos_arccos_lemma1) |
23045
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2182 |
apply (simp add: sin_arccos_lemma1) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2183 |
apply (simp add: power_divide) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2184 |
apply (simp add: real_sqrt_mult [symmetric]) |
95e04f335940
add lemmas about inverse functions; cleaned up proof of polar_ex
huffman
parents:
23043
diff
changeset
|
2185 |
apply (simp add: right_diff_distrib) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2186 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2187 |
|
15229 | 2188 |
lemma polar_ex2: |
22978
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2189 |
"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
|
2190 |
apply (insert polar_ex1 [where x=x and y="-y"], simp, clarify) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2191 |
apply (rule_tac x = r in exI) |
22978
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2192 |
apply (rule_tac x = "-a" in exI, simp) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2193 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2194 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2195 |
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
|
2196 |
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
|
2197 |
apply (erule polar_ex1) |
1cd8cc21a7c3
clean up polar_Ex proofs; remove unnecessary lemmas
huffman
parents:
22977
diff
changeset
|
2198 |
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
|
2199 |
apply (erule polar_ex2) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2200 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2201 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2202 |
|
23043 | 2203 |
subsection {* Theorems about Limits *} |
2204 |
||
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2205 |
(* need to rename second isCont_inverse *) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2206 |
|
15229 | 2207 |
lemma isCont_inv_fun: |
20561
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
huffman
parents:
20552
diff
changeset
|
2208 |
fixes f g :: "real \<Rightarrow> real" |
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
huffman
parents:
20552
diff
changeset
|
2209 |
shows "[| 0 < d; \<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z; |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2210 |
\<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2211 |
==> isCont g (f x)" |
22722
704de05715b4
lemma isCont_inv_fun is same as isCont_inverse_function
huffman
parents:
22721
diff
changeset
|
2212 |
by (rule isCont_inverse_function) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2213 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2214 |
lemma isCont_inv_fun_inv: |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20516
diff
changeset
|
2215 |
fixes f g :: "real \<Rightarrow> real" |
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20516
diff
changeset
|
2216 |
shows "[| 0 < d; |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2217 |
\<forall>z. \<bar>z - x\<bar> \<le> d --> g(f(z)) = z; |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2218 |
\<forall>z. \<bar>z - x\<bar> \<le> d --> isCont f z |] |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2219 |
==> \<exists>e. 0 < e & |
15081 | 2220 |
(\<forall>y. 0 < \<bar>y - f(x)\<bar> & \<bar>y - f(x)\<bar> < e --> f(g(y)) = y)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2221 |
apply (drule isCont_inj_range) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2222 |
prefer 2 apply (assumption, assumption, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2223 |
apply (rule_tac x = e in exI, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2224 |
apply (rotate_tac 2) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2225 |
apply (drule_tac x = y in spec, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2226 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2227 |
|
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 |
text{*Bartle/Sherbert: Introduction to Real Analysis, Theorem 4.2.9, p. 110*} |
15229 | 2230 |
lemma LIM_fun_gt_zero: |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20516
diff
changeset
|
2231 |
"[| f -- c --> (l::real); 0 < l |] |
20561
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
huffman
parents:
20552
diff
changeset
|
2232 |
==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> 0 < f x)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2233 |
apply (auto simp add: LIM_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2234 |
apply (drule_tac x = "l/2" in spec, safe, force) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2235 |
apply (rule_tac x = s in exI) |
22998 | 2236 |
apply (auto simp only: abs_less_iff) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2237 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2238 |
|
15229 | 2239 |
lemma LIM_fun_less_zero: |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20516
diff
changeset
|
2240 |
"[| f -- c --> (l::real); l < 0 |] |
20561
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
huffman
parents:
20552
diff
changeset
|
2241 |
==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x < 0)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2242 |
apply (auto simp add: LIM_def) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2243 |
apply (drule_tac x = "-l/2" in spec, safe, force) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2244 |
apply (rule_tac x = s in exI) |
22998 | 2245 |
apply (auto simp only: abs_less_iff) |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2246 |
done |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2247 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2248 |
|
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2249 |
lemma LIM_fun_not_zero: |
20552
2c31dd358c21
generalized types of many constants to work over arbitrary vector spaces;
huffman
parents:
20516
diff
changeset
|
2250 |
"[| f -- c --> (l::real); l \<noteq> 0 |] |
20561
6a6d8004322f
generalize type of (NS)LIM to work on functions with vector space domain types
huffman
parents:
20552
diff
changeset
|
2251 |
==> \<exists>r. 0 < r & (\<forall>x::real. x \<noteq> c & \<bar>c - x\<bar> < r --> f x \<noteq> 0)" |
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2252 |
apply (cut_tac x = l and y = 0 in linorder_less_linear, auto) |
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2253 |
apply (drule LIM_fun_less_zero) |
15241 | 2254 |
apply (drule_tac [3] LIM_fun_gt_zero) |
2255 |
apply force+ |
|
15077
89840837108e
converting Hyperreal/Transcendental to Isar script
paulson
parents:
15013
diff
changeset
|
2256 |
done |
20432
07ec57376051
lin_arith_prover: splitting reverted because of performance loss
webertj
parents:
20256
diff
changeset
|
2257 |
|
12196 | 2258 |
end |