68630
|
1 |
(*
|
|
2 |
File: Multiseries_Expansion.thy
|
|
3 |
Author: Manuel Eberl, TU München
|
|
4 |
|
|
5 |
Asymptotic bases and Multiseries expansions of real-valued functions.
|
|
6 |
Contains automation to prove limits and asymptotic relationships between such functions.
|
|
7 |
*)
|
|
8 |
section \<open>Multiseries expansions\<close>
|
|
9 |
theory Multiseries_Expansion
|
|
10 |
imports
|
|
11 |
"HOL-Library.BNF_Corec"
|
|
12 |
"HOL-Library.Landau_Symbols"
|
|
13 |
Lazy_Eval
|
|
14 |
Inst_Existentials
|
|
15 |
Eventuallize
|
|
16 |
begin
|
|
17 |
|
|
18 |
(* TODO Move *)
|
|
19 |
lemma real_powr_at_top:
|
|
20 |
assumes "(p::real) > 0"
|
|
21 |
shows "filterlim (\<lambda>x. x powr p) at_top at_top"
|
|
22 |
proof (subst filterlim_cong[OF refl refl])
|
|
23 |
show "LIM x at_top. exp (p * ln x) :> at_top"
|
|
24 |
by (rule filterlim_compose[OF exp_at_top filterlim_tendsto_pos_mult_at_top[OF tendsto_const]])
|
|
25 |
(simp_all add: ln_at_top assms)
|
|
26 |
show "eventually (\<lambda>x. x powr p = exp (p * ln x)) at_top"
|
|
27 |
using eventually_gt_at_top[of 0] by eventually_elim (simp add: powr_def)
|
|
28 |
qed
|
|
29 |
|
|
30 |
|
|
31 |
subsection \<open>Streams and lazy lists\<close>
|
|
32 |
|
|
33 |
codatatype 'a msstream = MSSCons 'a "'a msstream"
|
|
34 |
|
|
35 |
primrec mssnth :: "'a msstream \<Rightarrow> nat \<Rightarrow> 'a" where
|
|
36 |
"mssnth xs 0 = (case xs of MSSCons x xs \<Rightarrow> x)"
|
|
37 |
| "mssnth xs (Suc n) = (case xs of MSSCons x xs \<Rightarrow> mssnth xs n)"
|
|
38 |
|
|
39 |
codatatype 'a msllist = MSLNil | MSLCons 'a "'a msllist"
|
|
40 |
for map: mslmap
|
|
41 |
|
|
42 |
fun lcoeff where
|
|
43 |
"lcoeff MSLNil n = 0"
|
|
44 |
| "lcoeff (MSLCons x xs) 0 = x"
|
|
45 |
| "lcoeff (MSLCons x xs) (Suc n) = lcoeff xs n"
|
|
46 |
|
|
47 |
primcorec msllist_of_msstream :: "'a msstream \<Rightarrow> 'a msllist" where
|
|
48 |
"msllist_of_msstream xs = (case xs of MSSCons x xs \<Rightarrow> MSLCons x (msllist_of_msstream xs))"
|
|
49 |
|
|
50 |
lemma msllist_of_msstream_MSSCons [simp]:
|
|
51 |
"msllist_of_msstream (MSSCons x xs) = MSLCons x (msllist_of_msstream xs)"
|
|
52 |
by (subst msllist_of_msstream.code) simp
|
|
53 |
|
|
54 |
lemma lcoeff_llist_of_stream [simp]: "lcoeff (msllist_of_msstream xs) n = mssnth xs n"
|
|
55 |
by (induction "msllist_of_msstream xs" n arbitrary: xs rule: lcoeff.induct;
|
|
56 |
subst msllist_of_msstream.code) (auto split: msstream.splits)
|
|
57 |
|
|
58 |
|
|
59 |
subsection \<open>Convergent power series\<close>
|
|
60 |
|
|
61 |
subsubsection \<open>Definition\<close>
|
|
62 |
|
|
63 |
definition convergent_powser :: "real msllist \<Rightarrow> bool" where
|
|
64 |
"convergent_powser cs \<longleftrightarrow> (\<exists>R>0. \<forall>x. abs x < R \<longrightarrow> summable (\<lambda>n. lcoeff cs n * x ^ n))"
|
|
65 |
|
|
66 |
lemma convergent_powser_stl:
|
|
67 |
assumes "convergent_powser (MSLCons c cs)"
|
|
68 |
shows "convergent_powser cs"
|
|
69 |
proof -
|
|
70 |
from assms obtain R where "R > 0" "\<And>x. abs x < R \<Longrightarrow> summable (\<lambda>n. lcoeff (MSLCons c cs) n * x ^ n)"
|
|
71 |
unfolding convergent_powser_def by blast
|
|
72 |
hence "summable (\<lambda>n. lcoeff cs n * x ^ n)" if "abs x < R" for x
|
|
73 |
using that by (subst (asm) summable_powser_split_head [symmetric]) simp
|
|
74 |
thus ?thesis using \<open>R > 0\<close> by (auto simp: convergent_powser_def)
|
|
75 |
qed
|
|
76 |
|
|
77 |
|
|
78 |
definition powser :: "real msllist \<Rightarrow> real \<Rightarrow> real" where
|
|
79 |
"powser cs x = suminf (\<lambda>n. lcoeff cs n * x ^ n)"
|
|
80 |
|
|
81 |
lemma isCont_powser:
|
|
82 |
assumes "convergent_powser cs"
|
|
83 |
shows "isCont (powser cs) 0"
|
|
84 |
proof -
|
|
85 |
from assms obtain R where R: "R > 0" "\<And>x. abs x < R \<Longrightarrow> summable (\<lambda>n. lcoeff cs n * x ^ n)"
|
|
86 |
unfolding convergent_powser_def by blast
|
|
87 |
hence *: "summable (\<lambda>n. lcoeff cs n * (R/2) ^ n)" by (intro R) simp_all
|
|
88 |
from \<open>R > 0\<close> show ?thesis unfolding powser_def
|
|
89 |
by (intro isCont_powser[OF *]) simp_all
|
|
90 |
qed
|
|
91 |
|
|
92 |
lemma powser_MSLNil [simp]: "powser MSLNil = (\<lambda>_. 0)"
|
|
93 |
by (simp add: fun_eq_iff powser_def)
|
|
94 |
|
|
95 |
lemma powser_MSLCons:
|
|
96 |
assumes "convergent_powser (MSLCons c cs)"
|
|
97 |
shows "eventually (\<lambda>x. powser (MSLCons c cs) x = x * powser cs x + c) (nhds 0)"
|
|
98 |
proof -
|
|
99 |
from assms obtain R where R: "R > 0" "\<And>x. abs x < R \<Longrightarrow> summable (\<lambda>n. lcoeff (MSLCons c cs) n * x ^ n)"
|
|
100 |
unfolding convergent_powser_def by blast
|
|
101 |
from R have "powser (MSLCons c cs) x = x * powser cs x + c" if "abs x < R" for x
|
|
102 |
using that unfolding powser_def by (subst powser_split_head) simp_all
|
|
103 |
moreover have "eventually (\<lambda>x. abs x < R) (nhds 0)"
|
|
104 |
using \<open>R > 0\<close> filterlim_ident[of "nhds (0::real)"]
|
|
105 |
by (subst (asm) tendsto_iff) (simp add: dist_real_def)
|
|
106 |
ultimately show ?thesis by (auto elim: eventually_mono)
|
|
107 |
qed
|
|
108 |
|
|
109 |
definition convergent_powser' :: "real msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool" where
|
|
110 |
"convergent_powser' cs f \<longleftrightarrow> (\<exists>R>0. \<forall>x. abs x < R \<longrightarrow> (\<lambda>n. lcoeff cs n * x ^ n) sums f x)"
|
|
111 |
|
|
112 |
lemma convergent_powser'_imp_convergent_powser:
|
|
113 |
"convergent_powser' cs f \<Longrightarrow> convergent_powser cs"
|
|
114 |
unfolding convergent_powser_def convergent_powser'_def by (auto simp: sums_iff)
|
|
115 |
|
|
116 |
lemma convergent_powser'_eventually:
|
|
117 |
assumes "convergent_powser' cs f"
|
|
118 |
shows "eventually (\<lambda>x. powser cs x = f x) (nhds 0)"
|
|
119 |
proof -
|
|
120 |
from assms obtain R where "R > 0" "\<And>x. abs x < R \<Longrightarrow> (\<lambda>n. lcoeff cs n * x ^ n) sums f x"
|
|
121 |
unfolding convergent_powser'_def by blast
|
|
122 |
hence "powser cs x = f x" if "abs x < R" for x
|
|
123 |
using that by (simp add: powser_def sums_iff)
|
|
124 |
moreover have "eventually (\<lambda>x. abs x < R) (nhds 0)"
|
|
125 |
using \<open>R > 0\<close> filterlim_ident[of "nhds (0::real)"]
|
|
126 |
by (subst (asm) tendsto_iff) (simp add: dist_real_def)
|
|
127 |
ultimately show ?thesis by (auto elim: eventually_mono)
|
|
128 |
qed
|
|
129 |
|
|
130 |
|
|
131 |
subsubsection \<open>Geometric series\<close>
|
|
132 |
|
|
133 |
primcorec mssalternate :: "'a \<Rightarrow> 'a \<Rightarrow> 'a msstream" where
|
|
134 |
"mssalternate a b = MSSCons a (mssalternate b a)"
|
|
135 |
|
|
136 |
lemma case_mssalternate [simp]:
|
|
137 |
"(case mssalternate a b of MSSCons c d \<Rightarrow> f c d) = f a (mssalternate b a)"
|
|
138 |
by (subst mssalternate.code) simp
|
|
139 |
|
|
140 |
lemma mssnth_mssalternate: "mssnth (mssalternate a b) n = (if even n then a else b)"
|
|
141 |
by (induction n arbitrary: a b; subst mssalternate.code) simp_all
|
|
142 |
|
|
143 |
lemma convergent_powser'_geometric:
|
|
144 |
"convergent_powser' (msllist_of_msstream (mssalternate 1 (-1))) (\<lambda>x. inverse (1 + x))"
|
|
145 |
proof -
|
|
146 |
have "(\<lambda>n. lcoeff (msllist_of_msstream (mssalternate 1 (-1))) n * x ^ n) sums (inverse (1 + x))"
|
|
147 |
if "abs x < 1" for x :: real
|
|
148 |
proof -
|
|
149 |
have "(\<lambda>n. (-1) ^ n * x ^ n) sums (inverse (1 + x))"
|
|
150 |
using that geometric_sums[of "-x"] by (simp add: field_simps power_mult_distrib [symmetric])
|
|
151 |
also have "(\<lambda>n. (-1) ^ n * x ^ n) =
|
|
152 |
(\<lambda>n. lcoeff (msllist_of_msstream (mssalternate 1 (-1))) n * x ^ n)"
|
|
153 |
by (auto simp add: mssnth_mssalternate fun_eq_iff)
|
|
154 |
finally show ?thesis .
|
|
155 |
qed
|
|
156 |
thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1])
|
|
157 |
qed
|
|
158 |
|
|
159 |
|
|
160 |
subsubsection \<open>Exponential series\<close>
|
|
161 |
|
|
162 |
primcorec exp_series_stream_aux :: "real \<Rightarrow> real \<Rightarrow> real msstream" where
|
|
163 |
"exp_series_stream_aux acc n = MSSCons acc (exp_series_stream_aux (acc / n) (n + 1))"
|
|
164 |
|
|
165 |
lemma mssnth_exp_series_stream_aux:
|
|
166 |
"mssnth (exp_series_stream_aux (1 / fact n) (n + 1)) m = 1 / fact (n + m)"
|
|
167 |
proof (induction m arbitrary: n)
|
|
168 |
case (0 n)
|
|
169 |
thus ?case by (subst exp_series_stream_aux.code) simp
|
|
170 |
next
|
|
171 |
case (Suc m n)
|
|
172 |
from Suc.IH[of "n + 1"] show ?case
|
|
173 |
by (subst exp_series_stream_aux.code) (simp add: algebra_simps)
|
|
174 |
qed
|
|
175 |
|
|
176 |
definition exp_series_stream :: "real msstream" where
|
|
177 |
"exp_series_stream = exp_series_stream_aux 1 1"
|
|
178 |
|
|
179 |
lemma mssnth_exp_series_stream:
|
|
180 |
"mssnth exp_series_stream n = 1 / fact n"
|
|
181 |
unfolding exp_series_stream_def using mssnth_exp_series_stream_aux[of 0 n] by simp
|
|
182 |
|
|
183 |
lemma convergent_powser'_exp:
|
|
184 |
"convergent_powser' (msllist_of_msstream exp_series_stream) exp"
|
|
185 |
proof -
|
|
186 |
have "(\<lambda>n. lcoeff (msllist_of_msstream exp_series_stream) n * x ^ n) sums exp x" for x :: real
|
|
187 |
using exp_converges[of x] by (simp add: mssnth_exp_series_stream divide_simps)
|
|
188 |
thus ?thesis by (auto intro: exI[of _ "1::real"] simp: convergent_powser'_def)
|
|
189 |
qed
|
|
190 |
|
|
191 |
|
|
192 |
subsubsection \<open>Logarithm series\<close>
|
|
193 |
|
|
194 |
primcorec ln_series_stream_aux :: "bool \<Rightarrow> real \<Rightarrow> real msstream" where
|
|
195 |
"ln_series_stream_aux b n =
|
|
196 |
MSSCons (if b then -inverse n else inverse n) (ln_series_stream_aux (\<not>b) (n+1))"
|
|
197 |
|
|
198 |
lemma mssnth_ln_series_stream_aux:
|
|
199 |
"mssnth (ln_series_stream_aux b x) n =
|
|
200 |
(if b then - ((-1)^n) * inverse (x + real n) else ((-1)^n) * inverse (x + real n))"
|
|
201 |
by (induction n arbitrary: b x; subst ln_series_stream_aux.code) simp_all
|
|
202 |
|
|
203 |
definition ln_series_stream :: "real msstream" where
|
|
204 |
"ln_series_stream = MSSCons 0 (ln_series_stream_aux False 1)"
|
|
205 |
|
|
206 |
lemma mssnth_ln_series_stream:
|
|
207 |
"mssnth ln_series_stream n = (-1) ^ Suc n / real n"
|
|
208 |
unfolding ln_series_stream_def
|
|
209 |
by (cases n) (simp_all add: mssnth_ln_series_stream_aux field_simps)
|
|
210 |
|
|
211 |
lemma ln_series':
|
|
212 |
assumes "abs (x::real) < 1"
|
|
213 |
shows "(\<lambda>n. - ((-x)^n) / of_nat n) sums ln (1 + x)"
|
|
214 |
proof -
|
|
215 |
have "\<forall>n\<ge>1. norm (-((-x)^n)) / of_nat n \<le> norm x ^ n / 1"
|
|
216 |
by (intro exI[of _ "1::nat"] allI impI frac_le) (simp_all add: power_abs)
|
|
217 |
hence "\<exists>N. \<forall>n\<ge>N. norm (-((-x)^n) / of_nat n) \<le> norm x ^ n"
|
|
218 |
by (intro exI[of _ 1]) simp_all
|
|
219 |
moreover from assms have "summable (\<lambda>n. norm x ^ n)"
|
|
220 |
by (intro summable_geometric) simp_all
|
|
221 |
ultimately have *: "summable (\<lambda>n. - ((-x)^n) / of_nat n)"
|
|
222 |
by (rule summable_comparison_test)
|
|
223 |
moreover from assms have "0 < 1 + x" "1 + x < 2" by simp_all
|
|
224 |
from ln_series[OF this]
|
|
225 |
have "ln (1 + x) = (\<Sum>n. - ((-x) ^ Suc n) / real (Suc n))"
|
|
226 |
by (simp_all add: power_minus' mult_ac)
|
|
227 |
hence "ln (1 + x) = (\<Sum>n. - ((-x) ^ n / real n))"
|
|
228 |
by (subst (asm) suminf_split_head[OF *]) simp_all
|
|
229 |
ultimately show ?thesis by (simp add: sums_iff)
|
|
230 |
qed
|
|
231 |
|
|
232 |
lemma convergent_powser'_ln:
|
|
233 |
"convergent_powser' (msllist_of_msstream ln_series_stream) (\<lambda>x. ln (1 + x))"
|
|
234 |
proof -
|
|
235 |
have "(\<lambda>n. lcoeff (msllist_of_msstream ln_series_stream)n * x ^ n) sums ln (1 + x)"
|
|
236 |
if "abs x < 1" for x
|
|
237 |
proof -
|
|
238 |
from that have "(\<lambda>n. - ((- x) ^ n) / real n) sums ln (1 + x)" by (rule ln_series')
|
|
239 |
also have "(\<lambda>n. - ((- x) ^ n) / real n) =
|
|
240 |
(\<lambda>n. lcoeff (msllist_of_msstream ln_series_stream) n * x ^ n)"
|
|
241 |
by (auto simp: fun_eq_iff mssnth_ln_series_stream power_mult_distrib [symmetric])
|
|
242 |
finally show ?thesis .
|
|
243 |
qed
|
|
244 |
thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1])
|
|
245 |
qed
|
|
246 |
|
|
247 |
|
|
248 |
subsubsection \<open>Generalized binomial series\<close>
|
|
249 |
|
|
250 |
primcorec gbinomial_series_aux :: "bool \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real msllist" where
|
|
251 |
"gbinomial_series_aux abort x n acc =
|
|
252 |
(if abort \<and> acc = 0 then MSLNil else
|
|
253 |
MSLCons acc (gbinomial_series_aux abort x (n + 1) ((x - n) / (n + 1) * acc)))"
|
|
254 |
|
|
255 |
lemma gbinomial_series_abort [simp]: "gbinomial_series_aux True x n 0 = MSLNil"
|
|
256 |
by (subst gbinomial_series_aux.code) simp
|
|
257 |
|
|
258 |
definition gbinomial_series :: "bool \<Rightarrow> real \<Rightarrow> real msllist" where
|
|
259 |
"gbinomial_series abort x = gbinomial_series_aux abort x 0 1"
|
|
260 |
|
|
261 |
|
|
262 |
(* TODO Move *)
|
|
263 |
lemma gbinomial_Suc_rec:
|
|
264 |
"(x gchoose (Suc k)) = (x gchoose k) * ((x - of_nat k) / (of_nat k + 1))"
|
|
265 |
proof -
|
|
266 |
have "((x - 1) + 1) gchoose (Suc k) = x * (x - 1 gchoose k) / (1 + of_nat k)"
|
|
267 |
by (subst gbinomial_factors) simp
|
|
268 |
also have "x * (x - 1 gchoose k) = (x - of_nat k) * (x gchoose k)"
|
|
269 |
by (rule gbinomial_absorb_comp [symmetric])
|
|
270 |
finally show ?thesis by (simp add: algebra_simps)
|
|
271 |
qed
|
|
272 |
|
|
273 |
lemma lcoeff_gbinomial_series_aux:
|
|
274 |
"lcoeff (gbinomial_series_aux abort x m (x gchoose m)) n = (x gchoose (n + m))"
|
|
275 |
proof (induction n arbitrary: m)
|
|
276 |
case 0
|
|
277 |
show ?case by (subst gbinomial_series_aux.code) simp
|
|
278 |
next
|
|
279 |
case (Suc n m)
|
|
280 |
show ?case
|
|
281 |
proof (cases "abort \<and> (x gchoose m) = 0")
|
|
282 |
case True
|
|
283 |
with gbinomial_mult_fact[of m x] obtain k where "x = real k" "k < m"
|
|
284 |
by auto
|
|
285 |
hence "(x gchoose Suc (n + m)) = 0"
|
|
286 |
using gbinomial_mult_fact[of "Suc (n + m)" x]
|
|
287 |
by (simp add: gbinomial_altdef_of_nat)
|
|
288 |
with True show ?thesis by (subst gbinomial_series_aux.code) simp
|
|
289 |
next
|
|
290 |
case False
|
|
291 |
hence "lcoeff (gbinomial_series_aux abort x m (x gchoose m)) (Suc n) =
|
|
292 |
lcoeff (gbinomial_series_aux abort x (Suc m) ((x gchoose m) *
|
|
293 |
((x - real m) / (real m + 1)))) n"
|
|
294 |
by (subst gbinomial_series_aux.code) (auto simp: algebra_simps)
|
|
295 |
also have "((x gchoose m) * ((x - real m) / (real m + 1))) = x gchoose (Suc m)"
|
|
296 |
by (rule gbinomial_Suc_rec [symmetric])
|
|
297 |
also have "lcoeff (gbinomial_series_aux abort x (Suc m) \<dots>) n = x gchoose (n + Suc m)"
|
|
298 |
by (rule Suc.IH)
|
|
299 |
finally show ?thesis by simp
|
|
300 |
qed
|
|
301 |
qed
|
|
302 |
|
|
303 |
lemma lcoeff_gbinomial_series [simp]:
|
|
304 |
"lcoeff (gbinomial_series abort x) n = (x gchoose n)"
|
|
305 |
using lcoeff_gbinomial_series_aux[of abort x 0 n] by (simp add: gbinomial_series_def)
|
|
306 |
|
|
307 |
lemma gbinomial_ratio_limit:
|
|
308 |
fixes a :: "'a :: real_normed_field"
|
|
309 |
assumes "a \<notin> \<nat>"
|
|
310 |
shows "(\<lambda>n. (a gchoose n) / (a gchoose Suc n)) \<longlonglongrightarrow> -1"
|
|
311 |
proof (rule Lim_transform_eventually)
|
|
312 |
let ?f = "\<lambda>n. inverse (a / of_nat (Suc n) - of_nat n / of_nat (Suc n))"
|
|
313 |
from eventually_gt_at_top[of "0::nat"]
|
|
314 |
show "eventually (\<lambda>n. ?f n = (a gchoose n) /(a gchoose Suc n)) sequentially"
|
|
315 |
proof eventually_elim
|
|
316 |
fix n :: nat assume n: "n > 0"
|
|
317 |
then obtain q where q: "n = Suc q" by (cases n) blast
|
|
318 |
let ?P = "\<Prod>i=0..<n. a - of_nat i"
|
|
319 |
from n have "(a gchoose n) / (a gchoose Suc n) = (of_nat (Suc n) :: 'a) *
|
|
320 |
(?P / (\<Prod>i=0..n. a - of_nat i))"
|
|
321 |
by (simp add: gbinomial_prod_rev atLeastLessThanSuc_atLeastAtMost)
|
|
322 |
also from q have "(\<Prod>i=0..n. a - of_nat i) = ?P * (a - of_nat n)"
|
|
323 |
by (simp add: prod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost)
|
|
324 |
also have "?P / \<dots> = (?P / ?P) / (a - of_nat n)" by (rule divide_divide_eq_left[symmetric])
|
|
325 |
also from assms have "?P / ?P = 1" by auto
|
|
326 |
also have "of_nat (Suc n) * (1 / (a - of_nat n)) =
|
|
327 |
inverse (inverse (of_nat (Suc n)) * (a - of_nat n))" by (simp add: field_simps)
|
|
328 |
also have "inverse (of_nat (Suc n)) * (a - of_nat n) =
|
|
329 |
a / of_nat (Suc n) - of_nat n / of_nat (Suc n)"
|
|
330 |
by (simp add: field_simps del: of_nat_Suc)
|
|
331 |
finally show "?f n = (a gchoose n) / (a gchoose Suc n)" by simp
|
|
332 |
qed
|
|
333 |
|
|
334 |
have "(\<lambda>n. norm a / (of_nat (Suc n))) \<longlonglongrightarrow> 0"
|
|
335 |
unfolding divide_inverse
|
|
336 |
by (intro tendsto_mult_right_zero LIMSEQ_inverse_real_of_nat)
|
|
337 |
hence "(\<lambda>n. a / of_nat (Suc n)) \<longlonglongrightarrow> 0"
|
|
338 |
by (subst tendsto_norm_zero_iff[symmetric]) (simp add: norm_divide del: of_nat_Suc)
|
|
339 |
hence "?f \<longlonglongrightarrow> inverse (0 - 1)"
|
|
340 |
by (intro tendsto_inverse tendsto_diff LIMSEQ_n_over_Suc_n) simp_all
|
|
341 |
thus "?f \<longlonglongrightarrow> -1" by simp
|
|
342 |
qed
|
|
343 |
|
|
344 |
lemma summable_gbinomial:
|
|
345 |
fixes a z :: real
|
|
346 |
assumes "norm z < 1"
|
|
347 |
shows "summable (\<lambda>n. (a gchoose n) * z ^ n)"
|
|
348 |
proof (cases "z = 0 \<or> a \<in> \<nat>")
|
|
349 |
case False
|
|
350 |
define r where "r = (norm z + 1) / 2"
|
|
351 |
from assms have r: "r > norm z" "r < 1" by (simp_all add: r_def field_simps)
|
|
352 |
from False have "abs z > 0" by auto
|
|
353 |
from False have "a \<notin> \<nat>" by (auto elim!: Nats_cases)
|
|
354 |
hence *: "(\<lambda>n. norm (z / ((a gchoose n) / (a gchoose (Suc n))))) \<longlonglongrightarrow> norm (z / (-1))"
|
|
355 |
by (intro tendsto_norm tendsto_divide tendsto_const gbinomial_ratio_limit) simp_all
|
|
356 |
hence "\<forall>\<^sub>F x in at_top. norm (z / ((a gchoose x) / (a gchoose Suc x))) > 0"
|
|
357 |
using assms False by (intro order_tendstoD) auto
|
|
358 |
hence nz: "\<forall>\<^sub>F x in at_top. (a gchoose x) \<noteq> 0" by eventually_elim auto
|
|
359 |
|
|
360 |
have "\<forall>\<^sub>F x in at_top. norm (z / ((a gchoose x) / (a gchoose Suc x))) < r"
|
|
361 |
using assms r by (intro order_tendstoD(2)[OF *]) simp_all
|
|
362 |
with nz have "\<forall>\<^sub>F n in at_top. norm ((a gchoose (Suc n)) * z) \<le> r * norm (a gchoose n)"
|
|
363 |
by eventually_elim (simp add: field_simps abs_mult split: if_splits)
|
|
364 |
hence "\<forall>\<^sub>F n in at_top. norm (z ^ n) * norm ((a gchoose (Suc n)) * z) \<le>
|
|
365 |
norm (z ^ n) * (r * norm (a gchoose n))"
|
|
366 |
by eventually_elim (insert False, intro mult_left_mono, auto)
|
|
367 |
hence "\<forall>\<^sub>F n in at_top. norm ((a gchoose (Suc n)) * z ^ (Suc n)) \<le>
|
|
368 |
r * norm ((a gchoose n) * z ^ n)"
|
|
369 |
by (simp add: abs_mult mult_ac)
|
|
370 |
then obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm ((a gchoose (Suc n)) * z ^ (Suc n)) \<le>
|
|
371 |
r * norm ((a gchoose n) * z ^ n)"
|
|
372 |
unfolding eventually_at_top_linorder by blast
|
|
373 |
show "summable (\<lambda>n. (a gchoose n) * z ^ n)"
|
|
374 |
by (intro summable_ratio_test[OF r(2), of N] N)
|
|
375 |
next
|
|
376 |
case True
|
|
377 |
thus ?thesis
|
|
378 |
proof
|
|
379 |
assume "a \<in> \<nat>"
|
|
380 |
then obtain n where [simp]: "a = of_nat n" by (auto elim: Nats_cases)
|
|
381 |
from eventually_gt_at_top[of n]
|
|
382 |
have "eventually (\<lambda>n. (a gchoose n) * z ^ n = 0) at_top"
|
|
383 |
by eventually_elim (simp add: binomial_gbinomial [symmetric])
|
|
384 |
from summable_cong[OF this] show ?thesis by simp
|
|
385 |
qed auto
|
|
386 |
qed
|
|
387 |
|
|
388 |
lemma gen_binomial_real:
|
|
389 |
fixes z :: real
|
|
390 |
assumes "norm z < 1"
|
|
391 |
shows "(\<lambda>n. (a gchoose n) * z^n) sums (1 + z) powr a"
|
|
392 |
proof -
|
|
393 |
define K where "K = 1 - (1 - norm z) / 2"
|
|
394 |
from assms have K: "K > 0" "K < 1" "norm z < K"
|
|
395 |
unfolding K_def by (auto simp: field_simps intro!: add_pos_nonneg)
|
|
396 |
let ?f = "\<lambda>n. a gchoose n" and ?f' = "diffs (\<lambda>n. a gchoose n)"
|
|
397 |
have summable_strong: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < 1" for z using that
|
|
398 |
by (intro summable_gbinomial)
|
|
399 |
with K have summable: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < K" for z using that by auto
|
|
400 |
hence summable': "summable (\<lambda>n. ?f' n * z ^ n)" if "norm z < K" for z using that
|
|
401 |
by (intro termdiff_converges[of _ K]) simp_all
|
|
402 |
|
|
403 |
define f f' where [abs_def]: "f z = (\<Sum>n. ?f n * z ^ n)" "f' z = (\<Sum>n. ?f' n * z ^ n)" for z
|
|
404 |
{
|
|
405 |
fix z :: real assume z: "norm z < K"
|
|
406 |
from summable_mult2[OF summable'[OF z], of z]
|
|
407 |
have summable1: "summable (\<lambda>n. ?f' n * z ^ Suc n)" by (simp add: mult_ac)
|
|
408 |
hence summable2: "summable (\<lambda>n. of_nat n * ?f n * z^n)"
|
|
409 |
unfolding diffs_def by (subst (asm) summable_Suc_iff)
|
|
410 |
|
|
411 |
have "(1 + z) * f' z = (\<Sum>n. ?f' n * z^n) + (\<Sum>n. ?f' n * z^Suc n)"
|
|
412 |
unfolding f_f'_def using summable' z by (simp add: algebra_simps suminf_mult)
|
|
413 |
also have "(\<Sum>n. ?f' n * z^n) = (\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n)"
|
|
414 |
by (intro suminf_cong) (simp add: diffs_def)
|
|
415 |
also have "(\<Sum>n. ?f' n * z^Suc n) = (\<Sum>n. of_nat n * ?f n * z ^ n)"
|
|
416 |
using summable1 suminf_split_initial_segment[OF summable1] unfolding diffs_def
|
|
417 |
by (subst suminf_split_head, subst (asm) summable_Suc_iff) simp_all
|
|
418 |
also have "(\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n) + (\<Sum>n. of_nat n * ?f n * z^n) =
|
|
419 |
(\<Sum>n. a * ?f n * z^n)"
|
|
420 |
by (subst gbinomial_mult_1, subst suminf_add)
|
|
421 |
(insert summable'[OF z] summable2,
|
|
422 |
simp_all add: summable_powser_split_head algebra_simps diffs_def)
|
|
423 |
also have "\<dots> = a * f z" unfolding f_f'_def
|
|
424 |
by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac)
|
|
425 |
finally have "a * f z = (1 + z) * f' z" by simp
|
|
426 |
} note deriv = this
|
|
427 |
|
|
428 |
have [derivative_intros]: "(f has_field_derivative f' z) (at z)" if "norm z < of_real K" for z
|
|
429 |
unfolding f_f'_def using K that
|
|
430 |
by (intro termdiffs_strong[of "?f" K z] summable_strong) simp_all
|
|
431 |
have "f 0 = (\<Sum>n. if n = 0 then 1 else 0)" unfolding f_f'_def by (intro suminf_cong) simp
|
|
432 |
also have "\<dots> = 1" using sums_single[of 0 "\<lambda>_. 1::real"] unfolding sums_iff by simp
|
|
433 |
finally have [simp]: "f 0 = 1" .
|
|
434 |
|
|
435 |
have "f z * (1 + z) powr (-a) = f 0 * (1 + 0) powr (-a)"
|
|
436 |
proof (rule DERIV_isconst3[where ?f = "\<lambda>x. f x * (1 + x) powr (-a)"])
|
|
437 |
fix z :: real assume z': "z \<in> {-K<..<K}"
|
|
438 |
hence z: "norm z < K" using K by auto
|
|
439 |
with K have nz: "1 + z \<noteq> 0" by (auto dest!: minus_unique)
|
|
440 |
from z K have "norm z < 1" by simp
|
|
441 |
hence "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative
|
|
442 |
f' z * (1 + z) powr (-a) - a * f z * (1 + z) powr (-a-1)) (at z)" using z
|
|
443 |
by (auto intro!: derivative_eq_intros)
|
|
444 |
also from z have "a * f z = (1 + z) * f' z" by (rule deriv)
|
|
445 |
also have "f' z * (1 + z) powr - a - (1 + z) * f' z * (1 + z) powr (- a - 1) = 0"
|
|
446 |
using \<open>norm z < 1\<close> by (auto simp add: field_simps powr_diff)
|
|
447 |
finally show "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z)" .
|
|
448 |
qed (insert K, auto)
|
|
449 |
also have "f 0 * (1 + 0) powr -a = 1" by simp
|
|
450 |
finally have "f z = (1 + z) powr a" using K
|
|
451 |
by (simp add: field_simps dist_real_def powr_minus)
|
|
452 |
with summable K show ?thesis unfolding f_f'_def by (simp add: sums_iff)
|
|
453 |
qed
|
|
454 |
|
|
455 |
lemma convergent_powser'_gbinomial:
|
|
456 |
"convergent_powser' (gbinomial_series abort p) (\<lambda>x. (1 + x) powr p)"
|
|
457 |
proof -
|
|
458 |
have "(\<lambda>n. lcoeff (gbinomial_series abort p) n * x ^ n) sums (1 + x) powr p" if "abs x < 1" for x
|
|
459 |
using that gen_binomial_real[of x p] by simp
|
|
460 |
thus ?thesis unfolding convergent_powser'_def by (force intro!: exI[of _ 1])
|
|
461 |
qed
|
|
462 |
|
|
463 |
lemma convergent_powser'_gbinomial':
|
|
464 |
"convergent_powser' (gbinomial_series abort (real n)) (\<lambda>x. (1 + x) ^ n)"
|
|
465 |
proof -
|
|
466 |
{
|
|
467 |
fix x :: real
|
|
468 |
have "(\<lambda>k. if k \<in> {0..n} then real (n choose k) * x ^ k else 0) sums (x + 1) ^ n"
|
|
469 |
using sums_If_finite_set[of "{..n}" "\<lambda>k. real (n choose k) * x ^ k"]
|
|
470 |
by (subst binomial_ring) simp
|
|
471 |
also have "(\<lambda>k. if k \<in> {0..n} then real (n choose k) * x ^ k else 0) =
|
|
472 |
(\<lambda>k. lcoeff (gbinomial_series abort (real n)) k * x ^ k)"
|
|
473 |
by (simp add: fun_eq_iff binomial_gbinomial [symmetric])
|
|
474 |
finally have "\<dots> sums (1 + x) ^ n" by (simp add: add_ac)
|
|
475 |
}
|
|
476 |
thus ?thesis unfolding convergent_powser'_def
|
|
477 |
by (auto intro!: exI[of _ 1])
|
|
478 |
qed
|
|
479 |
|
|
480 |
|
|
481 |
subsubsection \<open>Sine and cosine\<close>
|
|
482 |
|
|
483 |
primcorec sin_series_stream_aux :: "bool \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real msstream" where
|
|
484 |
"sin_series_stream_aux b acc m =
|
|
485 |
MSSCons (if b then -inverse acc else inverse acc)
|
|
486 |
(sin_series_stream_aux (\<not>b) (acc * (m + 1) * (m + 2)) (m + 2))"
|
|
487 |
|
|
488 |
lemma mssnth_sin_series_stream_aux:
|
|
489 |
"mssnth (sin_series_stream_aux b (fact m) m) n =
|
|
490 |
(if b then -1 else 1) * (-1) ^ n / (fact (2 * n + m))"
|
|
491 |
proof (induction n arbitrary: b m)
|
|
492 |
case (0 b m)
|
|
493 |
thus ?case by (subst sin_series_stream_aux.code) (simp add: field_simps)
|
|
494 |
next
|
|
495 |
case (Suc n b m)
|
|
496 |
show ?case using Suc.IH[of "\<not>b" "m + 2"]
|
|
497 |
by (subst sin_series_stream_aux.code) (auto simp: field_simps)
|
|
498 |
qed
|
|
499 |
|
|
500 |
definition sin_series_stream where
|
|
501 |
"sin_series_stream = sin_series_stream_aux False 1 1"
|
|
502 |
|
|
503 |
lemma mssnth_sin_series_stream:
|
|
504 |
"mssnth sin_series_stream n = (-1) ^ n / fact (2 * n + 1)"
|
|
505 |
using mssnth_sin_series_stream_aux[of False 1 n] unfolding sin_series_stream_def by simp
|
|
506 |
|
|
507 |
definition cos_series_stream where
|
|
508 |
"cos_series_stream = sin_series_stream_aux False 1 0"
|
|
509 |
|
|
510 |
lemma mssnth_cos_series_stream:
|
|
511 |
"mssnth cos_series_stream n = (-1) ^ n / fact (2 * n)"
|
|
512 |
using mssnth_sin_series_stream_aux[of False 0 n] unfolding cos_series_stream_def by simp
|
|
513 |
|
|
514 |
lemma summable_pre_sin: "summable (\<lambda>n. mssnth sin_series_stream n * (x::real) ^ n)"
|
|
515 |
proof -
|
|
516 |
have *: "summable (\<lambda>n. 1 / fact n * abs x ^ n)"
|
|
517 |
using exp_converges[of "abs x"] by (simp add: sums_iff field_simps)
|
|
518 |
{
|
|
519 |
fix n :: nat
|
|
520 |
have "\<bar>x\<bar> ^ n / fact (2 * n + 1) \<le> \<bar>x\<bar> ^ n / fact n"
|
|
521 |
by (intro divide_left_mono fact_mono) auto
|
|
522 |
hence "norm (mssnth sin_series_stream n * x ^ n) \<le> 1 / fact n * abs x ^ n"
|
|
523 |
by (simp add: mssnth_sin_series_stream abs_mult power_abs field_simps)
|
|
524 |
}
|
|
525 |
thus ?thesis
|
|
526 |
by (intro summable_comparison_test[OF _ *]) (auto intro!: exI[of _ 0])
|
|
527 |
qed
|
|
528 |
|
|
529 |
lemma summable_pre_cos: "summable (\<lambda>n. mssnth cos_series_stream n * (x::real) ^ n)"
|
|
530 |
proof -
|
|
531 |
have *: "summable (\<lambda>n. 1 / fact n * abs x ^ n)"
|
|
532 |
using exp_converges[of "abs x"] by (simp add: sums_iff field_simps)
|
|
533 |
{
|
|
534 |
fix n :: nat
|
|
535 |
have "\<bar>x\<bar> ^ n / fact (2 * n) \<le> \<bar>x\<bar> ^ n / fact n"
|
|
536 |
by (intro divide_left_mono fact_mono) auto
|
|
537 |
hence "norm (mssnth cos_series_stream n * x ^ n) \<le> 1 / fact n * abs x ^ n"
|
|
538 |
by (simp add: mssnth_cos_series_stream abs_mult power_abs field_simps)
|
|
539 |
}
|
|
540 |
thus ?thesis
|
|
541 |
by (intro summable_comparison_test[OF _ *]) (auto intro!: exI[of _ 0])
|
|
542 |
qed
|
|
543 |
|
|
544 |
lemma cos_conv_pre_cos:
|
|
545 |
"cos x = powser (msllist_of_msstream cos_series_stream) (x ^ 2)"
|
|
546 |
proof -
|
|
547 |
have "(\<lambda>n. cos_coeff (2 * n) * x ^ (2 * n)) sums cos x"
|
|
548 |
using cos_converges[of x]
|
|
549 |
by (subst sums_mono_reindex[of "\<lambda>n. 2 * n"])
|
|
550 |
(auto simp: strict_mono_def cos_coeff_def elim!: evenE)
|
|
551 |
also have "(\<lambda>n. cos_coeff (2 * n) * x ^ (2 * n)) =
|
|
552 |
(\<lambda>n. mssnth cos_series_stream n * (x ^ 2) ^ n)"
|
|
553 |
by (simp add: fun_eq_iff mssnth_cos_series_stream cos_coeff_def power_mult)
|
|
554 |
finally have sums: "(\<lambda>n. mssnth cos_series_stream n * x\<^sup>2 ^ n) sums cos x" .
|
|
555 |
thus ?thesis by (simp add: powser_def sums_iff)
|
|
556 |
qed
|
|
557 |
|
|
558 |
lemma sin_conv_pre_sin:
|
|
559 |
"sin x = x * powser (msllist_of_msstream sin_series_stream) (x ^ 2)"
|
|
560 |
proof -
|
|
561 |
have "(\<lambda>n. sin_coeff (2 * n + 1) * x ^ (2 * n + 1)) sums sin x"
|
|
562 |
using sin_converges[of x]
|
|
563 |
by (subst sums_mono_reindex[of "\<lambda>n. 2 * n + 1"])
|
|
564 |
(auto simp: strict_mono_def sin_coeff_def elim!: oddE)
|
|
565 |
also have "(\<lambda>n. sin_coeff (2 * n + 1) * x ^ (2 * n + 1)) =
|
|
566 |
(\<lambda>n. x * mssnth sin_series_stream n * (x ^ 2) ^ n)"
|
|
567 |
by (simp add: fun_eq_iff mssnth_sin_series_stream sin_coeff_def power_mult [symmetric] algebra_simps)
|
|
568 |
finally have sums: "(\<lambda>n. x * mssnth sin_series_stream n * x\<^sup>2 ^ n) sums sin x" .
|
|
569 |
thus ?thesis using summable_pre_sin[of "x^2"]
|
|
570 |
by (simp add: powser_def sums_iff suminf_mult [symmetric] mult.assoc)
|
|
571 |
qed
|
|
572 |
|
|
573 |
lemma convergent_powser_sin_series_stream:
|
|
574 |
"convergent_powser (msllist_of_msstream sin_series_stream)"
|
|
575 |
(is "convergent_powser ?cs")
|
|
576 |
proof -
|
|
577 |
show ?thesis using summable_pre_sin unfolding convergent_powser_def
|
|
578 |
by (intro exI[of _ 1]) auto
|
|
579 |
qed
|
|
580 |
|
|
581 |
lemma convergent_powser_cos_series_stream:
|
|
582 |
"convergent_powser (msllist_of_msstream cos_series_stream)"
|
|
583 |
(is "convergent_powser ?cs")
|
|
584 |
proof -
|
|
585 |
show ?thesis using summable_pre_cos unfolding convergent_powser_def
|
|
586 |
by (intro exI[of _ 1]) auto
|
|
587 |
qed
|
|
588 |
|
|
589 |
|
|
590 |
subsubsection \<open>Arc tangent\<close>
|
|
591 |
|
|
592 |
definition arctan_coeffs :: "nat \<Rightarrow> 'a :: {real_normed_div_algebra, banach}" where
|
|
593 |
"arctan_coeffs n = (if odd n then (-1) ^ (n div 2) / of_nat n else 0)"
|
|
594 |
|
|
595 |
lemma arctan_converges:
|
|
596 |
assumes "norm x < 1"
|
|
597 |
shows "(\<lambda>n. arctan_coeffs n * x ^ n) sums arctan x"
|
|
598 |
proof -
|
|
599 |
have A: "(\<lambda>n. diffs arctan_coeffs n * x ^ n) sums (1 / (1 + x^2))" if "norm x < 1" for x :: real
|
|
600 |
proof -
|
|
601 |
let ?f = "\<lambda>n. (if odd n then 0 else (-1) ^ (n div 2)) * x ^ n"
|
|
602 |
from that have "norm x ^ 2 < 1 ^ 2" by (intro power_strict_mono) simp_all
|
|
603 |
hence "(\<lambda>n. (-(x^2))^n) sums (1 / (1 - (-(x^2))))" by (intro geometric_sums) simp_all
|
|
604 |
also have "1 - (-(x^2)) = 1 + x^2" by simp
|
|
605 |
also have "(\<lambda>n. (-(x^2))^n) = (\<lambda>n. ?f (2*n))" by (auto simp: fun_eq_iff power_minus' power_mult)
|
|
606 |
also have "range (( *) (2::nat)) = {n. even n}"
|
|
607 |
by (auto elim!: evenE)
|
|
608 |
hence "(\<lambda>n. ?f (2*n)) sums (1 / (1 + x^2)) \<longleftrightarrow> ?f sums (1 / (1 + x^2))"
|
|
609 |
by (intro sums_mono_reindex) (simp_all add: strict_mono_Suc_iff)
|
|
610 |
also have "?f = (\<lambda>n. diffs arctan_coeffs n * x ^ n)"
|
|
611 |
by (simp add: fun_eq_iff arctan_coeffs_def diffs_def)
|
|
612 |
finally show "(\<lambda>n. diffs arctan_coeffs n * x ^ n) sums (1 / (1 + x^2))" .
|
|
613 |
qed
|
|
614 |
|
|
615 |
have B: "summable (\<lambda>n. arctan_coeffs n * x ^ n)" if "norm x < 1" for x :: real
|
|
616 |
proof (rule summable_comparison_test)
|
|
617 |
show "\<exists>N. \<forall>n\<ge>N. norm (arctan_coeffs n * x ^ n) \<le> 1 * norm x ^ n"
|
|
618 |
unfolding norm_mult norm_power
|
|
619 |
by (intro exI[of _ "0::nat"] allI impI mult_right_mono)
|
|
620 |
(simp_all add: arctan_coeffs_def divide_simps)
|
|
621 |
from that show "summable (\<lambda>n. 1 * norm x ^ n)"
|
|
622 |
by (intro summable_mult summable_geometric) simp_all
|
|
623 |
qed
|
|
624 |
|
|
625 |
define F :: "real \<Rightarrow> real" where "F = (\<lambda>x. \<Sum>n. arctan_coeffs n * x ^ n)"
|
|
626 |
have [derivative_intros]:
|
|
627 |
"(F has_field_derivative (1 / (1 + x^2))) (at x)" if "norm x < 1" for x :: real
|
|
628 |
proof -
|
|
629 |
define K :: real where "K = (1 + norm x) / 2"
|
|
630 |
from that have K: "norm x < K" "K < 1" by (simp_all add: K_def field_simps)
|
|
631 |
have "(F has_field_derivative (\<Sum>n. diffs arctan_coeffs n * x ^ n)) (at x)"
|
|
632 |
unfolding F_def using K by (intro termdiffs_strong[where K = K] B) simp_all
|
|
633 |
also from A[OF that] have "(\<Sum>n. diffs arctan_coeffs n * x ^ n) = 1 / (1 + x^2)"
|
|
634 |
by (simp add: sums_iff)
|
|
635 |
finally show ?thesis .
|
|
636 |
qed
|
|
637 |
from assms have "arctan x - F x = arctan 0 - F 0"
|
|
638 |
by (intro DERIV_isconst3[of "-1" 1 _ _ "\<lambda>x. arctan x - F x"])
|
|
639 |
(auto intro!: derivative_eq_intros simp: field_simps)
|
|
640 |
hence "F x = arctan x" by (simp add: F_def arctan_coeffs_def)
|
|
641 |
with B[OF assms] show ?thesis by (simp add: sums_iff F_def)
|
|
642 |
qed
|
|
643 |
|
|
644 |
primcorec arctan_series_stream_aux :: "bool \<Rightarrow> real \<Rightarrow> real msstream" where
|
|
645 |
"arctan_series_stream_aux b n =
|
|
646 |
MSSCons (if b then -inverse n else inverse n) (arctan_series_stream_aux (\<not>b) (n + 2))"
|
|
647 |
|
|
648 |
lemma mssnth_arctan_series_stream_aux:
|
|
649 |
"mssnth (arctan_series_stream_aux b n) m = (-1) ^ (if b then Suc m else m) / (2*m + n)"
|
|
650 |
by (induction m arbitrary: b n; subst arctan_series_stream_aux.code)
|
|
651 |
(auto simp: field_simps split: if_splits)
|
|
652 |
|
|
653 |
definition arctan_series_stream where
|
|
654 |
"arctan_series_stream = arctan_series_stream_aux False 1"
|
|
655 |
|
|
656 |
lemma mssnth_arctan_series_stream:
|
|
657 |
"mssnth arctan_series_stream n = (-1) ^ n / (2 * n + 1)"
|
|
658 |
by (simp add: arctan_series_stream_def mssnth_arctan_series_stream_aux)
|
|
659 |
|
|
660 |
lemma summable_pre_arctan:
|
|
661 |
assumes "norm x < 1"
|
|
662 |
shows "summable (\<lambda>n. mssnth arctan_series_stream n * x ^ n)" (is "summable ?f")
|
|
663 |
proof (rule summable_comparison_test)
|
|
664 |
show "summable (\<lambda>n. abs x ^ n)" using assms by (intro summable_geometric) auto
|
|
665 |
show "\<exists>N. \<forall>n\<ge>N. norm (?f n) \<le> abs x ^ n"
|
|
666 |
proof (intro exI[of _ 0] allI impI)
|
|
667 |
fix n :: nat
|
|
668 |
have "norm (?f n) = \<bar>x\<bar> ^ n / (2 * real n + 1)"
|
|
669 |
by (simp add: mssnth_arctan_series_stream abs_mult power_abs)
|
|
670 |
also have "\<dots> \<le> \<bar>x\<bar> ^ n / 1" by (intro divide_left_mono) auto
|
|
671 |
finally show "norm (?f n) \<le> abs x ^ n" by simp
|
|
672 |
qed
|
|
673 |
qed
|
|
674 |
|
|
675 |
lemma arctan_conv_pre_arctan:
|
|
676 |
assumes "norm x < 1"
|
|
677 |
shows "arctan x = x * powser (msllist_of_msstream arctan_series_stream) (x ^ 2)"
|
|
678 |
proof -
|
|
679 |
from assms have "norm x * norm x < 1 * 1" by (intro mult_strict_mono) auto
|
|
680 |
hence norm_less: "norm (x ^ 2) < 1" by (simp add: power2_eq_square)
|
|
681 |
have "(\<lambda>n. arctan_coeffs n * x ^ n) sums arctan x"
|
|
682 |
by (intro arctan_converges assms)
|
|
683 |
also have "?this \<longleftrightarrow> (\<lambda>n. arctan_coeffs (2 * n + 1) * x ^ (2 * n + 1)) sums arctan x"
|
|
684 |
by (intro sums_mono_reindex [symmetric])
|
|
685 |
(auto simp: arctan_coeffs_def strict_mono_def elim!: oddE)
|
|
686 |
also have "(\<lambda>n. arctan_coeffs (2 * n + 1) * x ^ (2 * n + 1)) =
|
|
687 |
(\<lambda>n. x * mssnth arctan_series_stream n * (x ^ 2) ^ n)"
|
|
688 |
by (simp add: fun_eq_iff mssnth_arctan_series_stream
|
|
689 |
power_mult [symmetric] arctan_coeffs_def mult_ac)
|
|
690 |
finally have "(\<lambda>n. x * mssnth arctan_series_stream n * x\<^sup>2 ^ n) sums arctan x" .
|
|
691 |
thus ?thesis using summable_pre_arctan[OF norm_less] assms
|
|
692 |
by (simp add: powser_def sums_iff suminf_mult [symmetric] mult.assoc)
|
|
693 |
qed
|
|
694 |
|
|
695 |
lemma convergent_powser_arctan:
|
|
696 |
"convergent_powser (msllist_of_msstream arctan_series_stream)"
|
|
697 |
unfolding convergent_powser_def
|
|
698 |
using summable_pre_arctan by (intro exI[of _ 1]) auto
|
|
699 |
|
|
700 |
lemma arctan_inverse_pos: "x > 0 \<Longrightarrow> arctan x = pi / 2 - arctan (inverse x)"
|
|
701 |
using arctan_inverse[of x] by (simp add: field_simps)
|
|
702 |
|
|
703 |
lemma arctan_inverse_neg: "x < 0 \<Longrightarrow> arctan x = -pi / 2 - arctan (inverse x)"
|
|
704 |
using arctan_inverse[of x] by (simp add: field_simps)
|
|
705 |
|
|
706 |
|
|
707 |
|
|
708 |
subsection \<open>Multiseries\<close>
|
|
709 |
|
|
710 |
subsubsection \<open>Asymptotic bases\<close>
|
|
711 |
|
|
712 |
type_synonym basis = "(real \<Rightarrow> real) list"
|
|
713 |
|
|
714 |
lemma transp_ln_smallo_ln: "transp (\<lambda>f g. (\<lambda>x::real. ln (g x)) \<in> o(\<lambda>x. ln (f x)))"
|
|
715 |
by (rule transpI, erule landau_o.small.trans)
|
|
716 |
|
|
717 |
definition basis_wf :: "basis \<Rightarrow> bool" where
|
|
718 |
"basis_wf basis \<longleftrightarrow> (\<forall>f\<in>set basis. filterlim f at_top at_top) \<and>
|
|
719 |
sorted_wrt (\<lambda>f g. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (f x))) basis"
|
|
720 |
|
|
721 |
lemma basis_wf_Nil [simp]: "basis_wf []"
|
|
722 |
by (simp add: basis_wf_def)
|
|
723 |
|
|
724 |
lemma basis_wf_Cons:
|
|
725 |
"basis_wf (f # basis) \<longleftrightarrow>
|
|
726 |
filterlim f at_top at_top \<and> (\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (f x))) \<and> basis_wf basis"
|
|
727 |
unfolding basis_wf_def by auto
|
|
728 |
|
|
729 |
(* TODO: Move *)
|
|
730 |
lemma sorted_wrt_snoc:
|
|
731 |
assumes trans: "transp P" and "sorted_wrt P xs" "P (last xs) y"
|
|
732 |
shows "sorted_wrt P (xs @ [y])"
|
|
733 |
proof (subst sorted_wrt_append, intro conjI ballI)
|
|
734 |
fix x y' assume x: "x \<in> set xs" and y': "y' \<in> set [y]"
|
|
735 |
from y' have [simp]: "y' = y" by simp
|
|
736 |
show "P x y'"
|
|
737 |
proof (cases "x = last xs")
|
|
738 |
case False
|
|
739 |
from x have eq: "xs = butlast xs @ [last xs]"
|
|
740 |
by (subst append_butlast_last_id) auto
|
|
741 |
from x and False have x': "x \<in> set (butlast xs)" by (subst (asm) eq) auto
|
|
742 |
have "sorted_wrt P xs" by fact
|
|
743 |
also note eq
|
|
744 |
finally have "P x (last xs)" using x'
|
|
745 |
by (subst (asm) sorted_wrt_append) auto
|
|
746 |
with \<open>P (last xs) y\<close> have "P x y" using transpD[OF trans] by blast
|
|
747 |
thus ?thesis by simp
|
|
748 |
qed (insert assms, auto)
|
|
749 |
qed (insert assms, auto)
|
|
750 |
|
|
751 |
lemma basis_wf_snoc:
|
|
752 |
assumes "bs \<noteq> []"
|
|
753 |
assumes "basis_wf bs" "filterlim b at_top at_top"
|
|
754 |
assumes "(\<lambda>x. ln (b x)) \<in> o(\<lambda>x. ln (last bs x))"
|
|
755 |
shows "basis_wf (bs @ [b])"
|
|
756 |
proof -
|
|
757 |
have "transp (\<lambda>f g. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (f x)))"
|
|
758 |
by (auto simp: transp_def intro: landau_o.small.trans)
|
|
759 |
thus ?thesis using assms
|
|
760 |
by (auto simp add: basis_wf_def sorted_wrt_snoc[OF transp_ln_smallo_ln])
|
|
761 |
qed
|
|
762 |
|
|
763 |
lemma basis_wf_singleton: "basis_wf [b] \<longleftrightarrow> filterlim b at_top at_top"
|
|
764 |
by (simp add: basis_wf_Cons)
|
|
765 |
|
|
766 |
lemma basis_wf_many: "basis_wf (b # b' # bs) \<longleftrightarrow>
|
|
767 |
filterlim b at_top at_top \<and> (\<lambda>x. ln (b' x)) \<in> o(\<lambda>x. ln (b x)) \<and> basis_wf (b' # bs)"
|
|
768 |
unfolding basis_wf_def by (subst sorted_wrt2[OF transp_ln_smallo_ln]) auto
|
|
769 |
|
|
770 |
|
|
771 |
subsubsection \<open>Monomials\<close>
|
|
772 |
|
|
773 |
type_synonym monom = "real \<times> real list"
|
|
774 |
|
|
775 |
definition eval_monom :: "monom \<Rightarrow> basis \<Rightarrow> real \<Rightarrow> real" where
|
|
776 |
"eval_monom = (\<lambda>(c,es) basis x. c * prod_list (map (\<lambda>(b,e). b x powr e) (zip basis es)))"
|
|
777 |
|
|
778 |
lemma eval_monom_Nil1 [simp]: "eval_monom (c, []) basis = (\<lambda>_. c)"
|
|
779 |
by (simp add: eval_monom_def split: prod.split)
|
|
780 |
|
|
781 |
lemma eval_monom_Nil2 [simp]: "eval_monom monom [] = (\<lambda>_. fst monom)"
|
|
782 |
by (simp add: eval_monom_def split: prod.split)
|
|
783 |
|
|
784 |
lemma eval_monom_Cons:
|
|
785 |
"eval_monom (c, e # es) (b # basis) = (\<lambda>x. eval_monom (c, es) basis x * b x powr e)"
|
|
786 |
by (simp add: eval_monom_def fun_eq_iff split: prod.split)
|
|
787 |
|
|
788 |
definition inverse_monom :: "monom \<Rightarrow> monom" where
|
|
789 |
"inverse_monom monom = (case monom of (c, es) \<Rightarrow> (inverse c, map uminus es))"
|
|
790 |
|
|
791 |
lemma length_snd_inverse_monom [simp]:
|
|
792 |
"length (snd (inverse_monom monom)) = length (snd monom)"
|
|
793 |
by (simp add: inverse_monom_def split: prod.split)
|
|
794 |
|
|
795 |
lemma eval_monom_pos:
|
|
796 |
assumes "basis_wf basis" "fst monom > 0"
|
|
797 |
shows "eventually (\<lambda>x. eval_monom monom basis x > 0) at_top"
|
|
798 |
proof (cases monom)
|
|
799 |
case (Pair c es)
|
|
800 |
with assms have "c > 0" by simp
|
|
801 |
with assms(1) show ?thesis unfolding Pair
|
|
802 |
proof (induction es arbitrary: basis)
|
|
803 |
case (Cons e es)
|
|
804 |
note A = this
|
|
805 |
thus ?case
|
|
806 |
proof (cases basis)
|
|
807 |
case (Cons b basis')
|
|
808 |
with A(1)[of basis'] A(2,3)
|
|
809 |
have A: "\<forall>\<^sub>F x in at_top. eval_monom (c, es) basis' x > 0"
|
|
810 |
and B: "eventually (\<lambda>x. b x > 0) at_top"
|
|
811 |
by (simp_all add: basis_wf_Cons filterlim_at_top_dense)
|
|
812 |
thus ?thesis by eventually_elim (simp add: Cons eval_monom_Cons)
|
|
813 |
qed simp
|
|
814 |
qed simp
|
|
815 |
qed
|
|
816 |
|
|
817 |
lemma eval_monom_uminus: "eval_monom (-c, es) basis x = -eval_monom (c, es) basis x"
|
|
818 |
by (simp add: eval_monom_def)
|
|
819 |
|
|
820 |
lemma eval_monom_neg:
|
|
821 |
assumes "basis_wf basis" "fst monom < 0"
|
|
822 |
shows "eventually (\<lambda>x. eval_monom monom basis x < 0) at_top"
|
|
823 |
proof -
|
|
824 |
from assms have "eventually (\<lambda>x. eval_monom (-fst monom, snd monom) basis x > 0) at_top"
|
|
825 |
by (intro eval_monom_pos) simp_all
|
|
826 |
thus ?thesis by (simp add: eval_monom_uminus)
|
|
827 |
qed
|
|
828 |
|
|
829 |
lemma eval_monom_nonzero:
|
|
830 |
assumes "basis_wf basis" "fst monom \<noteq> 0"
|
|
831 |
shows "eventually (\<lambda>x. eval_monom monom basis x \<noteq> 0) at_top"
|
|
832 |
proof (cases "fst monom" "0 :: real" rule: linorder_cases)
|
|
833 |
case greater
|
|
834 |
with assms have "eventually (\<lambda>x. eval_monom monom basis x > 0) at_top"
|
|
835 |
by (simp add: eval_monom_pos)
|
|
836 |
thus ?thesis by eventually_elim simp
|
|
837 |
next
|
|
838 |
case less
|
|
839 |
with assms have "eventually (\<lambda>x. eval_monom (-fst monom, snd monom) basis x > 0) at_top"
|
|
840 |
by (simp add: eval_monom_pos)
|
|
841 |
thus ?thesis by eventually_elim (simp add: eval_monom_uminus)
|
|
842 |
qed (insert assms, simp_all)
|
|
843 |
|
|
844 |
|
|
845 |
subsubsection \<open>Typeclass for multiseries\<close>
|
|
846 |
|
|
847 |
class multiseries = plus + minus + times + uminus + inverse +
|
|
848 |
fixes is_expansion :: "'a \<Rightarrow> basis \<Rightarrow> bool"
|
|
849 |
and expansion_level :: "'a itself \<Rightarrow> nat"
|
|
850 |
and eval :: "'a \<Rightarrow> real \<Rightarrow> real"
|
|
851 |
and zero_expansion :: 'a
|
|
852 |
and const_expansion :: "real \<Rightarrow> 'a"
|
|
853 |
and powr_expansion :: "bool \<Rightarrow> 'a \<Rightarrow> real \<Rightarrow> 'a"
|
|
854 |
and power_expansion :: "bool \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
|
|
855 |
and trimmed :: "'a \<Rightarrow> bool"
|
|
856 |
and dominant_term :: "'a \<Rightarrow> monom"
|
|
857 |
|
|
858 |
assumes is_expansion_length:
|
|
859 |
"is_expansion F basis \<Longrightarrow> length basis = expansion_level (TYPE('a))"
|
|
860 |
assumes is_expansion_zero:
|
|
861 |
"basis_wf basis \<Longrightarrow> length basis = expansion_level (TYPE('a)) \<Longrightarrow>
|
|
862 |
is_expansion zero_expansion basis"
|
|
863 |
assumes is_expansion_const:
|
|
864 |
"basis_wf basis \<Longrightarrow> length basis = expansion_level (TYPE('a)) \<Longrightarrow>
|
|
865 |
is_expansion (const_expansion c) basis"
|
|
866 |
assumes is_expansion_uminus:
|
|
867 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion (-F) basis"
|
|
868 |
assumes is_expansion_add:
|
|
869 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow>
|
|
870 |
is_expansion (F + G) basis"
|
|
871 |
assumes is_expansion_minus:
|
|
872 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow>
|
|
873 |
is_expansion (F - G) basis"
|
|
874 |
assumes is_expansion_mult:
|
|
875 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow>
|
|
876 |
is_expansion (F * G) basis"
|
|
877 |
assumes is_expansion_inverse:
|
|
878 |
"basis_wf basis \<Longrightarrow> trimmed F \<Longrightarrow> is_expansion F basis \<Longrightarrow>
|
|
879 |
is_expansion (inverse F) basis"
|
|
880 |
assumes is_expansion_divide:
|
|
881 |
"basis_wf basis \<Longrightarrow> trimmed G \<Longrightarrow> is_expansion F basis \<Longrightarrow> is_expansion G basis \<Longrightarrow>
|
|
882 |
is_expansion (F / G) basis"
|
|
883 |
assumes is_expansion_powr:
|
|
884 |
"basis_wf basis \<Longrightarrow> trimmed F \<Longrightarrow> fst (dominant_term F) > 0 \<Longrightarrow> is_expansion F basis \<Longrightarrow>
|
|
885 |
is_expansion (powr_expansion abort F p) basis"
|
|
886 |
assumes is_expansion_power:
|
|
887 |
"basis_wf basis \<Longrightarrow> trimmed F \<Longrightarrow> is_expansion F basis \<Longrightarrow>
|
|
888 |
is_expansion (power_expansion abort F n) basis"
|
|
889 |
|
|
890 |
assumes is_expansion_imp_smallo:
|
|
891 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> filterlim b at_top at_top \<Longrightarrow>
|
|
892 |
(\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (b x))) \<Longrightarrow> e > 0 \<Longrightarrow> eval F \<in> o(\<lambda>x. b x powr e)"
|
|
893 |
assumes is_expansion_imp_smallomega:
|
|
894 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> filterlim b at_top at_top \<Longrightarrow> trimmed F \<Longrightarrow>
|
|
895 |
(\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (b x))) \<Longrightarrow> e < 0 \<Longrightarrow> eval F \<in> \<omega>(\<lambda>x. b x powr e)"
|
|
896 |
assumes trimmed_imp_eventually_sgn:
|
|
897 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> trimmed F \<Longrightarrow>
|
|
898 |
eventually (\<lambda>x. sgn (eval F x) = sgn (fst (dominant_term F))) at_top"
|
|
899 |
assumes trimmed_imp_eventually_nz:
|
|
900 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> trimmed F \<Longrightarrow>
|
|
901 |
eventually (\<lambda>x. eval F x \<noteq> 0) at_top"
|
|
902 |
assumes trimmed_imp_dominant_term_nz: "trimmed F \<Longrightarrow> fst (dominant_term F) \<noteq> 0"
|
|
903 |
|
|
904 |
assumes dominant_term:
|
|
905 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow> trimmed F \<Longrightarrow>
|
|
906 |
eval F \<sim>[at_top] eval_monom (dominant_term F) basis"
|
|
907 |
assumes dominant_term_bigo:
|
|
908 |
"basis_wf basis \<Longrightarrow> is_expansion F basis \<Longrightarrow>
|
|
909 |
eval F \<in> O(eval_monom (1, snd (dominant_term F)) basis)"
|
|
910 |
assumes length_dominant_term:
|
|
911 |
"length (snd (dominant_term F)) = expansion_level TYPE('a)"
|
|
912 |
assumes fst_dominant_term_uminus [simp]: "fst (dominant_term (- F)) = -fst (dominant_term F)"
|
|
913 |
assumes trimmed_uminus_iff [simp]: "trimmed (-F) \<longleftrightarrow> trimmed F"
|
|
914 |
|
|
915 |
assumes add_zero_expansion_left [simp]: "zero_expansion + F = F"
|
|
916 |
assumes add_zero_expansion_right [simp]: "F + zero_expansion = F"
|
|
917 |
|
|
918 |
assumes eval_zero [simp]: "eval zero_expansion x = 0"
|
|
919 |
assumes eval_const [simp]: "eval (const_expansion c) x = c"
|
|
920 |
assumes eval_uminus [simp]: "eval (-F) = (\<lambda>x. -eval F x)"
|
|
921 |
assumes eval_plus [simp]: "eval (F + G) = (\<lambda>x. eval F x + eval G x)"
|
|
922 |
assumes eval_minus [simp]: "eval (F - G) = (\<lambda>x. eval F x - eval G x)"
|
|
923 |
assumes eval_times [simp]: "eval (F * G) = (\<lambda>x. eval F x * eval G x)"
|
|
924 |
assumes eval_inverse [simp]: "eval (inverse F) = (\<lambda>x. inverse (eval F x))"
|
|
925 |
assumes eval_divide [simp]: "eval (F / G) = (\<lambda>x. eval F x / eval G x)"
|
|
926 |
assumes eval_powr [simp]: "eval (powr_expansion abort F p) = (\<lambda>x. eval F x powr p)"
|
|
927 |
assumes eval_power [simp]: "eval (power_expansion abort F n) = (\<lambda>x. eval F x ^ n)"
|
|
928 |
assumes minus_eq_plus_uminus: "F - G = F + (-G)"
|
|
929 |
assumes times_const_expansion_1: "const_expansion 1 * F = F"
|
|
930 |
assumes trimmed_const_expansion: "trimmed (const_expansion c) \<longleftrightarrow> c \<noteq> 0"
|
|
931 |
begin
|
|
932 |
|
|
933 |
definition trimmed_pos where
|
|
934 |
"trimmed_pos F \<longleftrightarrow> trimmed F \<and> fst (dominant_term F) > 0"
|
|
935 |
|
|
936 |
definition trimmed_neg where
|
|
937 |
"trimmed_neg F \<longleftrightarrow> trimmed F \<and> fst (dominant_term F) < 0"
|
|
938 |
|
|
939 |
lemma trimmed_pos_uminus: "trimmed_neg F \<Longrightarrow> trimmed_pos (-F)"
|
|
940 |
by (simp add: trimmed_neg_def trimmed_pos_def)
|
|
941 |
|
|
942 |
lemma trimmed_pos_imp_trimmed: "trimmed_pos x \<Longrightarrow> trimmed x"
|
|
943 |
by (simp add: trimmed_pos_def)
|
|
944 |
|
|
945 |
lemma trimmed_neg_imp_trimmed: "trimmed_neg x \<Longrightarrow> trimmed x"
|
|
946 |
by (simp add: trimmed_neg_def)
|
|
947 |
|
|
948 |
end
|
|
949 |
|
|
950 |
|
|
951 |
subsubsection \<open>Zero-rank expansions\<close>
|
|
952 |
|
|
953 |
instantiation real :: multiseries
|
|
954 |
begin
|
|
955 |
|
|
956 |
inductive is_expansion_real :: "real \<Rightarrow> basis \<Rightarrow> bool" where
|
|
957 |
"is_expansion_real c []"
|
|
958 |
|
|
959 |
definition expansion_level_real :: "real itself \<Rightarrow> nat" where
|
|
960 |
expansion_level_real_def [simp]: "expansion_level_real _ = 0"
|
|
961 |
|
|
962 |
definition eval_real :: "real \<Rightarrow> real \<Rightarrow> real" where
|
|
963 |
eval_real_def [simp]: "eval_real c = (\<lambda>_. c)"
|
|
964 |
|
|
965 |
definition zero_expansion_real :: "real" where
|
|
966 |
zero_expansion_real_def [simp]: "zero_expansion_real = 0"
|
|
967 |
|
|
968 |
definition const_expansion_real :: "real \<Rightarrow> real" where
|
|
969 |
const_expansion_real_def [simp]: "const_expansion_real x = x"
|
|
970 |
|
|
971 |
definition powr_expansion_real :: "bool \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where
|
|
972 |
powr_expansion_real_def [simp]: "powr_expansion_real _ x p = x powr p"
|
|
973 |
|
|
974 |
definition power_expansion_real :: "bool \<Rightarrow> real \<Rightarrow> nat \<Rightarrow> real" where
|
|
975 |
power_expansion_real_def [simp]: "power_expansion_real _ x n = x ^ n"
|
|
976 |
|
|
977 |
definition trimmed_real :: "real \<Rightarrow> bool" where
|
|
978 |
trimmed_real_def [simp]: "trimmed_real x \<longleftrightarrow> x \<noteq> 0"
|
|
979 |
|
|
980 |
definition dominant_term_real :: "real \<Rightarrow> monom" where
|
|
981 |
dominant_term_real_def [simp]: "dominant_term_real c = (c, [])"
|
|
982 |
|
|
983 |
instance proof
|
|
984 |
fix basis :: basis and b :: "real \<Rightarrow> real" and e F :: real
|
|
985 |
assume *: "basis_wf basis" "filterlim b at_top at_top" "is_expansion F basis" "0 < e"
|
|
986 |
have "(\<lambda>x. b x powr e) \<in> \<omega>(\<lambda>_. 1)"
|
|
987 |
by (intro smallomegaI_filterlim_at_infinity filterlim_at_top_imp_at_infinity)
|
|
988 |
(auto intro!: filterlim_compose[OF real_powr_at_top] * )
|
|
989 |
with * show "eval F \<in> o(\<lambda>x. b x powr e)"
|
|
990 |
by (cases "F = 0") (auto elim!: is_expansion_real.cases simp: smallomega_iff_smallo)
|
|
991 |
next
|
|
992 |
fix basis :: basis and b :: "real \<Rightarrow> real" and e F :: real
|
|
993 |
assume *: "basis_wf basis" "filterlim b at_top at_top" "is_expansion F basis"
|
|
994 |
"e < 0" "trimmed F"
|
|
995 |
from * have pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
996 |
have "(\<lambda>x. b x powr -e) \<in> \<omega>(\<lambda>_. 1)"
|
|
997 |
by (intro smallomegaI_filterlim_at_infinity filterlim_at_top_imp_at_infinity)
|
|
998 |
(auto intro!: filterlim_compose[OF real_powr_at_top] * )
|
|
999 |
with pos have "(\<lambda>x. b x powr e) \<in> o(\<lambda>_. 1)" unfolding powr_minus
|
|
1000 |
by (subst (asm) landau_omega.small.inverse_eq2)
|
|
1001 |
(auto elim: eventually_mono simp: smallomega_iff_smallo)
|
|
1002 |
with * show "eval F \<in> \<omega>(\<lambda>x. b x powr e)"
|
|
1003 |
by (auto elim!: is_expansion_real.cases simp: smallomega_iff_smallo)
|
|
1004 |
qed (auto intro!: is_expansion_real.intros elim!: is_expansion_real.cases)
|
|
1005 |
|
|
1006 |
end
|
|
1007 |
|
|
1008 |
lemma eval_real: "eval (c :: real) x = c" by simp
|
|
1009 |
|
|
1010 |
|
|
1011 |
subsubsection \<open>Operations on multiseries\<close>
|
|
1012 |
|
|
1013 |
lemma ms_aux_cases [case_names MSLNil MSLCons]:
|
|
1014 |
fixes xs :: "('a \<times> real) msllist"
|
|
1015 |
obtains "xs = MSLNil" | c e xs' where "xs = MSLCons (c, e) xs'"
|
|
1016 |
proof (cases xs)
|
|
1017 |
case (MSLCons x xs')
|
|
1018 |
with that(2)[of "fst x" "snd x" xs'] show ?thesis by auto
|
|
1019 |
qed auto
|
|
1020 |
|
|
1021 |
definition ms_aux_hd_exp :: "('a \<times> real) msllist \<Rightarrow> real option" where
|
|
1022 |
"ms_aux_hd_exp xs = (case xs of MSLNil \<Rightarrow> None | MSLCons (_, e) _ \<Rightarrow> Some e)"
|
|
1023 |
|
|
1024 |
primrec ms_exp_gt :: "real \<Rightarrow> real option \<Rightarrow> bool" where
|
|
1025 |
"ms_exp_gt x None = True"
|
|
1026 |
| "ms_exp_gt x (Some y) \<longleftrightarrow> x > y"
|
|
1027 |
|
|
1028 |
lemma ms_aux_hd_exp_MSLNil [simp]: "ms_aux_hd_exp MSLNil = None"
|
|
1029 |
by (simp add: ms_aux_hd_exp_def split: prod.split)
|
|
1030 |
|
|
1031 |
lemma ms_aux_hd_exp_MSLCons [simp]: "ms_aux_hd_exp (MSLCons x xs) = Some (snd x)"
|
|
1032 |
by (simp add: ms_aux_hd_exp_def split: prod.split)
|
|
1033 |
|
|
1034 |
coinductive is_expansion_aux ::
|
|
1035 |
"('a :: multiseries \<times> real) msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> basis \<Rightarrow> bool" where
|
|
1036 |
is_expansion_MSLNil:
|
|
1037 |
"eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> length basis = Suc (expansion_level TYPE('a)) \<Longrightarrow>
|
|
1038 |
is_expansion_aux MSLNil f basis"
|
|
1039 |
| is_expansion_MSLCons:
|
|
1040 |
"is_expansion_aux xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis) \<Longrightarrow>
|
|
1041 |
is_expansion C basis \<Longrightarrow>
|
|
1042 |
(\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')) \<Longrightarrow> ms_exp_gt e (ms_aux_hd_exp xs) \<Longrightarrow>
|
|
1043 |
is_expansion_aux (MSLCons (C, e) xs) f (b # basis)"
|
|
1044 |
|
|
1045 |
inductive_cases is_expansion_aux_MSLCons: "is_expansion_aux (MSLCons (c, e) xs) f basis"
|
|
1046 |
|
|
1047 |
lemma is_expansion_aux_basis_nonempty: "is_expansion_aux F f basis \<Longrightarrow> basis \<noteq> []"
|
|
1048 |
by (erule is_expansion_aux.cases) auto
|
|
1049 |
|
|
1050 |
lemma is_expansion_aux_expansion_level:
|
|
1051 |
assumes "is_expansion_aux (G :: ('a::multiseries \<times> real) msllist) g basis"
|
|
1052 |
shows "length basis = Suc (expansion_level TYPE('a))"
|
|
1053 |
using assms by (cases rule: is_expansion_aux.cases) (auto dest: is_expansion_length)
|
|
1054 |
|
|
1055 |
lemma is_expansion_aux_imp_smallo:
|
|
1056 |
assumes "is_expansion_aux xs f basis" "ms_exp_gt p (ms_aux_hd_exp xs)"
|
|
1057 |
shows "f \<in> o(\<lambda>x. hd basis x powr p)"
|
|
1058 |
using assms
|
|
1059 |
proof (cases rule: is_expansion_aux.cases)
|
|
1060 |
case is_expansion_MSLNil
|
|
1061 |
show ?thesis by (simp add: landau_o.small.in_cong[OF is_expansion_MSLNil(2)])
|
|
1062 |
next
|
|
1063 |
case (is_expansion_MSLCons xs C b e basis)
|
|
1064 |
with assms have "f \<in> o(\<lambda>x. b x powr p)"
|
|
1065 |
by (intro is_expansion_MSLCons) (simp_all add: ms_aux_hd_exp_def)
|
|
1066 |
thus ?thesis by (simp add: is_expansion_MSLCons)
|
|
1067 |
qed
|
|
1068 |
|
|
1069 |
lemma is_expansion_aux_imp_smallo':
|
|
1070 |
assumes "is_expansion_aux xs f basis"
|
|
1071 |
obtains e where "f \<in> o(\<lambda>x. hd basis x powr e)"
|
|
1072 |
proof -
|
|
1073 |
define e where "e = (case ms_aux_hd_exp xs of None \<Rightarrow> 0 | Some e \<Rightarrow> e + 1)"
|
|
1074 |
have "ms_exp_gt e (ms_aux_hd_exp xs)"
|
|
1075 |
by (auto simp add: e_def ms_aux_hd_exp_def split: msllist.splits)
|
|
1076 |
from assms this have "f \<in> o(\<lambda>x. hd basis x powr e)" by (rule is_expansion_aux_imp_smallo)
|
|
1077 |
from that[OF this] show ?thesis .
|
|
1078 |
qed
|
|
1079 |
|
|
1080 |
lemma is_expansion_aux_imp_smallo'':
|
|
1081 |
assumes "is_expansion_aux xs f basis" "ms_exp_gt e' (ms_aux_hd_exp xs)"
|
|
1082 |
obtains e where "e < e'" "f \<in> o(\<lambda>x. hd basis x powr e)"
|
|
1083 |
proof -
|
|
1084 |
define e where "e = (case ms_aux_hd_exp xs of None \<Rightarrow> e' - 1 | Some e \<Rightarrow> (e' + e) / 2)"
|
|
1085 |
from assms have "ms_exp_gt e (ms_aux_hd_exp xs)" "e < e'"
|
|
1086 |
by (cases xs; simp add: e_def field_simps)+
|
|
1087 |
from assms(1) this(1) have "f \<in> o(\<lambda>x. hd basis x powr e)" by (rule is_expansion_aux_imp_smallo)
|
|
1088 |
from that[OF \<open>e < e'\<close> this] show ?thesis .
|
|
1089 |
qed
|
|
1090 |
|
|
1091 |
definition trimmed_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> bool" where
|
|
1092 |
"trimmed_ms_aux xs = (case xs of MSLNil \<Rightarrow> False | MSLCons (C, _) _ \<Rightarrow> trimmed C)"
|
|
1093 |
|
|
1094 |
lemma not_trimmed_ms_aux_MSLNil [simp]: "\<not>trimmed_ms_aux MSLNil"
|
|
1095 |
by (simp add: trimmed_ms_aux_def)
|
|
1096 |
|
|
1097 |
lemma trimmed_ms_aux_MSLCons: "trimmed_ms_aux (MSLCons x xs) = trimmed (fst x)"
|
|
1098 |
by (simp add: trimmed_ms_aux_def split: prod.split)
|
|
1099 |
|
|
1100 |
lemma trimmed_ms_aux_imp_nz:
|
|
1101 |
assumes "basis_wf basis" "is_expansion_aux xs f basis" "trimmed_ms_aux xs"
|
|
1102 |
shows "eventually (\<lambda>x. f x \<noteq> 0) at_top"
|
|
1103 |
proof (cases xs rule: ms_aux_cases)
|
|
1104 |
case (MSLCons C e xs')
|
|
1105 |
from assms this obtain b basis' where *: "basis = b # basis'"
|
|
1106 |
"is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
1107 |
"ms_exp_gt e (ms_aux_hd_exp xs')" "is_expansion C basis'" "\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')"
|
|
1108 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
1109 |
from assms(1,3) * have nz: "eventually (\<lambda>x. eval C x \<noteq> 0) at_top"
|
|
1110 |
by (auto simp: trimmed_ms_aux_def MSLCons basis_wf_Cons
|
|
1111 |
intro: trimmed_imp_eventually_nz[of basis'])
|
|
1112 |
from assms(1) * have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1113 |
hence b_nz: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
1114 |
|
|
1115 |
from is_expansion_aux_imp_smallo''[OF *(2,3)]
|
|
1116 |
obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
|
|
1117 |
unfolding list.sel by blast
|
|
1118 |
note this(2)
|
|
1119 |
also have "\<dots> = o(\<lambda>x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric])
|
|
1120 |
also from assms * e' have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))"
|
|
1121 |
by (intro is_expansion_imp_smallomega[OF _ *(4)])
|
|
1122 |
(simp_all add: MSLCons basis_wf_Cons trimmed_ms_aux_def)
|
|
1123 |
hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
|
|
1124 |
by (intro landau_o.small_big_mult is_expansion_imp_smallomega)
|
|
1125 |
(simp_all add: smallomega_iff_smallo)
|
|
1126 |
finally have "(\<lambda>x. f x - eval C x * b x powr e + eval C x * b x powr e) \<in>
|
|
1127 |
\<Theta>(\<lambda>x. eval C x * b x powr e)"
|
|
1128 |
by (subst bigtheta_sym, subst landau_theta.plus_absorb1) simp_all
|
|
1129 |
hence theta: "f \<in> \<Theta>(\<lambda>x. eval C x * b x powr e)" by simp
|
|
1130 |
have "eventually (\<lambda>x. eval C x * b x powr e \<noteq> 0) at_top"
|
|
1131 |
using b_nz nz by eventually_elim simp
|
|
1132 |
thus ?thesis by (subst eventually_nonzero_bigtheta [OF theta])
|
|
1133 |
qed (insert assms, simp_all add: trimmed_ms_aux_def)
|
|
1134 |
|
|
1135 |
lemma is_expansion_aux_imp_smallomega:
|
|
1136 |
assumes "basis_wf basis" "is_expansion_aux xs f basis" "filterlim b' at_top at_top"
|
|
1137 |
"trimmed_ms_aux xs" "\<forall>g\<in>set basis. (\<lambda>x. ln (g x)) \<in> o(\<lambda>x. ln (b' x))" "r < 0"
|
|
1138 |
shows "f \<in> \<omega>(\<lambda>x. b' x powr r)"
|
|
1139 |
proof (cases xs rule: ms_aux_cases)
|
|
1140 |
case (MSLCons C e xs')
|
|
1141 |
from assms this obtain b basis' where *: "basis = b # basis'" "trimmed C"
|
|
1142 |
"is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
1143 |
"ms_exp_gt e (ms_aux_hd_exp xs')"
|
|
1144 |
"is_expansion C basis'" "\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')"
|
|
1145 |
by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def)
|
|
1146 |
from assms * have nz: "eventually (\<lambda>x. eval C x \<noteq> 0) at_top"
|
|
1147 |
by (intro trimmed_imp_eventually_nz[of basis']) (simp_all add: basis_wf_Cons)
|
|
1148 |
from assms(1) * have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1149 |
hence b_nz: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
1150 |
|
|
1151 |
from is_expansion_aux_imp_smallo''[OF *(3,4)]
|
|
1152 |
obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
|
|
1153 |
unfolding list.sel by blast
|
|
1154 |
note this(2)
|
|
1155 |
also have "\<dots> = o(\<lambda>x. b x powr (e' - e) * b x powr e)" by (simp add: powr_add [symmetric])
|
|
1156 |
also from assms * e' have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))"
|
|
1157 |
by (intro is_expansion_imp_smallomega[OF _ *(5)])
|
|
1158 |
(simp_all add: MSLCons basis_wf_Cons trimmed_ms_aux_def)
|
|
1159 |
hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
|
|
1160 |
by (intro landau_o.small_big_mult is_expansion_imp_smallomega)
|
|
1161 |
(simp_all add: smallomega_iff_smallo)
|
|
1162 |
finally have "(\<lambda>x. f x - eval C x * b x powr e + eval C x * b x powr e) \<in>
|
|
1163 |
\<Theta>(\<lambda>x. eval C x * b x powr e)"
|
|
1164 |
by (subst bigtheta_sym, subst landau_theta.plus_absorb1) simp_all
|
|
1165 |
hence theta: "f \<in> \<Theta>(\<lambda>x. eval C x * b x powr e)" by simp
|
|
1166 |
also from * assms e' have "(\<lambda>x. eval C x * b x powr e) \<in> \<omega>(\<lambda>x. b x powr (e' - e) * b x powr e)"
|
|
1167 |
by (intro landau_omega.small_big_mult is_expansion_imp_smallomega[of basis'])
|
|
1168 |
(simp_all add: basis_wf_Cons)
|
|
1169 |
also have "\<dots> = \<omega>(\<lambda>x. b x powr e')" by (simp add: powr_add [symmetric])
|
|
1170 |
also from *(1) assms have "(\<lambda>x. b x powr e') \<in> \<omega>(\<lambda>x. b' x powr r)"
|
|
1171 |
unfolding smallomega_iff_smallo by (intro ln_smallo_imp_flat') (simp_all add: basis_wf_Cons)
|
|
1172 |
finally show ?thesis .
|
|
1173 |
qed (insert assms, simp_all add: trimmed_ms_aux_def)
|
|
1174 |
|
|
1175 |
definition dominant_term_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> monom" where
|
|
1176 |
"dominant_term_ms_aux xs =
|
|
1177 |
(case xs of MSLNil \<Rightarrow> (0, replicate (Suc (expansion_level TYPE('a))) 0)
|
|
1178 |
| MSLCons (C, e) _ \<Rightarrow> case dominant_term C of (c, es) \<Rightarrow> (c, e # es))"
|
|
1179 |
|
|
1180 |
lemma dominant_term_ms_aux_MSLCons:
|
|
1181 |
"dominant_term_ms_aux (MSLCons (C, e) xs) = map_prod id (\<lambda>es. e # es) (dominant_term C)"
|
|
1182 |
by (simp add: dominant_term_ms_aux_def case_prod_unfold map_prod_def)
|
|
1183 |
|
|
1184 |
lemma length_dominant_term_ms_aux [simp]:
|
|
1185 |
"length (snd (dominant_term_ms_aux (xs :: ('a :: multiseries \<times> real) msllist))) =
|
|
1186 |
Suc (expansion_level TYPE('a))"
|
|
1187 |
proof (cases xs rule: ms_aux_cases)
|
|
1188 |
case (MSLCons C e xs')
|
|
1189 |
hence "length (snd (dominant_term_ms_aux xs)) = Suc (length (snd (dominant_term C)))"
|
|
1190 |
by (simp add: dominant_term_ms_aux_def split: prod.splits)
|
|
1191 |
also note length_dominant_term
|
|
1192 |
finally show ?thesis .
|
|
1193 |
qed (simp_all add: dominant_term_ms_aux_def split: msllist.splits prod.splits)
|
|
1194 |
|
|
1195 |
definition const_ms_aux :: "real \<Rightarrow> ('a :: multiseries \<times> real) msllist" where
|
|
1196 |
"const_ms_aux c = MSLCons (const_expansion c, 0) MSLNil"
|
|
1197 |
|
|
1198 |
definition uminus_ms_aux :: "('a :: uminus \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1199 |
"uminus_ms_aux xs = mslmap (map_prod uminus id) xs"
|
|
1200 |
|
|
1201 |
corec (friend) plus_ms_aux :: "('a :: plus \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1202 |
"plus_ms_aux xs ys =
|
|
1203 |
(case (xs, ys) of
|
|
1204 |
(MSLNil, MSLNil) \<Rightarrow> MSLNil
|
|
1205 |
| (MSLNil, MSLCons y ys) \<Rightarrow> MSLCons y ys
|
|
1206 |
| (MSLCons x xs, MSLNil) \<Rightarrow> MSLCons x xs
|
|
1207 |
| (MSLCons x xs, MSLCons y ys) \<Rightarrow>
|
|
1208 |
if snd x > snd y then MSLCons x (plus_ms_aux xs (MSLCons y ys))
|
|
1209 |
else if snd x < snd y then MSLCons y (plus_ms_aux (MSLCons x xs) ys)
|
|
1210 |
else MSLCons (fst x + fst y, snd x) (plus_ms_aux xs ys))"
|
|
1211 |
|
|
1212 |
context
|
|
1213 |
begin
|
|
1214 |
|
|
1215 |
friend_of_corec mslmap where
|
|
1216 |
"mslmap (f :: 'a \<Rightarrow> 'a) xs =
|
|
1217 |
(case xs of MSLNil \<Rightarrow> MSLNil | MSLCons x xs \<Rightarrow> MSLCons (f x) (mslmap f xs))"
|
|
1218 |
by (simp split: msllist.splits, transfer_prover)
|
|
1219 |
|
|
1220 |
corec (friend) times_ms_aux
|
|
1221 |
:: "('a :: {plus,times} \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1222 |
"times_ms_aux xs ys =
|
|
1223 |
(case (xs, ys) of
|
|
1224 |
(MSLNil, ys) \<Rightarrow> MSLNil
|
|
1225 |
| (xs, MSLNil) \<Rightarrow> MSLNil
|
|
1226 |
| (MSLCons x xs, MSLCons y ys) \<Rightarrow>
|
|
1227 |
MSLCons (fst x * fst y, snd x + snd y)
|
|
1228 |
(plus_ms_aux (mslmap (map_prod (( *) (fst x)) ((+) (snd x))) ys)
|
|
1229 |
(times_ms_aux xs (MSLCons y ys))))"
|
|
1230 |
|
|
1231 |
definition scale_shift_ms_aux :: "('a :: times \<times> real) \<Rightarrow> ('a \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1232 |
"scale_shift_ms_aux = (\<lambda>(a,b) xs. mslmap (map_prod (( *) a) ((+) b)) xs)"
|
|
1233 |
|
|
1234 |
lemma times_ms_aux_altdef:
|
|
1235 |
"times_ms_aux xs ys =
|
|
1236 |
(case (xs, ys) of
|
|
1237 |
(MSLNil, ys) \<Rightarrow> MSLNil
|
|
1238 |
| (xs, MSLNil) \<Rightarrow> MSLNil
|
|
1239 |
| (MSLCons x xs, MSLCons y ys) \<Rightarrow>
|
|
1240 |
MSLCons (fst x * fst y, snd x + snd y)
|
|
1241 |
(plus_ms_aux (scale_shift_ms_aux x ys) (times_ms_aux xs (MSLCons y ys))))"
|
|
1242 |
by (subst times_ms_aux.code) (simp_all add: scale_shift_ms_aux_def split: msllist.splits)
|
|
1243 |
end
|
|
1244 |
|
|
1245 |
corec powser_ms_aux :: "real msllist \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1246 |
"powser_ms_aux cs xs = (case cs of MSLNil \<Rightarrow> MSLNil | MSLCons c cs \<Rightarrow>
|
|
1247 |
MSLCons (const_expansion c, 0) (times_ms_aux xs (powser_ms_aux cs xs)))"
|
|
1248 |
|
|
1249 |
definition minus_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> _" where
|
|
1250 |
"minus_ms_aux xs ys = plus_ms_aux xs (uminus_ms_aux ys)"
|
|
1251 |
|
|
1252 |
lemma is_expansion_aux_coinduct [consumes 1, case_names is_expansion_aux]:
|
|
1253 |
fixes xs :: "('a :: multiseries \<times> real) msllist"
|
|
1254 |
assumes "X xs f basis" "(\<And>xs f basis. X xs f basis \<Longrightarrow>
|
|
1255 |
(xs = MSLNil \<and> (\<forall>\<^sub>F x in at_top. f x = 0) \<and> length basis = Suc (expansion_level TYPE('a))) \<or>
|
|
1256 |
(\<exists>xs' b e basis' C. xs = MSLCons (C, e) xs' \<and> basis = b # basis' \<and>
|
|
1257 |
(X xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')) \<and> is_expansion C basis' \<and>
|
|
1258 |
(\<forall>x>e. f \<in> o(\<lambda>xa. b xa powr x)) \<and> ms_exp_gt e (ms_aux_hd_exp xs')))"
|
|
1259 |
shows "is_expansion_aux xs f basis"
|
|
1260 |
using assms(1)
|
|
1261 |
proof (coinduction arbitrary: xs f)
|
|
1262 |
case (is_expansion_aux xs f)
|
|
1263 |
from assms(2)[OF is_expansion_aux] show ?case by blast
|
|
1264 |
qed
|
|
1265 |
|
|
1266 |
lemma is_expansion_aux_cong:
|
|
1267 |
assumes "is_expansion_aux F f basis"
|
|
1268 |
assumes "eventually (\<lambda>x. f x = g x) at_top"
|
|
1269 |
shows "is_expansion_aux F g basis"
|
|
1270 |
using assms
|
|
1271 |
proof (coinduction arbitrary: F f g rule: is_expansion_aux_coinduct)
|
|
1272 |
case (is_expansion_aux F f g)
|
|
1273 |
from this have ev: "eventually (\<lambda>x. g x = f x) at_top" by (simp add: eq_commute)
|
|
1274 |
from ev have [simp]: "eventually (\<lambda>x. g x = 0) at_top \<longleftrightarrow> eventually (\<lambda>x. f x = 0) at_top"
|
|
1275 |
by (rule eventually_subst')
|
|
1276 |
from ev have [simp]: "(\<forall>\<^sub>F x in at_top. h x = g x - h' x) \<longleftrightarrow>
|
|
1277 |
(\<forall>\<^sub>F x in at_top. h x = f x - h' x)" for h h'
|
|
1278 |
by (rule eventually_subst')
|
|
1279 |
note [simp] = landau_o.small.in_cong[OF ev]
|
|
1280 |
from is_expansion_aux(1) show ?case
|
|
1281 |
by (cases rule: is_expansion_aux.cases) force+
|
|
1282 |
qed
|
|
1283 |
|
|
1284 |
lemma scale_shift_ms_aux_conv_mslmap:
|
|
1285 |
"scale_shift_ms_aux x = mslmap (map_prod (( *) (fst x)) ((+) (snd x)))"
|
|
1286 |
by (rule ext) (simp add: scale_shift_ms_aux_def map_prod_def case_prod_unfold)
|
|
1287 |
|
|
1288 |
fun inverse_ms_aux :: "('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1289 |
"inverse_ms_aux (MSLCons x xs) =
|
|
1290 |
(let c' = inverse (fst x)
|
|
1291 |
in scale_shift_ms_aux (c', -snd x)
|
|
1292 |
(powser_ms_aux (msllist_of_msstream (mssalternate 1 (-1)))
|
|
1293 |
(scale_shift_ms_aux (c', -snd x) xs)))"
|
|
1294 |
|
|
1295 |
(* TODO: Move? *)
|
|
1296 |
lemma inverse_prod_list: "inverse (prod_list xs :: 'a :: field) = prod_list (map inverse xs)"
|
|
1297 |
by (induction xs) simp_all
|
|
1298 |
|
|
1299 |
lemma eval_inverse_monom:
|
|
1300 |
"eval_monom (inverse_monom monom) basis = (\<lambda>x. inverse (eval_monom monom basis x))"
|
|
1301 |
by (rule ext) (simp add: eval_monom_def inverse_monom_def zip_map2 o_def case_prod_unfold
|
|
1302 |
inverse_prod_list powr_minus)
|
|
1303 |
|
|
1304 |
fun powr_ms_aux :: "bool \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> real \<Rightarrow> ('a \<times> real) msllist" where
|
|
1305 |
"powr_ms_aux abort (MSLCons x xs) p =
|
|
1306 |
scale_shift_ms_aux (powr_expansion abort (fst x) p, snd x * p)
|
|
1307 |
(powser_ms_aux (gbinomial_series abort p)
|
|
1308 |
(scale_shift_ms_aux (inverse (fst x), -snd x) xs))"
|
|
1309 |
|
|
1310 |
fun power_ms_aux :: "bool \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> nat \<Rightarrow> ('a \<times> real) msllist" where
|
|
1311 |
"power_ms_aux abort (MSLCons x xs) n =
|
|
1312 |
scale_shift_ms_aux (power_expansion abort (fst x) n, snd x * real n)
|
|
1313 |
(powser_ms_aux (gbinomial_series abort (real n))
|
|
1314 |
(scale_shift_ms_aux (inverse (fst x), -snd x) xs))"
|
|
1315 |
|
|
1316 |
definition sin_ms_aux' :: "('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1317 |
"sin_ms_aux' xs = times_ms_aux xs (powser_ms_aux (msllist_of_msstream sin_series_stream)
|
|
1318 |
(times_ms_aux xs xs))"
|
|
1319 |
|
|
1320 |
definition cos_ms_aux' :: "('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
1321 |
"cos_ms_aux' xs = powser_ms_aux (msllist_of_msstream cos_series_stream) (times_ms_aux xs xs)"
|
|
1322 |
|
|
1323 |
subsubsection \<open>Basic properties of multiseries operations\<close>
|
|
1324 |
|
|
1325 |
lemma uminus_ms_aux_MSLNil [simp]: "uminus_ms_aux MSLNil = MSLNil"
|
|
1326 |
by (simp add: uminus_ms_aux_def)
|
|
1327 |
|
|
1328 |
lemma uminus_ms_aux_MSLCons: "uminus_ms_aux (MSLCons x xs) = MSLCons (-fst x, snd x) (uminus_ms_aux xs)"
|
|
1329 |
by (simp add: uminus_ms_aux_def map_prod_def split: prod.splits)
|
|
1330 |
|
|
1331 |
lemma uminus_ms_aux_MSLNil_iff [simp]: "uminus_ms_aux xs = MSLNil \<longleftrightarrow> xs = MSLNil"
|
|
1332 |
by (simp add: uminus_ms_aux_def)
|
|
1333 |
|
|
1334 |
lemma hd_exp_uminus [simp]: "ms_aux_hd_exp (uminus_ms_aux xs) = ms_aux_hd_exp xs"
|
|
1335 |
by (simp add: uminus_ms_aux_def ms_aux_hd_exp_def split: msllist.splits prod.split)
|
|
1336 |
|
|
1337 |
lemma scale_shift_ms_aux_MSLNil_iff [simp]: "scale_shift_ms_aux x F = MSLNil \<longleftrightarrow> F = MSLNil"
|
|
1338 |
by (auto simp: scale_shift_ms_aux_def split: prod.split)
|
|
1339 |
|
|
1340 |
lemma scale_shift_ms_aux_MSLNil [simp]: "scale_shift_ms_aux x MSLNil = MSLNil"
|
|
1341 |
by (simp add: scale_shift_ms_aux_def split: prod.split)
|
|
1342 |
|
|
1343 |
lemma scale_shift_ms_aux_1 [simp]:
|
|
1344 |
"scale_shift_ms_aux (1 :: real, 0) xs = xs"
|
|
1345 |
by (simp add: scale_shift_ms_aux_def map_prod_def msllist.map_id)
|
|
1346 |
|
|
1347 |
lemma scale_shift_ms_aux_MSLCons:
|
|
1348 |
"scale_shift_ms_aux x (MSLCons y F) = MSLCons (fst x * fst y, snd x + snd y) (scale_shift_ms_aux x F)"
|
|
1349 |
by (auto simp: scale_shift_ms_aux_def map_prod_def split: prod.split)
|
|
1350 |
|
|
1351 |
lemma hd_exp_basis:
|
|
1352 |
"ms_aux_hd_exp (scale_shift_ms_aux x xs) = map_option ((+) (snd x)) (ms_aux_hd_exp xs)"
|
|
1353 |
by (auto simp: ms_aux_hd_exp_def scale_shift_ms_aux_MSLCons split: msllist.split)
|
|
1354 |
|
|
1355 |
lemma plus_ms_aux_MSLNil_iff: "plus_ms_aux F G = MSLNil \<longleftrightarrow> F = MSLNil \<and> G = MSLNil"
|
|
1356 |
by (subst plus_ms_aux.code) (simp split: msllist.splits)
|
|
1357 |
|
|
1358 |
lemma plus_ms_aux_MSLNil [simp]: "plus_ms_aux F MSLNil = F" "plus_ms_aux MSLNil G = G"
|
|
1359 |
by (subst plus_ms_aux.code, simp split: msllist.splits)+
|
|
1360 |
|
|
1361 |
lemma plus_ms_aux_MSLCons:
|
|
1362 |
"plus_ms_aux (MSLCons x xs) (MSLCons y ys) =
|
|
1363 |
(if snd x > snd y then MSLCons x (plus_ms_aux xs (MSLCons y ys))
|
|
1364 |
else if snd x < snd y then MSLCons y (plus_ms_aux (MSLCons x xs) ys)
|
|
1365 |
else MSLCons (fst x + fst y, snd x) (plus_ms_aux xs ys))"
|
|
1366 |
by (subst plus_ms_aux.code) simp
|
|
1367 |
|
|
1368 |
lemma hd_exp_plus:
|
|
1369 |
"ms_aux_hd_exp (plus_ms_aux xs ys) = combine_options max (ms_aux_hd_exp xs) (ms_aux_hd_exp ys)"
|
|
1370 |
by (cases xs; cases ys)
|
|
1371 |
(simp_all add: plus_ms_aux_MSLCons ms_aux_hd_exp_def max_def split: prod.split)
|
|
1372 |
|
|
1373 |
lemma minus_ms_aux_MSLNil_left [simp]: "minus_ms_aux MSLNil ys = uminus_ms_aux ys"
|
|
1374 |
by (simp add: minus_ms_aux_def)
|
|
1375 |
|
|
1376 |
lemma minus_ms_aux_MSLNil_right [simp]: "minus_ms_aux xs MSLNil = xs"
|
|
1377 |
by (simp add: minus_ms_aux_def)
|
|
1378 |
|
|
1379 |
lemma times_ms_aux_MSLNil_iff: "times_ms_aux F G = MSLNil \<longleftrightarrow> F = MSLNil \<or> G = MSLNil"
|
|
1380 |
by (subst times_ms_aux.code) (simp split: msllist.splits)
|
|
1381 |
|
|
1382 |
lemma times_ms_aux_MSLNil [simp]: "times_ms_aux MSLNil G = MSLNil" "times_ms_aux F MSLNil = MSLNil"
|
|
1383 |
by (subst times_ms_aux.code, simp split: msllist.splits)+
|
|
1384 |
|
|
1385 |
lemma times_ms_aux_MSLCons: "times_ms_aux (MSLCons x xs) (MSLCons y ys) =
|
|
1386 |
MSLCons (fst x * fst y, snd x + snd y) (plus_ms_aux (scale_shift_ms_aux x ys)
|
|
1387 |
(times_ms_aux xs (MSLCons y ys)))"
|
|
1388 |
by (subst times_ms_aux_altdef) simp
|
|
1389 |
|
|
1390 |
lemma hd_exp_times:
|
|
1391 |
"ms_aux_hd_exp (times_ms_aux xs ys) =
|
|
1392 |
(case (ms_aux_hd_exp xs, ms_aux_hd_exp ys) of (Some x, Some y) \<Rightarrow> Some (x + y) | _ \<Rightarrow> None)"
|
|
1393 |
by (cases xs; cases ys) (simp_all add: times_ms_aux_MSLCons ms_aux_hd_exp_def split: prod.split)
|
|
1394 |
|
|
1395 |
|
|
1396 |
subsubsection \<open>Induction upto friends for multiseries\<close>
|
|
1397 |
|
|
1398 |
inductive ms_closure ::
|
|
1399 |
"(('a :: multiseries \<times> real) msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> basis \<Rightarrow> bool) \<Rightarrow>
|
|
1400 |
('a :: multiseries \<times> real) msllist \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> basis \<Rightarrow> bool" for P where
|
|
1401 |
ms_closure_embed:
|
|
1402 |
"P xs f basis \<Longrightarrow> ms_closure P xs f basis"
|
|
1403 |
| ms_closure_cong:
|
|
1404 |
"ms_closure P xs f basis \<Longrightarrow> eventually (\<lambda>x. f x = g x) at_top \<Longrightarrow> ms_closure P xs g basis"
|
|
1405 |
| ms_closure_MSLCons:
|
|
1406 |
"ms_closure P xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis) \<Longrightarrow>
|
|
1407 |
is_expansion C basis \<Longrightarrow> (\<And>e'. e' > e \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')) \<Longrightarrow>
|
|
1408 |
ms_exp_gt e (ms_aux_hd_exp xs) \<Longrightarrow>
|
|
1409 |
ms_closure P (MSLCons (C, e) xs) f (b # basis)"
|
|
1410 |
| ms_closure_add:
|
|
1411 |
"ms_closure P xs f basis \<Longrightarrow> ms_closure P ys g basis \<Longrightarrow>
|
|
1412 |
ms_closure P (plus_ms_aux xs ys) (\<lambda>x. f x + g x) basis"
|
|
1413 |
| ms_closure_mult:
|
|
1414 |
"ms_closure P xs f basis \<Longrightarrow> ms_closure P ys g basis \<Longrightarrow>
|
|
1415 |
ms_closure P (times_ms_aux xs ys) (\<lambda>x. f x * g x) basis"
|
|
1416 |
| ms_closure_basis:
|
|
1417 |
"ms_closure P xs f (b # basis) \<Longrightarrow> is_expansion C basis \<Longrightarrow>
|
|
1418 |
ms_closure P (scale_shift_ms_aux (C,e) xs) (\<lambda>x. eval C x * b x powr e * f x) (b # basis)"
|
|
1419 |
| ms_closure_embed':
|
|
1420 |
"is_expansion_aux xs f basis \<Longrightarrow> ms_closure P xs f basis"
|
|
1421 |
|
|
1422 |
lemma is_expansion_aux_coinduct_upto [consumes 2, case_names A B]:
|
|
1423 |
fixes xs :: "('a :: multiseries \<times> real) msllist"
|
|
1424 |
assumes *: "X xs f basis" and basis: "basis_wf basis"
|
|
1425 |
and step: "\<And>xs f basis. X xs f basis \<Longrightarrow>
|
|
1426 |
(xs = MSLNil \<and> eventually (\<lambda>x. f x = 0) at_top \<and> length basis = Suc (expansion_level TYPE('a))) \<or>
|
|
1427 |
(\<exists>xs' b e basis' C. xs = MSLCons (C, e) xs' \<and> basis = b # basis' \<and>
|
|
1428 |
ms_closure X xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis') \<and>
|
|
1429 |
is_expansion C basis' \<and> (\<forall>e'>e. f \<in> o(\<lambda>x. b x powr e')) \<and> ms_exp_gt e (ms_aux_hd_exp xs'))"
|
|
1430 |
shows "is_expansion_aux xs f basis"
|
|
1431 |
proof -
|
|
1432 |
from * have "ms_closure X xs f basis" by (rule ms_closure_embed[of X xs f basis])
|
|
1433 |
thus ?thesis
|
|
1434 |
proof (coinduction arbitrary: xs f rule: is_expansion_aux_coinduct)
|
|
1435 |
case (is_expansion_aux xs f)
|
|
1436 |
from this and basis show ?case
|
|
1437 |
proof (induction rule: ms_closure.induct)
|
|
1438 |
case (ms_closure_embed xs f basis)
|
|
1439 |
from step[OF ms_closure_embed(1)] show ?case by auto
|
|
1440 |
next
|
|
1441 |
case (ms_closure_MSLCons xs f C b e basis)
|
|
1442 |
thus ?case by auto
|
|
1443 |
next
|
|
1444 |
case (ms_closure_cong xs f basis g)
|
|
1445 |
note ev = \<open>\<forall>\<^sub>F x in at_top. f x = g x\<close>
|
|
1446 |
hence ev_zero_iff: "eventually (\<lambda>x. f x = 0) at_top \<longleftrightarrow> eventually (\<lambda>x. g x = 0) at_top"
|
|
1447 |
by (rule eventually_subst')
|
|
1448 |
have *: "ms_closure X xs' (\<lambda>x. f x - c x * b x powr e) basis \<longleftrightarrow>
|
|
1449 |
ms_closure X xs' (\<lambda>x. g x - c x * b x powr e) basis" for xs' c b e
|
|
1450 |
by (rule iffI; erule ms_closure.intros(2)) (insert ev, auto elim: eventually_mono)
|
|
1451 |
from ms_closure_cong show ?case
|
|
1452 |
by (simp add: ev_zero_iff * landau_o.small.in_cong[OF ev])
|
|
1453 |
next
|
|
1454 |
case (ms_closure_embed' xs f basis)
|
|
1455 |
from this show ?case
|
|
1456 |
by (cases rule: is_expansion_aux.cases) (force intro: ms_closure.intros(7))+
|
|
1457 |
next
|
|
1458 |
|
|
1459 |
case (ms_closure_basis xs f b basis C e)
|
|
1460 |
show ?case
|
|
1461 |
proof (cases xs rule: ms_aux_cases)
|
|
1462 |
case MSLNil
|
|
1463 |
with ms_closure_basis show ?thesis
|
|
1464 |
by (auto simp: scale_shift_ms_aux_def split: prod.split elim: eventually_mono)
|
|
1465 |
next
|
|
1466 |
case (MSLCons C' e' xs')
|
|
1467 |
with ms_closure_basis have IH:
|
|
1468 |
"ms_closure X (MSLCons (C', e') xs') f (b # basis)"
|
|
1469 |
"is_expansion C basis" "xs = MSLCons (C', e') xs'"
|
|
1470 |
"ms_closure X xs' (\<lambda>x. f x - eval C' x * b x powr e') (b # basis)"
|
|
1471 |
"ms_exp_gt e' (ms_aux_hd_exp xs')"
|
|
1472 |
"is_expansion C' basis" "\<And>x. x > e' \<Longrightarrow> f \<in> o(\<lambda>xa. b xa powr x)"
|
|
1473 |
by auto
|
|
1474 |
|
|
1475 |
{
|
|
1476 |
fix e'' :: real assume *: "e + e' < e''"
|
|
1477 |
define d where "d = (e'' - e - e') / 2"
|
|
1478 |
from * have "d > 0" by (simp add: d_def)
|
|
1479 |
hence "(\<lambda>x. eval C x * b x powr e * f x) \<in> o(\<lambda>x. b x powr d * b x powr e * b x powr (e'+d))"
|
|
1480 |
using ms_closure_basis(4) IH
|
|
1481 |
by (intro landau_o.small.mult[OF landau_o.small_big_mult] IH
|
|
1482 |
is_expansion_imp_smallo[OF _ IH(2)]) (simp_all add: basis_wf_Cons)
|
|
1483 |
also have "\<dots> = o(\<lambda>x. b x powr e'')" by (simp add: d_def powr_add [symmetric])
|
|
1484 |
finally have "(\<lambda>x. eval C x * b x powr e * f x) \<in> \<dots>" .
|
|
1485 |
}
|
|
1486 |
moreover have "ms_closure X xs' (\<lambda>x. f x - eval C' x * b x powr e') (b # basis)"
|
|
1487 |
using ms_closure_basis IH by auto
|
|
1488 |
note ms_closure.intros(6)[OF this IH(2), of e]
|
|
1489 |
moreover have "ms_exp_gt (e + e') (ms_aux_hd_exp (scale_shift_ms_aux (C, e) xs'))"
|
|
1490 |
using \<open>ms_exp_gt e' (ms_aux_hd_exp xs')\<close>
|
|
1491 |
by (cases xs') (simp_all add: hd_exp_basis scale_shift_ms_aux_def )
|
|
1492 |
ultimately show ?thesis using IH(2,6) MSLCons ms_closure_basis(4)
|
|
1493 |
by (auto simp: scale_shift_ms_aux_MSLCons algebra_simps powr_add basis_wf_Cons
|
|
1494 |
intro: is_expansion_mult)
|
|
1495 |
qed
|
|
1496 |
|
|
1497 |
next
|
|
1498 |
case (ms_closure_add xs f basis ys g)
|
|
1499 |
show ?case
|
|
1500 |
proof (cases xs ys rule: ms_aux_cases[case_product ms_aux_cases])
|
|
1501 |
case MSLNil_MSLNil
|
|
1502 |
with ms_closure_add
|
|
1503 |
have "eventually (\<lambda>x. f x = 0) at_top" "eventually (\<lambda>x. g x = 0) at_top"
|
|
1504 |
by simp_all
|
|
1505 |
hence "eventually (\<lambda>x. f x + g x = 0) at_top" by eventually_elim simp
|
|
1506 |
with MSLNil_MSLNil ms_closure_add show ?thesis by simp
|
|
1507 |
next
|
|
1508 |
case (MSLNil_MSLCons c e xs')
|
|
1509 |
with ms_closure_add have "eventually (\<lambda>x. f x = 0) at_top" by simp
|
|
1510 |
hence ev: "eventually (\<lambda>x. f x + g x = g x) at_top" by eventually_elim simp
|
|
1511 |
have *: "ms_closure X xs (\<lambda>x. f x + g x - y x) basis \<longleftrightarrow>
|
|
1512 |
ms_closure X xs (\<lambda>x. g x - y x) basis" for y basis xs
|
|
1513 |
by (rule iffI; erule ms_closure_cong) (insert ev, simp_all)
|
|
1514 |
from MSLNil_MSLCons ms_closure_add
|
|
1515 |
show ?thesis by (simp add: * landau_o.small.in_cong[OF ev])
|
|
1516 |
next
|
|
1517 |
case (MSLCons_MSLNil c e xs')
|
|
1518 |
with ms_closure_add have "eventually (\<lambda>x. g x = 0) at_top" by simp
|
|
1519 |
hence ev: "eventually (\<lambda>x. f x + g x = f x) at_top" by eventually_elim simp
|
|
1520 |
have *: "ms_closure X xs (\<lambda>x. f x + g x - y x) basis \<longleftrightarrow>
|
|
1521 |
ms_closure X xs (\<lambda>x. f x - y x) basis" for y basis xs
|
|
1522 |
by (rule iffI; erule ms_closure_cong) (insert ev, simp_all)
|
|
1523 |
from MSLCons_MSLNil ms_closure_add
|
|
1524 |
show ?thesis by (simp add: * landau_o.small.in_cong[OF ev])
|
|
1525 |
next
|
|
1526 |
case (MSLCons_MSLCons C1 e1 xs' C2 e2 ys')
|
|
1527 |
with ms_closure_add obtain b basis' where IH:
|
|
1528 |
"basis_wf (b # basis')" "basis = b # basis'"
|
|
1529 |
"ms_closure X (MSLCons (C1, e1) xs') f (b # basis')"
|
|
1530 |
"ms_closure X (MSLCons (C2, e2) ys') g (b # basis')"
|
|
1531 |
"xs = MSLCons (C1, e1) xs'" "ys = MSLCons (C2, e2) ys'"
|
|
1532 |
"ms_closure X xs' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
|
|
1533 |
"ms_closure X ys' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
|
|
1534 |
"is_expansion C1 basis'" "\<And>e. e > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e)"
|
|
1535 |
"is_expansion C2 basis'" "\<And>e. e > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e)"
|
|
1536 |
"ms_exp_gt e1 (ms_aux_hd_exp xs')" "ms_exp_gt e2 (ms_aux_hd_exp ys')"
|
|
1537 |
by blast
|
|
1538 |
have o: "(\<lambda>x. f x + g x) \<in> o(\<lambda>x. b x powr e)" if "e > max e1 e2" for e
|
|
1539 |
using that by (intro sum_in_smallo IH) simp_all
|
|
1540 |
|
|
1541 |
show ?thesis
|
|
1542 |
proof (cases e1 e2 rule: linorder_cases)
|
|
1543 |
case less
|
|
1544 |
have gt: "ms_exp_gt e2 (combine_options max (Some e1) (ms_aux_hd_exp ys'))"
|
|
1545 |
using \<open>ms_exp_gt e2 _\<close> less by (cases "ms_aux_hd_exp ys'") auto
|
|
1546 |
have "ms_closure X (plus_ms_aux xs ys')
|
|
1547 |
(\<lambda>x. f x + (g x - eval C2 x * b x powr e2)) (b # basis')"
|
|
1548 |
by (rule ms_closure.intros(4)) (insert IH, simp_all)
|
|
1549 |
with less IH(2,11,14) o gt show ?thesis
|
|
1550 |
by (intro disjI2) (simp add: MSLCons_MSLCons plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
|
|
1551 |
next
|
|
1552 |
case greater
|
|
1553 |
have gt: "ms_exp_gt e1 (combine_options max (ms_aux_hd_exp xs') (Some e2))"
|
|
1554 |
using \<open>ms_exp_gt e1 _\<close> greater by (cases "ms_aux_hd_exp xs'") auto
|
|
1555 |
have "ms_closure X (plus_ms_aux xs' ys)
|
|
1556 |
(\<lambda>x. (f x - eval C1 x * b x powr e1) + g x) (b # basis')"
|
|
1557 |
by (rule ms_closure.intros(4)) (insert IH, simp_all)
|
|
1558 |
with greater IH(2,9,13) o gt show ?thesis
|
|
1559 |
by (intro disjI2) (simp add: MSLCons_MSLCons plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
|
|
1560 |
next
|
|
1561 |
case equal
|
|
1562 |
have gt: "ms_exp_gt e2 (combine_options max (ms_aux_hd_exp xs')
|
|
1563 |
(ms_aux_hd_exp ys'))"
|
|
1564 |
using \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close> equal
|
|
1565 |
by (cases "ms_aux_hd_exp xs'"; cases "ms_aux_hd_exp ys'") auto
|
|
1566 |
have "ms_closure X (plus_ms_aux xs' ys')
|
|
1567 |
(\<lambda>x. (f x - eval C1 x * b x powr e1) + (g x - eval C2 x * b x powr e2)) (b # basis')"
|
|
1568 |
by (rule ms_closure.intros(4)) (insert IH, simp_all)
|
|
1569 |
with equal IH(1,2,9,11,13,14) o gt show ?thesis
|
|
1570 |
by (intro disjI2)
|
|
1571 |
(auto intro: is_expansion_add
|
|
1572 |
simp: plus_ms_aux_MSLCons basis_wf_Cons algebra_simps hd_exp_plus MSLCons_MSLCons)
|
|
1573 |
qed
|
|
1574 |
qed
|
|
1575 |
|
|
1576 |
next
|
|
1577 |
|
|
1578 |
case (ms_closure_mult xs f basis ys g)
|
|
1579 |
show ?case
|
|
1580 |
proof (cases "xs = MSLNil \<or> ys = MSLNil")
|
|
1581 |
case True
|
|
1582 |
with ms_closure_mult
|
|
1583 |
have "eventually (\<lambda>x. f x = 0) at_top \<or> eventually (\<lambda>x. g x = 0) at_top" by blast
|
|
1584 |
hence "eventually (\<lambda>x. f x * g x = 0) at_top" by (auto elim: eventually_mono)
|
|
1585 |
with ms_closure_mult True show ?thesis by auto
|
|
1586 |
next
|
|
1587 |
case False
|
|
1588 |
with ms_closure_mult obtain C1 e1 xs' C2 e2 ys' b basis' where IH:
|
|
1589 |
"xs = MSLCons (C1, e1) xs'" "ys = MSLCons (C2, e2) ys'"
|
|
1590 |
"basis_wf (b # basis')" "basis = b # basis'"
|
|
1591 |
"ms_closure X (MSLCons (C1, e1) xs') f (b # basis')"
|
|
1592 |
"ms_closure X (MSLCons (C2, e2) ys') g (b # basis')"
|
|
1593 |
"ms_closure X xs' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
|
|
1594 |
"ms_closure X ys' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
|
|
1595 |
"is_expansion C1 basis'" "\<And>e. e > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e)"
|
|
1596 |
"is_expansion C2 basis'" "\<And>e. e > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e)"
|
|
1597 |
"ms_exp_gt e1 (ms_aux_hd_exp xs')" "ms_exp_gt e2 (ms_aux_hd_exp ys')"
|
|
1598 |
by blast
|
|
1599 |
have o: "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr e)" if "e > e1 + e2" for e
|
|
1600 |
proof -
|
|
1601 |
define d where "d = (e - e1 - e2) / 2"
|
|
1602 |
from that have d: "d > 0" by (simp add: d_def)
|
|
1603 |
hence "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr (e1 + d) * b x powr (e2 + d))"
|
|
1604 |
by (intro landau_o.small.mult IH) simp_all
|
|
1605 |
also have "\<dots> = o(\<lambda>x. b x powr e)" by (simp add: d_def powr_add [symmetric])
|
|
1606 |
finally show ?thesis .
|
|
1607 |
qed
|
|
1608 |
|
|
1609 |
have "ms_closure X (plus_ms_aux (scale_shift_ms_aux (C1, e1) ys') (times_ms_aux xs' ys))
|
|
1610 |
(\<lambda>x. eval C1 x * b x powr e1 * (g x - eval C2 x * b x powr e2) +
|
|
1611 |
((f x - eval C1 x * b x powr e1) * g x)) (b # basis')"
|
|
1612 |
(is "ms_closure _ ?zs ?f _") using ms_closure_mult IH(4)
|
|
1613 |
by (intro ms_closure.intros(4-6) IH) simp_all
|
|
1614 |
also have "?f = (\<lambda>x. f x * g x - eval C1 x * eval C2 x * b x powr (e1 + e2))"
|
|
1615 |
by (intro ext) (simp add: algebra_simps powr_add)
|
|
1616 |
finally have "ms_closure X ?zs \<dots> (b # basis')" .
|
|
1617 |
moreover from \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close>
|
|
1618 |
have "ms_exp_gt (e1 + e2) (combine_options max (map_option ((+) e1)
|
|
1619 |
(ms_aux_hd_exp ys')) (case ms_aux_hd_exp xs' of None \<Rightarrow> None | Some x \<Rightarrow>
|
|
1620 |
(case Some e2 of None \<Rightarrow> None | Some y \<Rightarrow> Some (x + y))))"
|
|
1621 |
by (cases "ms_aux_hd_exp xs'"; cases "ms_aux_hd_exp ys'")
|
|
1622 |
(simp_all add: hd_exp_times hd_exp_plus hd_exp_basis IH)
|
|
1623 |
ultimately show ?thesis using IH(1,2,3,4,9,11) o
|
|
1624 |
by (auto simp: times_ms_aux_MSLCons basis_wf_Cons hd_exp_times hd_exp_plus hd_exp_basis
|
|
1625 |
intro: is_expansion_mult)
|
|
1626 |
qed
|
|
1627 |
qed
|
|
1628 |
qed
|
|
1629 |
qed
|
|
1630 |
|
|
1631 |
|
|
1632 |
|
|
1633 |
subsubsection \<open>Dominant terms\<close>
|
|
1634 |
|
|
1635 |
lemma one_smallo_powr: "e > (0::real) \<Longrightarrow> (\<lambda>_. 1) \<in> o(\<lambda>x. x powr e)"
|
|
1636 |
by (auto simp: smallomega_iff_smallo [symmetric] real_powr_at_top
|
|
1637 |
smallomegaI_filterlim_at_top_norm)
|
|
1638 |
|
|
1639 |
lemma powr_smallo_one: "e < (0::real) \<Longrightarrow> (\<lambda>x. x powr e) \<in> o(\<lambda>_. 1)"
|
|
1640 |
by (intro smalloI_tendsto) (auto intro!: tendsto_neg_powr filterlim_ident)
|
|
1641 |
|
|
1642 |
lemma eval_monom_smallo':
|
|
1643 |
assumes "basis_wf (b # basis)" "e > 0"
|
|
1644 |
shows "eval_monom x basis \<in> o(\<lambda>x. b x powr e)"
|
|
1645 |
proof (cases x)
|
|
1646 |
case (Pair c es)
|
|
1647 |
from assms show ?thesis unfolding Pair
|
|
1648 |
proof (induction es arbitrary: b basis e)
|
|
1649 |
case (Nil b basis e)
|
|
1650 |
hence b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1651 |
have "eval_monom (c, []) basis \<in> O(\<lambda>_. 1)" by simp
|
|
1652 |
also have "(\<lambda>_. 1) \<in> o(\<lambda>x. b x powr e)"
|
|
1653 |
by (rule landau_o.small.compose[OF _ b]) (insert Nil, simp add: one_smallo_powr)
|
|
1654 |
finally show ?case .
|
|
1655 |
next
|
|
1656 |
case (Cons e' es b basis e)
|
|
1657 |
show ?case
|
|
1658 |
proof (cases basis)
|
|
1659 |
case Nil
|
|
1660 |
from Cons have b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1661 |
from Nil have "eval_monom (c, e' # es) basis \<in> O(\<lambda>_. 1)" by simp
|
|
1662 |
also have "(\<lambda>_. 1) \<in> o(\<lambda>x. b x powr e)"
|
|
1663 |
by (rule landau_o.small.compose[OF _ b]) (insert Cons.prems, simp add: one_smallo_powr)
|
|
1664 |
finally show ?thesis .
|
|
1665 |
next
|
|
1666 |
case (Cons b' basis')
|
|
1667 |
from Cons.prems have "eval_monom (c, es) basis' \<in> o(\<lambda>x. b' x powr 1)"
|
|
1668 |
by (intro Cons.IH) (simp_all add: basis_wf_Cons Cons)
|
|
1669 |
hence "(\<lambda>x. eval_monom (c, es) basis' x * b' x powr e') \<in> o(\<lambda>x. b' x powr 1 * b' x powr e')"
|
|
1670 |
by (rule landau_o.small_big_mult) simp_all
|
|
1671 |
also have "\<dots> = o(\<lambda>x. b' x powr (1 + e'))" by (simp add: powr_add)
|
|
1672 |
also from Cons.prems have "(\<lambda>x. b' x powr (1 + e')) \<in> o(\<lambda>x. b x powr e)"
|
|
1673 |
by (intro ln_smallo_imp_flat) (simp_all add: basis_wf_Cons Cons)
|
|
1674 |
finally show ?thesis by (simp add: powr_add [symmetric] Cons eval_monom_Cons)
|
|
1675 |
qed
|
|
1676 |
qed
|
|
1677 |
qed
|
|
1678 |
|
|
1679 |
lemma eval_monom_smallomega':
|
|
1680 |
assumes "basis_wf (b # basis)" "e < 0"
|
|
1681 |
shows "eval_monom (1, es) basis \<in> \<omega>(\<lambda>x. b x powr e)"
|
|
1682 |
using assms
|
|
1683 |
proof (induction es arbitrary: b basis e)
|
|
1684 |
case (Nil b basis e)
|
|
1685 |
hence b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1686 |
have "eval_monom (1, []) basis \<in> \<Omega>(\<lambda>_. 1)" by simp
|
|
1687 |
also have "(\<lambda>_. 1) \<in> \<omega>(\<lambda>x. b x powr e)" unfolding smallomega_iff_smallo
|
|
1688 |
by (rule landau_o.small.compose[OF _ b]) (insert Nil, simp add: powr_smallo_one)
|
|
1689 |
finally show ?case .
|
|
1690 |
next
|
|
1691 |
case (Cons e' es b basis e)
|
|
1692 |
show ?case
|
|
1693 |
proof (cases basis)
|
|
1694 |
case Nil
|
|
1695 |
from Cons have b: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1696 |
from Nil have "eval_monom (1, e' # es) basis \<in> \<Omega>(\<lambda>_. 1)" by simp
|
|
1697 |
also have "(\<lambda>_. 1) \<in> \<omega>(\<lambda>x. b x powr e)" unfolding smallomega_iff_smallo
|
|
1698 |
by (rule landau_o.small.compose[OF _ b]) (insert Cons.prems, simp add: powr_smallo_one)
|
|
1699 |
finally show ?thesis .
|
|
1700 |
next
|
|
1701 |
case (Cons b' basis')
|
|
1702 |
from Cons.prems have "eval_monom (1, es) basis' \<in> \<omega>(\<lambda>x. b' x powr -1)"
|
|
1703 |
by (intro Cons.IH) (simp_all add: basis_wf_Cons Cons)
|
|
1704 |
hence "(\<lambda>x. eval_monom (1, es) basis' x * b' x powr e') \<in> \<omega>(\<lambda>x. b' x powr -1 * b' x powr e')"
|
|
1705 |
by (rule landau_omega.small_big_mult) simp_all
|
|
1706 |
also have "\<dots> = \<omega>(\<lambda>x. b' x powr (e' - 1))" by (simp add: powr_diff powr_minus field_simps)
|
|
1707 |
also from Cons.prems have "(\<lambda>x. b' x powr (e' - 1)) \<in> \<omega>(\<lambda>x. b x powr e)"
|
|
1708 |
unfolding smallomega_iff_smallo
|
|
1709 |
by (intro ln_smallo_imp_flat') (simp_all add: basis_wf_Cons Cons)
|
|
1710 |
finally show ?thesis by (simp add: powr_add [symmetric] Cons eval_monom_Cons)
|
|
1711 |
qed
|
|
1712 |
qed
|
|
1713 |
|
|
1714 |
lemma dominant_term_ms_aux:
|
|
1715 |
assumes basis: "basis_wf basis" and xs: "trimmed_ms_aux xs" "is_expansion_aux xs f basis"
|
|
1716 |
shows "f \<sim>[at_top] eval_monom (dominant_term_ms_aux xs) basis" (is ?thesis1)
|
|
1717 |
and "eventually (\<lambda>x. sgn (f x) = sgn (fst (dominant_term_ms_aux xs))) at_top" (is ?thesis2)
|
|
1718 |
proof -
|
|
1719 |
include asymp_equiv_notation
|
|
1720 |
from xs(2,1) obtain xs' C b e basis' where xs':
|
|
1721 |
"trimmed C" "xs = MSLCons (C, e) xs'" "basis = b # basis'"
|
|
1722 |
"is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
1723 |
"is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
|
|
1724 |
"ms_exp_gt e (ms_aux_hd_exp xs')"
|
|
1725 |
by (cases rule: is_expansion_aux.cases) (auto simp: trimmed_ms_aux_MSLCons)
|
|
1726 |
from is_expansion_aux_imp_smallo''[OF xs'(4,7)]
|
|
1727 |
obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
|
|
1728 |
unfolding list.sel by blast
|
|
1729 |
from basis xs' have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1730 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
1731 |
|
|
1732 |
note e'(2)
|
|
1733 |
also have "o(\<lambda>x. b x powr e') = o(\<lambda>x. b x powr (e' - e) * b x powr e)"
|
|
1734 |
by (simp add: powr_add [symmetric])
|
|
1735 |
also from xs' basis e'(1) have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))"
|
|
1736 |
by (intro is_expansion_imp_smallomega[of basis']) (simp_all add: basis_wf_Cons)
|
|
1737 |
hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
|
|
1738 |
by (intro landau_o.small_big_mult) (simp_all add: smallomega_iff_smallo)
|
|
1739 |
finally have *: "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)" .
|
|
1740 |
hence "f \<sim> (\<lambda>x. eval C x * b x powr e)" by (intro smallo_imp_asymp_equiv) simp
|
|
1741 |
also from basis xs' have "eval C \<sim> (\<lambda>x. eval_monom (dominant_term C) basis' x)"
|
|
1742 |
by (intro dominant_term) (simp_all add: basis_wf_Cons)
|
|
1743 |
also have "(\<lambda>a. eval_monom (dominant_term C) basis' a * b a powr e) =
|
|
1744 |
eval_monom (dominant_term_ms_aux xs) basis"
|
|
1745 |
by (auto simp add: dominant_term_ms_aux_def xs' eval_monom_Cons fun_eq_iff split: prod.split)
|
|
1746 |
finally show ?thesis1 by - (simp_all add: asymp_equiv_intros)
|
|
1747 |
|
|
1748 |
have "eventually (\<lambda>x. sgn (eval C x * b x powr e + (f x - eval C x * b x powr e)) =
|
|
1749 |
sgn (eval C x * b x powr e)) at_top"
|
|
1750 |
by (intro smallo_imp_eventually_sgn *)
|
|
1751 |
moreover have "eventually (\<lambda>x. sgn (eval C x) = sgn (fst (dominant_term C))) at_top"
|
|
1752 |
using xs' basis by (intro trimmed_imp_eventually_sgn[of basis']) (simp_all add: basis_wf_Cons)
|
|
1753 |
ultimately have "eventually (\<lambda>x. sgn (f x) = sgn (fst (dominant_term C))) at_top"
|
|
1754 |
using b_pos by eventually_elim (simp_all add: sgn_mult)
|
|
1755 |
thus ?thesis2 using xs'(2) by (auto simp: dominant_term_ms_aux_def split: prod.split)
|
|
1756 |
qed
|
|
1757 |
|
|
1758 |
lemma eval_monom_smallo:
|
|
1759 |
assumes "basis \<noteq> []" "basis_wf basis" "length es = length basis" "e > hd es"
|
|
1760 |
shows "eval_monom (c, es) basis \<in> o(\<lambda>x. hd basis x powr e)"
|
|
1761 |
proof (cases es; cases basis)
|
|
1762 |
fix b basis' e' es' assume [simp]: "basis = b # basis'" "es = e' # es'"
|
|
1763 |
from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1764 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top"
|
|
1765 |
using filterlim_at_top_dense by blast
|
|
1766 |
have "(\<lambda>x. eval_monom (1, es') basis' x * b x powr e') \<in> o(\<lambda>x. b x powr (e - e') * b x powr e')"
|
|
1767 |
using assms by (intro landau_o.small_big_mult eval_monom_smallo') auto
|
|
1768 |
also have "(\<lambda>x. b x powr (e - e') * b x powr e') \<in> \<Theta>(\<lambda>x. b x powr e)"
|
|
1769 |
by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff)
|
|
1770 |
finally show ?thesis by (simp add: eval_monom_def algebra_simps)
|
|
1771 |
qed (insert assms, auto)
|
|
1772 |
|
|
1773 |
lemma eval_monom_smallomega:
|
|
1774 |
assumes "basis \<noteq> []" "basis_wf basis" "length es = length basis" "e < hd es"
|
|
1775 |
shows "eval_monom (1, es) basis \<in> \<omega>(\<lambda>x. hd basis x powr e)"
|
|
1776 |
proof (cases es; cases basis)
|
|
1777 |
fix b basis' e' es' assume [simp]: "basis = b # basis'" "es = e' # es'"
|
|
1778 |
from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1779 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top"
|
|
1780 |
using filterlim_at_top_dense by blast
|
|
1781 |
have "(\<lambda>x. eval_monom (1, es') basis' x * b x powr e') \<in> \<omega>(\<lambda>x. b x powr (e - e') * b x powr e')"
|
|
1782 |
using assms by (intro landau_omega.small_big_mult eval_monom_smallomega') auto
|
|
1783 |
also have "(\<lambda>x. b x powr (e - e') * b x powr e') \<in> \<Theta>(\<lambda>x. b x powr e)"
|
|
1784 |
by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff)
|
|
1785 |
finally show ?thesis by (simp add: eval_monom_Cons)
|
|
1786 |
qed (insert assms, auto)
|
|
1787 |
|
|
1788 |
|
|
1789 |
subsubsection \<open>Correctness of multiseries operations\<close>
|
|
1790 |
|
|
1791 |
lemma drop_zero_ms_aux:
|
|
1792 |
assumes "eventually (\<lambda>x. eval C x = 0) at_top"
|
|
1793 |
assumes "is_expansion_aux (MSLCons (C, e) xs) f basis"
|
|
1794 |
shows "is_expansion_aux xs f basis"
|
|
1795 |
proof (rule is_expansion_aux_cong)
|
|
1796 |
from assms(2) show "is_expansion_aux xs (\<lambda>x. f x - eval C x * hd basis x powr e) basis"
|
|
1797 |
by (auto elim: is_expansion_aux_MSLCons)
|
|
1798 |
from assms(1) show "eventually (\<lambda>x. f x - eval C x * hd basis x powr e = f x) at_top"
|
|
1799 |
by eventually_elim simp
|
|
1800 |
qed
|
|
1801 |
|
|
1802 |
lemma dominant_term_ms_aux_bigo:
|
|
1803 |
assumes basis: "basis_wf basis" and xs: "is_expansion_aux xs f basis"
|
|
1804 |
shows "f \<in> O(eval_monom (1, snd (dominant_term_ms_aux xs)) basis)" (is ?thesis1)
|
|
1805 |
proof (cases xs rule: ms_aux_cases)
|
|
1806 |
case MSLNil
|
|
1807 |
with assms have "eventually (\<lambda>x. f x = 0) at_top"
|
|
1808 |
by (auto elim: is_expansion_aux.cases)
|
|
1809 |
hence "f \<in> \<Theta>(\<lambda>_. 0)" by (intro bigthetaI_cong) auto
|
|
1810 |
also have "(\<lambda>_. 0) \<in> O(eval_monom (1, snd (dominant_term_ms_aux xs)) basis)" by simp
|
|
1811 |
finally show ?thesis .
|
|
1812 |
next
|
|
1813 |
case [simp]: (MSLCons C e xs')
|
|
1814 |
from xs obtain b basis' where xs':
|
|
1815 |
"basis = b # basis'"
|
|
1816 |
"is_expansion_aux xs' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
1817 |
"is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
|
|
1818 |
"ms_exp_gt e (ms_aux_hd_exp xs')"
|
|
1819 |
by (cases rule: is_expansion_aux.cases) auto
|
|
1820 |
from is_expansion_aux_imp_smallo''[OF xs'(2,5)]
|
|
1821 |
obtain e' where e': "e' < e" "(\<lambda>x. f x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')"
|
|
1822 |
unfolding list.sel by blast
|
|
1823 |
from basis xs' have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1824 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
1825 |
|
|
1826 |
let ?h = "(\<lambda>x. eval_monom (1, snd (dominant_term C)) basis' x * b x powr e)"
|
|
1827 |
note e'(2)
|
|
1828 |
also have "(\<lambda>x. b x powr e') \<in> \<Theta>(\<lambda>x. b x powr (e' - e) * b x powr e)"
|
|
1829 |
by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: powr_diff)
|
|
1830 |
also have "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(?h)"
|
|
1831 |
unfolding smallomega_iff_smallo [symmetric] using e'(1) basis xs'
|
|
1832 |
by (intro landau_omega.small_big_mult landau_omega.big_refl eval_monom_smallomega')
|
|
1833 |
(simp_all add: basis_wf_Cons)
|
|
1834 |
finally have "(\<lambda>x. f x - eval C x * b x powr e) \<in> O(?h)" by (rule landau_o.small_imp_big)
|
|
1835 |
moreover have "(\<lambda>x. eval C x * b x powr e) \<in> O(?h)"
|
|
1836 |
using basis xs' by (intro landau_o.big.mult dominant_term_bigo) (auto simp: basis_wf_Cons)
|
|
1837 |
ultimately have "(\<lambda>x. f x - eval C x * b x powr e + eval C x * b x powr e) \<in> O(?h)"
|
|
1838 |
by (rule sum_in_bigo)
|
|
1839 |
hence "f \<in> O(?h)" by simp
|
|
1840 |
also have "?h = eval_monom (1, snd (dominant_term_ms_aux xs)) basis"
|
|
1841 |
using xs' by (auto simp: dominant_term_ms_aux_def case_prod_unfold eval_monom_Cons)
|
|
1842 |
finally show ?thesis .
|
|
1843 |
qed
|
|
1844 |
|
|
1845 |
lemma const_ms_aux:
|
|
1846 |
assumes basis: "basis_wf basis"
|
|
1847 |
and "length basis = Suc (expansion_level TYPE('a::multiseries))"
|
|
1848 |
shows "is_expansion_aux (const_ms_aux c :: ('a \<times> real) msllist) (\<lambda>_. c) basis"
|
|
1849 |
proof -
|
|
1850 |
from assms(2) obtain b basis' where [simp]: "basis = b # basis'" by (cases basis) simp_all
|
|
1851 |
from basis have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
1852 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
1853 |
|
|
1854 |
have "(\<lambda>_. c) \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e
|
|
1855 |
proof -
|
|
1856 |
have "(\<lambda>_. c) \<in> O(\<lambda>_. 1)" by (cases "c = 0") simp_all
|
|
1857 |
also from b_pos have "(\<lambda>_. 1) \<in> \<Theta>(\<lambda>x. b x powr 0)"
|
|
1858 |
by (intro bigthetaI_cong) (auto elim: eventually_mono)
|
|
1859 |
also from that b_limit have "(\<lambda>x. b x powr 0) \<in> o(\<lambda>x. b x powr e)"
|
|
1860 |
by (subst powr_smallo_iff) auto
|
|
1861 |
finally show ?thesis .
|
|
1862 |
qed
|
|
1863 |
with b_pos assms show ?thesis
|
|
1864 |
by (auto intro!: is_expansion_aux.intros simp: const_ms_aux_def is_expansion_const basis_wf_Cons
|
|
1865 |
elim: eventually_mono)
|
|
1866 |
qed
|
|
1867 |
|
|
1868 |
lemma uminus_ms_aux:
|
|
1869 |
assumes basis: "basis_wf basis"
|
|
1870 |
assumes F: "is_expansion_aux F f basis"
|
|
1871 |
shows "is_expansion_aux (uminus_ms_aux F) (\<lambda>x. -f x) basis"
|
|
1872 |
using F
|
|
1873 |
proof (coinduction arbitrary: F f rule: is_expansion_aux_coinduct)
|
|
1874 |
case (is_expansion_aux F f)
|
|
1875 |
note IH = is_expansion_aux
|
|
1876 |
show ?case
|
|
1877 |
proof (cases F rule: ms_aux_cases)
|
|
1878 |
case MSLNil
|
|
1879 |
with IH show ?thesis by (auto simp: uminus_ms_aux_def elim: is_expansion_aux.cases)
|
|
1880 |
next
|
|
1881 |
case (MSLCons C e xs)
|
|
1882 |
with IH obtain b basis'
|
|
1883 |
where IH: "basis = b # basis'"
|
|
1884 |
"is_expansion_aux xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
1885 |
"is_expansion C basis'"
|
|
1886 |
"\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp xs)"
|
|
1887 |
by (auto elim: is_expansion_aux_MSLCons)
|
|
1888 |
from basis IH have basis': "basis_wf basis'" by (simp add: basis_wf_Cons)
|
|
1889 |
with IH basis show ?thesis
|
|
1890 |
by (force simp: uminus_ms_aux_MSLCons MSLCons is_expansion_uminus)
|
|
1891 |
qed
|
|
1892 |
qed
|
|
1893 |
|
|
1894 |
lemma plus_ms_aux:
|
|
1895 |
assumes "basis_wf basis" "is_expansion_aux F f basis" "is_expansion_aux G g basis"
|
|
1896 |
shows "is_expansion_aux (plus_ms_aux F G) (\<lambda>x. f x + g x) basis"
|
|
1897 |
using assms
|
|
1898 |
proof (coinduction arbitrary: F f G g rule: is_expansion_aux_coinduct)
|
|
1899 |
case (is_expansion_aux F f G g)
|
|
1900 |
note IH = this
|
|
1901 |
show ?case
|
|
1902 |
proof (cases F G rule: ms_aux_cases[case_product ms_aux_cases])
|
|
1903 |
assume [simp]: "F = MSLNil" "G = MSLNil"
|
|
1904 |
with IH have *: "eventually (\<lambda>x. f x = 0) at_top" "eventually (\<lambda>x. g x = 0) at_top"
|
|
1905 |
and length: "length basis = Suc (expansion_level TYPE('a))"
|
|
1906 |
by (auto elim: is_expansion_aux.cases)
|
|
1907 |
from * have "eventually (\<lambda>x. f x + g x = 0) at_top" by eventually_elim simp
|
|
1908 |
with length show ?case by simp
|
|
1909 |
next
|
|
1910 |
assume [simp]: "F = MSLNil"
|
|
1911 |
fix C e G' assume G: "G = MSLCons (C, e) G'"
|
|
1912 |
from IH have f: "eventually (\<lambda>x. f x = 0) at_top" by (auto elim: is_expansion_aux.cases)
|
|
1913 |
from f have "eventually (\<lambda>x. f x + g x = g x) at_top" by eventually_elim simp
|
|
1914 |
note eq = landau_o.small.in_cong[OF this]
|
|
1915 |
from IH(3) G obtain b basis' where G':
|
|
1916 |
"basis = b # basis'"
|
|
1917 |
"is_expansion_aux G' (\<lambda>x. g x - eval C x * b x powr e) (b # basis')"
|
|
1918 |
"is_expansion C basis'" "\<forall>e'>e. g \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp G')"
|
|
1919 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
1920 |
show ?thesis
|
|
1921 |
by (rule disjI2, inst_existentials G' b e basis' C F f G' "\<lambda>x. g x - eval C x * b x powr e")
|
|
1922 |
(insert G' IH(1,2), simp_all add: G eq algebra_simps)
|
|
1923 |
next
|
|
1924 |
assume [simp]: "G = MSLNil"
|
|
1925 |
fix C e F' assume F: "F = MSLCons (C, e) F'"
|
|
1926 |
from IH have g: "eventually (\<lambda>x. g x = 0) at_top" by (auto elim: is_expansion_aux.cases)
|
|
1927 |
hence "eventually (\<lambda>x. f x + g x = f x) at_top" by eventually_elim simp
|
|
1928 |
note eq = landau_o.small.in_cong[OF this]
|
|
1929 |
from IH F obtain b basis' where F':
|
|
1930 |
"basis = b # basis'"
|
|
1931 |
"is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
1932 |
"is_expansion C basis'" "\<forall>e'>e. f \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp F')"
|
|
1933 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
1934 |
show ?thesis
|
|
1935 |
by (rule disjI2, inst_existentials F' b e basis' C F' "\<lambda>x. f x - eval C x * b x powr e" G g)
|
|
1936 |
(insert F g F' IH, simp_all add: eq algebra_simps)
|
|
1937 |
next
|
|
1938 |
fix C1 e1 F' C2 e2 G'
|
|
1939 |
assume F: "F = MSLCons (C1, e1) F'" and G: "G = MSLCons (C2, e2) G'"
|
|
1940 |
from IH F obtain b basis' where
|
|
1941 |
basis': "basis = b # basis'" and F':
|
|
1942 |
"is_expansion_aux F' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
|
|
1943 |
"is_expansion C1 basis'" "\<And>e'. e' > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')"
|
|
1944 |
"ms_exp_gt e1 (ms_aux_hd_exp F')"
|
|
1945 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
1946 |
from IH G basis' have G':
|
|
1947 |
"is_expansion_aux G' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
|
|
1948 |
"is_expansion C2 basis'" "\<And>e'. e' > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e')"
|
|
1949 |
"ms_exp_gt e2 (ms_aux_hd_exp G')"
|
|
1950 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
1951 |
hence *: "\<forall>x>max e1 e2. (\<lambda>x. f x + g x) \<in> o(\<lambda>xa. b xa powr x)"
|
|
1952 |
by (intro allI impI sum_in_smallo F' G') simp_all
|
|
1953 |
|
|
1954 |
consider "e1 > e2" | "e1 < e2" | "e1 = e2" by force
|
|
1955 |
thus ?thesis
|
|
1956 |
proof cases
|
|
1957 |
assume e: "e1 > e2"
|
|
1958 |
have gt: "ms_exp_gt e1 (combine_options max (ms_aux_hd_exp F') (Some e2))"
|
|
1959 |
using \<open>ms_exp_gt e1 _\<close> e by (cases "ms_aux_hd_exp F'") simp_all
|
|
1960 |
show ?thesis
|
|
1961 |
by (rule disjI2, inst_existentials "plus_ms_aux F' G" b e1 basis' C1 F'
|
|
1962 |
"\<lambda>x. f x - eval C1 x * b x powr e1" G g)
|
|
1963 |
(insert e F'(1,2,4) IH(1,3) basis'(1) * gt,
|
|
1964 |
simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
|
|
1965 |
next
|
|
1966 |
assume e: "e1 < e2"
|
|
1967 |
have gt: "ms_exp_gt e2 (combine_options max (Some e1) (ms_aux_hd_exp G'))"
|
|
1968 |
using \<open>ms_exp_gt e2 _\<close> e by (cases "ms_aux_hd_exp G'") simp_all
|
|
1969 |
show ?thesis
|
|
1970 |
by (rule disjI2, inst_existentials "plus_ms_aux F G'" b e2 basis' C2 F f G'
|
|
1971 |
"\<lambda>x. g x - eval C2 x * b x powr e2")
|
|
1972 |
(insert e G'(1,2,4) IH(1,2) basis'(1) * gt,
|
|
1973 |
simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus)
|
|
1974 |
next
|
|
1975 |
assume e: "e1 = e2"
|
|
1976 |
have gt: "ms_exp_gt e2 (combine_options max (ms_aux_hd_exp F') (ms_aux_hd_exp G'))"
|
|
1977 |
using \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close> e
|
|
1978 |
by (cases "ms_aux_hd_exp F'"; cases "ms_aux_hd_exp G'") simp_all
|
|
1979 |
show ?thesis
|
|
1980 |
by (rule disjI2, inst_existentials "plus_ms_aux F' G'" b e1 basis' "C1 + C2" F'
|
|
1981 |
"\<lambda>x. f x - eval C1 x * b x powr e1" G' "\<lambda>x. g x - eval C2 x * b x powr e2")
|
|
1982 |
(insert e F'(1,2,4) G'(1,2,4) IH(1) basis'(1) * gt,
|
|
1983 |
simp_all add: F G plus_ms_aux_MSLCons algebra_simps hd_exp_plus is_expansion_add basis_wf_Cons)
|
|
1984 |
qed
|
|
1985 |
qed
|
|
1986 |
qed
|
|
1987 |
|
|
1988 |
lemma minus_ms_aux:
|
|
1989 |
assumes "basis_wf basis" "is_expansion_aux F f basis" "is_expansion_aux G g basis"
|
|
1990 |
shows "is_expansion_aux (minus_ms_aux F G) (\<lambda>x. f x - g x) basis"
|
|
1991 |
proof -
|
|
1992 |
have "is_expansion_aux (minus_ms_aux F G) (\<lambda>x. f x + (- g x)) basis"
|
|
1993 |
unfolding minus_ms_aux_def by (intro plus_ms_aux uminus_ms_aux assms)
|
|
1994 |
thus ?thesis by simp
|
|
1995 |
qed
|
|
1996 |
|
|
1997 |
lemma scale_shift_ms_aux:
|
|
1998 |
assumes basis: "basis_wf (b # basis)"
|
|
1999 |
assumes F: "is_expansion_aux F f (b # basis)"
|
|
2000 |
assumes C: "is_expansion C basis"
|
|
2001 |
shows "is_expansion_aux (scale_shift_ms_aux (C, e) F) (\<lambda>x. eval C x * b x powr e * f x) (b # basis)"
|
|
2002 |
using F
|
|
2003 |
proof (coinduction arbitrary: F f rule: is_expansion_aux_coinduct)
|
|
2004 |
case (is_expansion_aux F f)
|
|
2005 |
note IH = is_expansion_aux
|
|
2006 |
show ?case
|
|
2007 |
proof (cases F rule: ms_aux_cases)
|
|
2008 |
case MSLNil
|
|
2009 |
with IH show ?thesis
|
|
2010 |
by (auto simp: scale_shift_ms_aux_def elim!: is_expansion_aux.cases eventually_mono)
|
|
2011 |
next
|
|
2012 |
case (MSLCons C' e' xs)
|
|
2013 |
with IH have IH: "is_expansion_aux xs (\<lambda>x. f x - eval C' x * b x powr e') (b # basis)"
|
|
2014 |
"is_expansion C' basis" "ms_exp_gt e' (ms_aux_hd_exp xs)"
|
|
2015 |
"\<And>e''. e' < e'' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'')"
|
|
2016 |
by (auto elim: is_expansion_aux_MSLCons)
|
|
2017 |
|
|
2018 |
have "(\<lambda>x. eval C x * b x powr e * f x) \<in> o(\<lambda>x. b x powr e'')" if "e'' > e + e'" for e''
|
|
2019 |
proof -
|
|
2020 |
define d where "d = (e'' - e - e') / 2"
|
|
2021 |
from that have d: "d > 0" by (simp add: d_def)
|
|
2022 |
have "(\<lambda>x. eval C x * b x powr e * f x) \<in> o(\<lambda>x. b x powr d * b x powr e * b x powr (e' + d))"
|
|
2023 |
using d basis
|
|
2024 |
by (intro landau_o.small.mult[OF landau_o.small_big_mult]
|
|
2025 |
is_expansion_imp_smallo[OF _ C] IH) (simp_all add: basis_wf_Cons)
|
|
2026 |
also have "\<dots> = o(\<lambda>x. b x powr e'')" by (simp add: d_def powr_add [symmetric])
|
|
2027 |
finally show ?thesis .
|
|
2028 |
qed
|
|
2029 |
moreover have "ms_exp_gt (e + e') (ms_aux_hd_exp (scale_shift_ms_aux (C, e) xs))"
|
|
2030 |
using IH(3) by (cases "ms_aux_hd_exp xs") (simp_all add: hd_exp_basis)
|
|
2031 |
ultimately show ?thesis using basis IH(1,2)
|
|
2032 |
by (intro disjI2, inst_existentials "scale_shift_ms_aux (C, e) xs" b "e + e'"
|
|
2033 |
basis "C * C'" xs "\<lambda>x. f x - eval C' x * b x powr e'")
|
|
2034 |
(simp_all add: scale_shift_ms_aux_MSLCons MSLCons is_expansion_mult basis_wf_Cons
|
|
2035 |
algebra_simps powr_add C)
|
|
2036 |
qed
|
|
2037 |
qed
|
|
2038 |
|
|
2039 |
lemma times_ms_aux:
|
|
2040 |
assumes basis: "basis_wf basis"
|
|
2041 |
assumes F: "is_expansion_aux F f basis" and G: "is_expansion_aux G g basis"
|
|
2042 |
shows "is_expansion_aux (times_ms_aux F G) (\<lambda>x. f x * g x) basis"
|
|
2043 |
using F G
|
|
2044 |
proof (coinduction arbitrary: F f G g rule: is_expansion_aux_coinduct_upto)
|
|
2045 |
case (B F f G g)
|
|
2046 |
note IH = this
|
|
2047 |
show ?case
|
|
2048 |
proof (cases "F = MSLNil \<or> G = MSLNil")
|
|
2049 |
case True
|
|
2050 |
with IH have "eventually (\<lambda>x. f x = 0) at_top \<or> eventually (\<lambda>x. g x = 0) at_top"
|
|
2051 |
and length: "length basis = Suc (expansion_level TYPE('a))"
|
|
2052 |
by (auto elim: is_expansion_aux.cases)
|
|
2053 |
from this(1) have "eventually (\<lambda>x. f x * g x = 0) at_top" by (auto elim!: eventually_mono)
|
|
2054 |
with length True show ?thesis by auto
|
|
2055 |
next
|
|
2056 |
case False
|
|
2057 |
from False obtain C1 e1 F' where F: "F = MSLCons (C1, e1) F'"
|
|
2058 |
by (cases F rule: ms_aux_cases) simp_all
|
|
2059 |
from False obtain C2 e2 G' where G: "G = MSLCons (C2, e2) G'"
|
|
2060 |
by (cases G rule: ms_aux_cases) simp_all
|
|
2061 |
|
|
2062 |
from IH(1) F obtain b basis' where
|
|
2063 |
basis': "basis = b # basis'" and F':
|
|
2064 |
"is_expansion_aux F' (\<lambda>x. f x - eval C1 x * b x powr e1) (b # basis')"
|
|
2065 |
"is_expansion C1 basis'" "\<And>e'. e' > e1 \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e')"
|
|
2066 |
"ms_exp_gt e1 (ms_aux_hd_exp F')"
|
|
2067 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
2068 |
from IH(2) G basis' have G':
|
|
2069 |
"is_expansion_aux G' (\<lambda>x. g x - eval C2 x * b x powr e2) (b # basis')"
|
|
2070 |
"is_expansion C2 basis'" "\<And>e'. e' > e2 \<Longrightarrow> g \<in> o(\<lambda>x. b x powr e')"
|
|
2071 |
"ms_exp_gt e2 (ms_aux_hd_exp G')"
|
|
2072 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
2073 |
let ?P = "(\<lambda>xs'' fa'' basisa''. \<exists>F f G. xs'' = times_ms_aux F G \<and>
|
|
2074 |
(\<exists>g. fa'' = (\<lambda>x. f x * g x) \<and> basisa'' = b # basis' \<and> is_expansion_aux F f
|
|
2075 |
(b # basis') \<and> is_expansion_aux G g (b # basis')))"
|
|
2076 |
have "ms_closure ?P (plus_ms_aux (scale_shift_ms_aux (C1, e1) G') (times_ms_aux F' G))
|
|
2077 |
(\<lambda>x. eval C1 x * b x powr e1 * (g x - eval C2 x * b x powr e2) +
|
|
2078 |
(f x - eval C1 x * b x powr e1) * g x) (b # basis')"
|
|
2079 |
(is "ms_closure _ ?zs _ _") using IH(2) basis' basis
|
|
2080 |
by (intro ms_closure_add ms_closure_embed'[OF scale_shift_ms_aux]
|
|
2081 |
ms_closure_mult[OF ms_closure_embed' ms_closure_embed'] F' G') simp_all
|
|
2082 |
hence "ms_closure ?P (plus_ms_aux (scale_shift_ms_aux (C1, e1) G') (times_ms_aux F' G))
|
|
2083 |
(\<lambda>x. f x * g x - eval C1 x * eval C2 x * b x powr (e1 + e2)) (b # basis')"
|
|
2084 |
by (simp add: algebra_simps powr_add)
|
|
2085 |
moreover have "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr e)" if "e > e1 + e2" for e
|
|
2086 |
proof -
|
|
2087 |
define d where "d = (e - e1 - e2) / 2"
|
|
2088 |
from that have "d > 0" by (simp add: d_def)
|
|
2089 |
hence "(\<lambda>x. f x * g x) \<in> o(\<lambda>x. b x powr (e1 + d) * b x powr (e2 + d))"
|
|
2090 |
by (intro landau_o.small.mult F' G') simp_all
|
|
2091 |
also have "\<dots> = o(\<lambda>x. b x powr e)"
|
|
2092 |
by (simp add: d_def powr_add [symmetric])
|
|
2093 |
finally show ?thesis .
|
|
2094 |
qed
|
|
2095 |
moreover from \<open>ms_exp_gt e1 _\<close> \<open>ms_exp_gt e2 _\<close>
|
|
2096 |
have "ms_exp_gt (e1 + e2) (ms_aux_hd_exp ?zs)"
|
|
2097 |
by (cases "ms_aux_hd_exp F'"; cases "ms_aux_hd_exp G'")
|
|
2098 |
(simp_all add: hd_exp_times hd_exp_plus hd_exp_basis G)
|
|
2099 |
ultimately show ?thesis using F'(2) G'(2) basis' basis
|
|
2100 |
by (simp add: times_ms_aux_MSLCons basis_wf_Cons is_expansion_mult F G)
|
|
2101 |
qed
|
|
2102 |
qed (insert basis, simp_all)
|
|
2103 |
|
|
2104 |
|
|
2105 |
|
|
2106 |
|
|
2107 |
lemma powser_ms_aux_MSLNil_iff [simp]: "powser_ms_aux cs f = MSLNil \<longleftrightarrow> cs = MSLNil"
|
|
2108 |
by (subst powser_ms_aux.code) (simp split: msllist.splits)
|
|
2109 |
|
|
2110 |
lemma powser_ms_aux_MSLNil [simp]: "powser_ms_aux MSLNil f = MSLNil"
|
|
2111 |
by (subst powser_ms_aux.code) simp
|
|
2112 |
|
|
2113 |
lemma powser_ms_aux_MSLNil' [simp]:
|
|
2114 |
"powser_ms_aux (MSLCons c cs) MSLNil = MSLCons (const_expansion c, 0) MSLNil"
|
|
2115 |
by (subst powser_ms_aux.code) simp
|
|
2116 |
|
|
2117 |
lemma powser_ms_aux_MSLCons:
|
|
2118 |
"powser_ms_aux (MSLCons c cs) f = MSLCons (const_expansion c, 0) (times_ms_aux f (powser_ms_aux cs f))"
|
|
2119 |
by (subst powser_ms_aux.code) simp
|
|
2120 |
|
|
2121 |
lemma hd_exp_powser: "ms_aux_hd_exp (powser_ms_aux cs f) = (if cs = MSLNil then None else Some 0)"
|
|
2122 |
by (subst powser_ms_aux.code) (simp split: msllist.splits)
|
|
2123 |
|
|
2124 |
lemma powser_ms_aux:
|
|
2125 |
assumes "convergent_powser cs" and basis: "basis_wf basis"
|
|
2126 |
assumes G: "is_expansion_aux G g basis" "ms_exp_gt 0 (ms_aux_hd_exp G)"
|
|
2127 |
shows "is_expansion_aux (powser_ms_aux cs G) (powser cs \<circ> g) basis"
|
|
2128 |
using assms(1)
|
|
2129 |
proof (coinduction arbitrary: cs rule: is_expansion_aux_coinduct_upto)
|
|
2130 |
case (B cs)
|
|
2131 |
show ?case
|
|
2132 |
proof (cases cs)
|
|
2133 |
case MSLNil
|
|
2134 |
with G show ?thesis by (auto simp: is_expansion_aux_expansion_level)
|
|
2135 |
next
|
|
2136 |
case (MSLCons c cs')
|
|
2137 |
from is_expansion_aux_basis_nonempty[OF G(1)]
|
|
2138 |
obtain b basis' where basis': "basis = b # basis'" by (cases basis) simp_all
|
|
2139 |
from B have conv: "convergent_powser cs'" by (auto simp: MSLCons dest: convergent_powser_stl)
|
|
2140 |
from basis basis' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
2141 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
2142 |
|
|
2143 |
from G(2) have "g \<in> o(\<lambda>x. hd basis x powr 0)"
|
|
2144 |
by (intro is_expansion_aux_imp_smallo[OF G(1)]) simp
|
|
2145 |
with basis' have "g \<in> o(\<lambda>x. b x powr 0)" by simp
|
|
2146 |
also have "(\<lambda>x. b x powr 0) \<in> \<Theta>(\<lambda>x. 1)"
|
|
2147 |
using b_pos by (intro bigthetaI_cong) (auto elim!: eventually_mono)
|
|
2148 |
finally have g_limit: "(g \<longlongrightarrow> 0) at_top" by (auto dest: smalloD_tendsto)
|
|
2149 |
|
|
2150 |
let ?P = "(\<lambda>xs'' fa'' basisa''. \<exists>cs. xs'' = powser_ms_aux cs G \<and>
|
|
2151 |
fa'' = powser cs \<circ> g \<and> basisa'' = b # basis' \<and> convergent_powser cs)"
|
|
2152 |
have "ms_closure ?P (times_ms_aux G (powser_ms_aux cs' G))
|
|
2153 |
(\<lambda>x. g x * (powser cs' \<circ> g) x) basis" using conv
|
|
2154 |
by (intro ms_closure_mult [OF ms_closure_embed' ms_closure_embed] G)
|
|
2155 |
(auto simp add: basis' o_def)
|
|
2156 |
also have "(\<lambda>x. g x * (powser cs' \<circ> g) x) = (\<lambda>x. x * powser cs' x) \<circ> g"
|
|
2157 |
by (simp add: o_def)
|
|
2158 |
finally have "ms_closure ?P (times_ms_aux G (powser_ms_aux cs' G))
|
|
2159 |
(\<lambda>x. (powser cs \<circ> g) x - c * b x powr 0) basis" unfolding o_def
|
|
2160 |
proof (rule ms_closure_cong)
|
|
2161 |
note b_pos
|
|
2162 |
moreover have "eventually (\<lambda>x. powser cs (g x) = g x * powser cs' (g x) + c) at_top"
|
|
2163 |
proof -
|
|
2164 |
from B have "eventually (\<lambda>x. powser cs x = x * powser cs' x + c) (nhds 0)"
|
|
2165 |
by (simp add: MSLCons powser_MSLCons)
|
|
2166 |
from this and g_limit show ?thesis by (rule eventually_compose_filterlim)
|
|
2167 |
qed
|
|
2168 |
ultimately show "eventually (\<lambda>x. g x * powser cs' (g x) =
|
|
2169 |
powser cs (g x) - c * b x powr 0) at_top"
|
|
2170 |
by eventually_elim simp
|
|
2171 |
qed
|
|
2172 |
moreover from basis G have "is_expansion (const_expansion c :: 'a) basis'"
|
|
2173 |
by (intro is_expansion_const)
|
|
2174 |
(auto dest: is_expansion_aux_expansion_level simp: basis' basis_wf_Cons)
|
|
2175 |
moreover have "powser cs \<circ> g \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e
|
|
2176 |
proof -
|
|
2177 |
have "((powser cs \<circ> g) \<longlongrightarrow> powser cs 0) at_top" unfolding o_def
|
|
2178 |
by (intro isCont_tendsto_compose[OF _ g_limit] isCont_powser B)
|
|
2179 |
hence "powser cs \<circ> g \<in> O(\<lambda>_. 1)"
|
|
2180 |
by (intro bigoI_tendsto[where c = "powser cs 0"]) (simp_all add: o_def)
|
|
2181 |
also from b_pos have "O(\<lambda>_. 1) = O(\<lambda>x. b x powr 0)"
|
|
2182 |
by (intro landau_o.big.cong) (auto elim: eventually_mono)
|
|
2183 |
also from that b_limit have "(\<lambda>x. b x powr 0) \<in> o(\<lambda>x. b x powr e)"
|
|
2184 |
by (subst powr_smallo_iff) simp_all
|
|
2185 |
finally show ?thesis .
|
|
2186 |
qed
|
|
2187 |
moreover from G have "ms_exp_gt 0 (ms_aux_hd_exp (times_ms_aux G (powser_ms_aux cs' G)))"
|
|
2188 |
by (cases "ms_aux_hd_exp G") (simp_all add: hd_exp_times MSLCons hd_exp_powser)
|
|
2189 |
ultimately show ?thesis
|
|
2190 |
by (simp add: MSLCons powser_ms_aux_MSLCons basis')
|
|
2191 |
qed
|
|
2192 |
qed (insert assms, simp_all)
|
|
2193 |
|
|
2194 |
lemma powser_ms_aux':
|
|
2195 |
assumes powser: "convergent_powser' cs f" and basis: "basis_wf basis"
|
|
2196 |
assumes G: "is_expansion_aux G g basis" "ms_exp_gt 0 (ms_aux_hd_exp G)"
|
|
2197 |
shows "is_expansion_aux (powser_ms_aux cs G) (f \<circ> g) basis"
|
|
2198 |
proof (rule is_expansion_aux_cong)
|
|
2199 |
from is_expansion_aux_basis_nonempty[OF G(1)] basis
|
|
2200 |
have "filterlim (hd basis) at_top at_top" by (cases basis) (simp_all add: basis_wf_Cons)
|
|
2201 |
hence pos: "eventually (\<lambda>x. hd basis x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
2202 |
from G(2) have "g \<in> o(\<lambda>x. hd basis x powr 0)"
|
|
2203 |
by (intro is_expansion_aux_imp_smallo[OF G(1)]) simp
|
|
2204 |
also have "(\<lambda>x. hd basis x powr 0) \<in> \<Theta>(\<lambda>x. 1)"
|
|
2205 |
using pos by (intro bigthetaI_cong) (auto elim!: eventually_mono)
|
|
2206 |
finally have g_limit: "(g \<longlongrightarrow> 0) at_top" by (auto dest: smalloD_tendsto)
|
|
2207 |
|
|
2208 |
from powser have "eventually (\<lambda>x. powser cs x = f x) (nhds 0)"
|
|
2209 |
by (rule convergent_powser'_eventually)
|
|
2210 |
from this and g_limit show "eventually (\<lambda>x. (powser cs \<circ> g) x = (f \<circ> g) x) at_top"
|
|
2211 |
unfolding o_def by (rule eventually_compose_filterlim)
|
|
2212 |
qed (rule assms powser_ms_aux convergent_powser'_imp_convergent_powser)+
|
|
2213 |
|
|
2214 |
lemma inverse_ms_aux:
|
|
2215 |
assumes basis: "basis_wf basis"
|
|
2216 |
assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F"
|
|
2217 |
shows "is_expansion_aux (inverse_ms_aux F) (\<lambda>x. inverse (f x)) basis"
|
|
2218 |
proof (cases F rule: ms_aux_cases)
|
|
2219 |
case (MSLCons C e F')
|
|
2220 |
from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C"
|
|
2221 |
"is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
2222 |
"is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
|
|
2223 |
"ms_exp_gt e (ms_aux_hd_exp F')"
|
|
2224 |
by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def)
|
|
2225 |
define f' where "f' = (\<lambda>x. f x - eval C x * b x powr e)"
|
|
2226 |
from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
2227 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
2228 |
moreover from F'(6)
|
|
2229 |
have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))"
|
|
2230 |
by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis)
|
|
2231 |
ultimately have
|
|
2232 |
"is_expansion_aux (inverse_ms_aux F)
|
|
2233 |
(\<lambda>x. eval (inverse C) x * b x powr (-e) *
|
|
2234 |
((\<lambda>x. inverse (1 + x)) \<circ> (\<lambda>x. eval (inverse C) x * b x powr (-e) * f' x)) x)
|
|
2235 |
(b # basis')" (is "is_expansion_aux _ ?h _")
|
|
2236 |
unfolding MSLCons inverse_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric]
|
|
2237 |
using basis F(2) F'(1,3,4)
|
|
2238 |
by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse)
|
|
2239 |
(simp_all add: convergent_powser'_geometric basis_wf_Cons trimmed_ms_aux_def MSLCons)
|
|
2240 |
also have "b # basis' = basis" by (simp add: F')
|
|
2241 |
finally show ?thesis
|
|
2242 |
proof (rule is_expansion_aux_cong, goal_cases)
|
|
2243 |
case 1
|
|
2244 |
from basis F' have "eventually (\<lambda>x. eval C x \<noteq> 0) at_top"
|
|
2245 |
by (intro trimmed_imp_eventually_nz[of basis']) (simp_all add: basis_wf_Cons)
|
|
2246 |
with b_pos show ?case by eventually_elim (simp add: o_def field_simps powr_minus f'_def)
|
|
2247 |
qed
|
|
2248 |
qed (insert assms, auto simp: trimmed_ms_aux_def)
|
|
2249 |
|
|
2250 |
lemma hd_exp_inverse:
|
|
2251 |
"xs \<noteq> MSLNil \<Longrightarrow> ms_aux_hd_exp (inverse_ms_aux xs) = map_option uminus (ms_aux_hd_exp xs)"
|
|
2252 |
by (cases xs) (auto simp: Let_def hd_exp_basis hd_exp_powser inverse_ms_aux.simps)
|
|
2253 |
|
|
2254 |
lemma eval_pos_if_dominant_term_pos:
|
|
2255 |
assumes "basis_wf basis" "is_expansion F basis" "trimmed F"
|
|
2256 |
"fst (dominant_term F) > 0"
|
|
2257 |
shows "eventually (\<lambda>x. eval F x > 0) at_top"
|
|
2258 |
proof -
|
|
2259 |
have "eval F \<sim>[at_top] eval_monom (dominant_term F) basis"
|
|
2260 |
by (intro dominant_term assms)
|
|
2261 |
from trimmed_imp_eventually_sgn[OF assms(1-3)]
|
|
2262 |
have "\<forall>\<^sub>F x in at_top. sgn (eval F x) = sgn (fst (dominant_term F))" .
|
|
2263 |
moreover from assms
|
|
2264 |
have "eventually (\<lambda>x. eval_monom (dominant_term F) basis x > 0) at_top"
|
|
2265 |
by (intro eval_monom_pos)
|
|
2266 |
ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits)
|
|
2267 |
qed
|
|
2268 |
|
|
2269 |
lemma eval_pos_if_dominant_term_pos':
|
|
2270 |
assumes "basis_wf basis" "trimmed_ms_aux F" "is_expansion_aux F f basis"
|
|
2271 |
"fst (dominant_term_ms_aux F) > 0"
|
|
2272 |
shows "eventually (\<lambda>x. f x > 0) at_top"
|
|
2273 |
proof -
|
|
2274 |
have "f \<sim>[at_top] eval_monom (dominant_term_ms_aux F) basis"
|
|
2275 |
by (intro dominant_term_ms_aux assms)
|
|
2276 |
from dominant_term_ms_aux(2)[OF assms(1-3)]
|
|
2277 |
have "\<forall>\<^sub>F x in at_top. sgn (f x) = sgn (fst (dominant_term_ms_aux F))" .
|
|
2278 |
moreover from assms
|
|
2279 |
have "eventually (\<lambda>x. eval_monom (dominant_term_ms_aux F) basis x > 0) at_top"
|
|
2280 |
by (intro eval_monom_pos)
|
|
2281 |
ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits)
|
|
2282 |
qed
|
|
2283 |
|
|
2284 |
lemma eval_neg_if_dominant_term_neg':
|
|
2285 |
assumes "basis_wf basis" "trimmed_ms_aux F" "is_expansion_aux F f basis"
|
|
2286 |
"fst (dominant_term_ms_aux F) < 0"
|
|
2287 |
shows "eventually (\<lambda>x. f x < 0) at_top"
|
|
2288 |
proof -
|
|
2289 |
have "f \<sim>[at_top] eval_monom (dominant_term_ms_aux F) basis"
|
|
2290 |
by (intro dominant_term_ms_aux assms)
|
|
2291 |
from dominant_term_ms_aux(2)[OF assms(1-3)]
|
|
2292 |
have "\<forall>\<^sub>F x in at_top. sgn (f x) = sgn (fst (dominant_term_ms_aux F))" .
|
|
2293 |
moreover from assms
|
|
2294 |
have "eventually (\<lambda>x. eval_monom (dominant_term_ms_aux F) basis x < 0) at_top"
|
|
2295 |
by (intro eval_monom_neg)
|
|
2296 |
ultimately show ?thesis by eventually_elim (insert assms, simp add: sgn_if split: if_splits)
|
|
2297 |
qed
|
|
2298 |
|
|
2299 |
lemma fst_dominant_term_ms_aux_MSLCons:
|
|
2300 |
"fst (dominant_term_ms_aux (MSLCons x xs)) = fst (dominant_term (fst x))"
|
|
2301 |
by (auto simp: dominant_term_ms_aux_def split: prod.splits)
|
|
2302 |
|
|
2303 |
lemma powr_ms_aux:
|
|
2304 |
assumes basis: "basis_wf basis"
|
|
2305 |
assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F" "fst (dominant_term_ms_aux F) > 0"
|
|
2306 |
shows "is_expansion_aux (powr_ms_aux abort F p) (\<lambda>x. f x powr p) basis"
|
|
2307 |
proof (cases F rule: ms_aux_cases)
|
|
2308 |
case (MSLCons C e F')
|
|
2309 |
from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C" "fst (dominant_term C) > 0"
|
|
2310 |
"is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
2311 |
"is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
|
|
2312 |
"ms_exp_gt e (ms_aux_hd_exp F')"
|
|
2313 |
by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def dominant_term_ms_aux_def
|
|
2314 |
split: prod.splits)
|
|
2315 |
define f' where "f' = (\<lambda>x. f x - eval C x * b x powr e)"
|
|
2316 |
from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
2317 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
2318 |
moreover from F' have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))"
|
|
2319 |
by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis)
|
|
2320 |
ultimately have
|
|
2321 |
"is_expansion_aux (powr_ms_aux abort F p)
|
|
2322 |
(\<lambda>x. eval (powr_expansion abort C p) x * b x powr (e * p) *
|
|
2323 |
((\<lambda>x. (1 + x) powr p) \<circ> (\<lambda>x. eval (inverse C) x * b x powr (-e) * f' x)) x)
|
|
2324 |
(b # basis')" (is "is_expansion_aux _ ?h _")
|
|
2325 |
unfolding MSLCons powr_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric]
|
|
2326 |
using basis F'(1-5)
|
|
2327 |
by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse is_expansion_powr)
|
|
2328 |
(simp_all add: MSLCons basis_wf_Cons convergent_powser'_gbinomial)
|
|
2329 |
also have "b # basis' = basis" by (simp add: F')
|
|
2330 |
finally show ?thesis
|
|
2331 |
proof (rule is_expansion_aux_cong, goal_cases)
|
|
2332 |
case 1
|
|
2333 |
from basis F' have "eventually (\<lambda>x. eval C x > 0) at_top"
|
|
2334 |
by (intro eval_pos_if_dominant_term_pos[of basis']) (simp_all add: basis_wf_Cons)
|
|
2335 |
moreover from basis F have "eventually (\<lambda>x. f x > 0) at_top"
|
|
2336 |
by (intro eval_pos_if_dominant_term_pos'[of basis F])
|
|
2337 |
ultimately show ?case using b_pos
|
|
2338 |
by eventually_elim
|
|
2339 |
(simp add: powr_powr [symmetric] powr_minus powr_mult powr_divide f'_def field_simps)
|
|
2340 |
qed
|
|
2341 |
qed (insert assms, auto simp: trimmed_ms_aux_def)
|
|
2342 |
|
|
2343 |
lemma power_ms_aux:
|
|
2344 |
assumes basis: "basis_wf basis"
|
|
2345 |
assumes F: "is_expansion_aux F f basis" "trimmed_ms_aux F"
|
|
2346 |
shows "is_expansion_aux (power_ms_aux abort F n) (\<lambda>x. f x ^ n) basis"
|
|
2347 |
proof (cases F rule: ms_aux_cases)
|
|
2348 |
case (MSLCons C e F')
|
|
2349 |
from F MSLCons obtain b basis' where F': "basis = b # basis'" "trimmed C"
|
|
2350 |
"is_expansion_aux F' (\<lambda>x. f x - eval C x * b x powr e) (b # basis')"
|
|
2351 |
"is_expansion C basis'" "(\<And>e'. e < e' \<Longrightarrow> f \<in> o(\<lambda>x. b x powr e'))"
|
|
2352 |
"ms_exp_gt e (ms_aux_hd_exp F')"
|
|
2353 |
by (auto elim!: is_expansion_aux_MSLCons simp: trimmed_ms_aux_def dominant_term_ms_aux_def
|
|
2354 |
split: prod.splits)
|
|
2355 |
define f' where "f' = (\<lambda>x. f x - eval C x * b x powr e)"
|
|
2356 |
from basis F' have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
2357 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
2358 |
moreover have "ms_exp_gt 0 (ms_aux_hd_exp (scale_shift_ms_aux (inverse C, - e) F'))"
|
|
2359 |
using F'(6) by (cases "ms_aux_hd_exp F'") (simp_all add: hd_exp_basis)
|
|
2360 |
ultimately have
|
|
2361 |
"is_expansion_aux (power_ms_aux abort F n)
|
|
2362 |
(\<lambda>x. eval (power_expansion abort C n) x * b x powr (e * real n) *
|
|
2363 |
((\<lambda>x. (1 + x) ^ n) \<circ> (\<lambda>x. eval (inverse C) x * b x powr (-e) * f' x)) x)
|
|
2364 |
(b # basis')" (is "is_expansion_aux _ ?h _")
|
|
2365 |
unfolding MSLCons power_ms_aux.simps Let_def fst_conv snd_conv f'_def eval_inverse [symmetric]
|
|
2366 |
using basis F'(1-4)
|
|
2367 |
by (intro scale_shift_ms_aux powser_ms_aux' is_expansion_inverse is_expansion_power)
|
|
2368 |
(simp_all add: MSLCons basis_wf_Cons convergent_powser'_gbinomial')
|
|
2369 |
also have "b # basis' = basis" by (simp add: F')
|
|
2370 |
finally show ?thesis
|
|
2371 |
proof (rule is_expansion_aux_cong, goal_cases)
|
|
2372 |
case 1
|
|
2373 |
from F'(1,2,4) basis have "eventually (\<lambda>x. eval C x \<noteq> 0) at_top"
|
|
2374 |
using trimmed_imp_eventually_nz[of basis' C] by (simp add: basis_wf_Cons)
|
|
2375 |
thus ?case using b_pos
|
|
2376 |
by eventually_elim
|
|
2377 |
(simp add: field_simps f'_def powr_minus powr_powr [symmetric] powr_realpow
|
|
2378 |
power_mult_distrib [symmetric] power_divide [symmetric]
|
|
2379 |
del: power_mult_distrib power_divide)
|
|
2380 |
qed
|
|
2381 |
qed (insert assms, auto simp: trimmed_ms_aux_def)
|
|
2382 |
|
|
2383 |
lemma snd_dominant_term_ms_aux_MSLCons:
|
|
2384 |
"snd (dominant_term_ms_aux (MSLCons (C, e) xs)) = e # snd (dominant_term C)"
|
|
2385 |
by (simp add: dominant_term_ms_aux_def case_prod_unfold)
|
|
2386 |
|
|
2387 |
|
|
2388 |
subsubsection \<open>Type-class instantiations\<close>
|
|
2389 |
|
|
2390 |
datatype 'a ms = MS "('a \<times> real) msllist" "real \<Rightarrow> real"
|
|
2391 |
|
|
2392 |
instantiation ms :: (uminus) uminus
|
|
2393 |
begin
|
|
2394 |
|
|
2395 |
primrec uminus_ms where
|
|
2396 |
"-(MS xs f) = MS (uminus_ms_aux xs) (\<lambda>x. -f x)"
|
|
2397 |
|
|
2398 |
instance ..
|
|
2399 |
end
|
|
2400 |
|
|
2401 |
instantiation ms :: (plus) plus
|
|
2402 |
begin
|
|
2403 |
|
|
2404 |
fun plus_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
|
|
2405 |
"MS xs f + MS ys g = MS (plus_ms_aux xs ys) (\<lambda>x. f x + g x)"
|
|
2406 |
|
|
2407 |
instance ..
|
|
2408 |
end
|
|
2409 |
|
|
2410 |
instantiation ms :: ("{plus,uminus}") minus
|
|
2411 |
begin
|
|
2412 |
|
|
2413 |
definition minus_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
|
|
2414 |
"F - G = F + -(G :: 'a ms)"
|
|
2415 |
|
|
2416 |
instance ..
|
|
2417 |
end
|
|
2418 |
|
|
2419 |
instantiation ms :: ("{plus,times}") times
|
|
2420 |
begin
|
|
2421 |
|
|
2422 |
fun times_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
|
|
2423 |
"MS xs f * MS ys g = MS (times_ms_aux xs ys) (\<lambda>x. f x * g x)"
|
|
2424 |
|
|
2425 |
instance ..
|
|
2426 |
end
|
|
2427 |
|
|
2428 |
instantiation ms :: (multiseries) inverse
|
|
2429 |
begin
|
|
2430 |
|
|
2431 |
fun inverse_ms :: "'a ms \<Rightarrow> 'a ms" where
|
|
2432 |
"inverse_ms (MS xs f) = MS (inverse_ms_aux xs) (\<lambda>x. inverse (f x))"
|
|
2433 |
|
|
2434 |
fun divide_ms :: "'a ms \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
|
|
2435 |
"divide_ms (MS xs f) (MS ys g) = MS (times_ms_aux xs (inverse_ms_aux ys)) (\<lambda>x. f x / g x)"
|
|
2436 |
|
|
2437 |
instance ..
|
|
2438 |
end
|
|
2439 |
|
|
2440 |
|
|
2441 |
instantiation ms :: (multiseries) multiseries
|
|
2442 |
begin
|
|
2443 |
|
|
2444 |
definition expansion_level_ms :: "'a ms itself \<Rightarrow> nat" where
|
|
2445 |
expansion_level_ms_def [simp]: "expansion_level_ms _ = Suc (expansion_level (TYPE('a)))"
|
|
2446 |
|
|
2447 |
definition zero_expansion_ms :: "'a ms" where
|
|
2448 |
"zero_expansion = MS MSLNil (\<lambda>_. 0)"
|
|
2449 |
|
|
2450 |
definition const_expansion_ms :: "real \<Rightarrow> 'a ms" where
|
|
2451 |
"const_expansion c = MS (const_ms_aux c) (\<lambda>_. c)"
|
|
2452 |
|
|
2453 |
primrec is_expansion_ms :: "'a ms \<Rightarrow> basis \<Rightarrow> bool" where
|
|
2454 |
"is_expansion (MS xs f) = is_expansion_aux xs f"
|
|
2455 |
|
|
2456 |
lemma is_expansion_ms_def': "is_expansion F = (case F of MS xs f \<Rightarrow> is_expansion_aux xs f)"
|
|
2457 |
by (simp add: is_expansion_ms_def split: ms.splits)
|
|
2458 |
|
|
2459 |
primrec eval_ms :: "'a ms \<Rightarrow> real \<Rightarrow> real" where
|
|
2460 |
"eval_ms (MS _ f) = f"
|
|
2461 |
|
|
2462 |
lemma eval_ms_def': "eval F = (case F of MS _ f \<Rightarrow> f)"
|
|
2463 |
by (cases F) simp_all
|
|
2464 |
|
|
2465 |
primrec powr_expansion_ms :: "bool \<Rightarrow> 'a ms \<Rightarrow> real \<Rightarrow> 'a ms" where
|
|
2466 |
"powr_expansion_ms abort (MS xs f) p = MS (powr_ms_aux abort xs p) (\<lambda>x. f x powr p)"
|
|
2467 |
|
|
2468 |
primrec power_expansion_ms :: "bool \<Rightarrow> 'a ms \<Rightarrow> nat \<Rightarrow> 'a ms" where
|
|
2469 |
"power_expansion_ms abort (MS xs f) n = MS (power_ms_aux abort xs n) (\<lambda>x. f x ^ n)"
|
|
2470 |
|
|
2471 |
primrec trimmed_ms :: "'a ms \<Rightarrow> bool" where
|
|
2472 |
"trimmed_ms (MS xs _) \<longleftrightarrow> trimmed_ms_aux xs"
|
|
2473 |
|
|
2474 |
primrec dominant_term_ms :: "'a ms \<Rightarrow> monom" where
|
|
2475 |
"dominant_term_ms (MS xs _) = dominant_term_ms_aux xs"
|
|
2476 |
|
|
2477 |
lemma length_dominant_term_ms:
|
|
2478 |
"length (snd (dominant_term (F :: 'a ms))) = Suc (expansion_level TYPE('a))"
|
|
2479 |
by (cases F) (simp_all add: length_dominant_term)
|
|
2480 |
|
|
2481 |
instance
|
|
2482 |
proof (standard, goal_cases length_basis zero const uminus plus minus times
|
|
2483 |
inverse divide powr power smallo smallomega trimmed dominant dominant_bigo)
|
|
2484 |
case (smallo basis F b e)
|
|
2485 |
from \<open>is_expansion F basis\<close> obtain xs f where F: "F = MS xs f" "is_expansion_aux xs f basis"
|
|
2486 |
by (simp add: is_expansion_ms_def' split: ms.splits)
|
|
2487 |
from F(2) have nonempty: "basis \<noteq> []" by (rule is_expansion_aux_basis_nonempty)
|
|
2488 |
with smallo have filterlim_hd_basis: "filterlim (hd basis) at_top at_top"
|
|
2489 |
by (cases basis) (simp_all add: basis_wf_Cons)
|
|
2490 |
from F(2) obtain e' where "f \<in> o(\<lambda>x. hd basis x powr e')"
|
|
2491 |
by (erule is_expansion_aux_imp_smallo')
|
|
2492 |
also from smallo nonempty filterlim_hd_basis have "(\<lambda>x. hd basis x powr e') \<in> o(\<lambda>x. b x powr e)"
|
|
2493 |
by (intro ln_smallo_imp_flat) auto
|
|
2494 |
finally show ?case by (simp add: F)
|
|
2495 |
next
|
|
2496 |
case (smallomega basis F b e)
|
|
2497 |
obtain xs f where F: "F = MS xs f" by (cases F) simp_all
|
|
2498 |
from this smallomega have *: "is_expansion_aux xs f basis" by simp
|
|
2499 |
with smallomega F have "f \<in> \<omega>(\<lambda>x. b x powr e)"
|
|
2500 |
by (intro is_expansion_aux_imp_smallomega [OF _ *])
|
|
2501 |
(simp_all add: is_expansion_ms_def' split: ms.splits)
|
|
2502 |
with F show ?case by simp
|
|
2503 |
next
|
|
2504 |
case (minus basis F G)
|
|
2505 |
thus ?case
|
|
2506 |
by (simp add: minus_ms_def is_expansion_ms_def' add_uminus_conv_diff [symmetric]
|
|
2507 |
plus_ms_aux uminus_ms_aux del: add_uminus_conv_diff split: ms.splits)
|
|
2508 |
next
|
|
2509 |
case (divide basis F G)
|
|
2510 |
have "G / F = G * inverse F" by (cases F; cases G) (simp add: divide_inverse)
|
|
2511 |
with divide show ?case
|
|
2512 |
by (simp add: is_expansion_ms_def' divide_inverse times_ms_aux inverse_ms_aux split: ms.splits)
|
|
2513 |
next
|
|
2514 |
fix c :: real
|
|
2515 |
show "trimmed (const_expansion c :: 'a ms) \<longleftrightarrow> c \<noteq> 0"
|
|
2516 |
by (simp add: const_expansion_ms_def trimmed_ms_aux_def const_ms_aux_def
|
|
2517 |
trimmed_const_expansion split: msllist.splits)
|
|
2518 |
next
|
|
2519 |
fix F :: "'a ms" assume "trimmed F"
|
|
2520 |
thus "fst (dominant_term F) \<noteq> 0"
|
|
2521 |
by (cases F) (auto simp: dominant_term_ms_aux_def trimmed_ms_aux_MSLCons case_prod_unfold
|
|
2522 |
trimmed_imp_dominant_term_nz split: msllist.splits)
|
|
2523 |
next
|
|
2524 |
fix F :: "'a ms"
|
|
2525 |
have "times_ms_aux (MSLCons (const_expansion 1, 0) MSLNil) xs = xs" for xs :: "('a \<times> real) msllist"
|
|
2526 |
proof (coinduction arbitrary: xs rule: msllist.coinduct_upto)
|
|
2527 |
case Eq_real_prod_msllist
|
|
2528 |
have "map_prod (( *) (const_expansion 1 :: 'a)) ((+) (0::real)) = (\<lambda>x. x)"
|
|
2529 |
by (rule ext) (simp add: map_prod_def times_const_expansion_1 case_prod_unfold)
|
|
2530 |
moreover have "mslmap \<dots> = (\<lambda>x. x)" by (rule ext) (simp add: msllist.map_ident)
|
|
2531 |
ultimately have "scale_shift_ms_aux (const_expansion 1 :: 'a, 0) = (\<lambda>x. x)"
|
|
2532 |
by (simp add: scale_shift_ms_aux_conv_mslmap)
|
|
2533 |
thus ?case
|
|
2534 |
by (cases xs rule: ms_aux_cases)
|
|
2535 |
(auto simp: times_ms_aux_MSLCons times_const_expansion_1
|
|
2536 |
times_ms_aux.corec.cong_refl)
|
|
2537 |
qed
|
|
2538 |
thus "const_expansion 1 * F = F"
|
|
2539 |
by (cases F) (simp add: const_expansion_ms_def const_ms_aux_def)
|
|
2540 |
next
|
|
2541 |
fix F :: "'a ms"
|
|
2542 |
show "fst (dominant_term (- F)) = - fst (dominant_term F)"
|
|
2543 |
"trimmed (-F) \<longleftrightarrow> trimmed F"
|
|
2544 |
by (cases F; simp add: dominant_term_ms_aux_def case_prod_unfold uminus_ms_aux_MSLCons
|
|
2545 |
trimmed_ms_aux_def split: msllist.splits)+
|
|
2546 |
next
|
|
2547 |
fix F :: "'a ms"
|
|
2548 |
show "zero_expansion + F = F" "F + zero_expansion = F"
|
|
2549 |
by (cases F; simp add: zero_expansion_ms_def)+
|
|
2550 |
qed (auto simp: eval_ms_def' const_expansion_ms_def trimmed_ms_aux_imp_nz is_expansion_ms_def'
|
|
2551 |
const_ms_aux uminus_ms_aux plus_ms_aux times_ms_aux inverse_ms_aux
|
|
2552 |
is_expansion_aux_expansion_level dominant_term_ms_aux length_dominant_term_ms
|
|
2553 |
minus_ms_def powr_ms_aux power_ms_aux zero_expansion_ms_def
|
|
2554 |
is_expansion_aux.intros(1) dominant_term_ms_aux_bigo
|
|
2555 |
split: ms.splits)
|
|
2556 |
|
|
2557 |
end
|
|
2558 |
|
|
2559 |
|
|
2560 |
subsubsection \<open>Changing the evaluation function of a multiseries\<close>
|
|
2561 |
|
|
2562 |
definition modify_eval :: "(real \<Rightarrow> real) \<Rightarrow> 'a ms \<Rightarrow> 'a ms" where
|
|
2563 |
"modify_eval f F = (case F of MS xs _ \<Rightarrow> MS xs f)"
|
|
2564 |
|
|
2565 |
lemma eval_modify_eval [simp]: "eval (modify_eval f F) = f"
|
|
2566 |
by (cases F) (simp add: modify_eval_def)
|
|
2567 |
|
|
2568 |
|
|
2569 |
subsubsection \<open>Scaling expansions\<close>
|
|
2570 |
|
|
2571 |
definition scale_ms :: "real \<Rightarrow> 'a :: multiseries \<Rightarrow> 'a" where
|
|
2572 |
"scale_ms c F = const_expansion c * F"
|
|
2573 |
|
|
2574 |
lemma scale_ms_real [simp]: "scale_ms c (c' :: real) = c * c'"
|
|
2575 |
by (simp add: scale_ms_def)
|
|
2576 |
|
|
2577 |
definition scale_ms_aux :: "real \<Rightarrow> ('a :: multiseries \<times> real) msllist \<Rightarrow> ('a \<times> real) msllist" where
|
|
2578 |
"scale_ms_aux c xs = scale_shift_ms_aux (const_expansion c, 0) xs"
|
|
2579 |
|
|
2580 |
|
|
2581 |
lemma scale_ms_aux_eq_MSLNil_iff [simp]: "scale_ms_aux x xs = MSLNil \<longleftrightarrow> xs = MSLNil"
|
|
2582 |
unfolding scale_ms_aux_def by simp
|
|
2583 |
|
|
2584 |
lemma times_ms_aux_singleton:
|
|
2585 |
"times_ms_aux (MSLCons (c, e) MSLNil) xs = scale_shift_ms_aux (c, e) xs"
|
|
2586 |
proof (coinduction arbitrary: xs rule: msllist.coinduct_strong)
|
|
2587 |
case (Eq_msllist xs)
|
|
2588 |
thus ?case
|
|
2589 |
by (cases xs rule: ms_aux_cases)
|
|
2590 |
(simp_all add: scale_shift_ms_aux_def times_ms_aux_MSLCons)
|
|
2591 |
qed
|
|
2592 |
|
|
2593 |
lemma scale_ms [simp]: "scale_ms c (MS xs f) = MS (scale_ms_aux c xs) (\<lambda>x. c * f x)"
|
|
2594 |
by (simp add: scale_ms_def scale_ms_aux_def const_expansion_ms_def const_ms_aux_def
|
|
2595 |
times_ms_aux_singleton)
|
|
2596 |
|
|
2597 |
lemma scale_ms_aux_MSLNil [simp]: "scale_ms_aux c MSLNil = MSLNil"
|
|
2598 |
by (simp add: scale_ms_aux_def)
|
|
2599 |
|
|
2600 |
lemma scale_ms_aux_MSLCons:
|
|
2601 |
"scale_ms_aux c (MSLCons (c', e) xs) = MSLCons (scale_ms c c', e) (scale_ms_aux c xs)"
|
|
2602 |
by (simp add: scale_ms_aux_def scale_shift_ms_aux_MSLCons scale_ms_def)
|
|
2603 |
|
|
2604 |
|
|
2605 |
subsubsection \<open>Additional convenience rules\<close>
|
|
2606 |
|
|
2607 |
lemma trimmed_pos_real_iff [simp]: "trimmed_pos (x :: real) \<longleftrightarrow> x > 0"
|
|
2608 |
by (auto simp: trimmed_pos_def)
|
|
2609 |
|
|
2610 |
lemma trimmed_pos_ms_iff:
|
|
2611 |
"trimmed_pos (MS xs f) \<longleftrightarrow> (case xs of MSLNil \<Rightarrow> False | MSLCons x xs \<Rightarrow> trimmed_pos (fst x))"
|
|
2612 |
by (auto simp add: trimmed_pos_def dominant_term_ms_aux_def trimmed_ms_aux_def
|
|
2613 |
split: msllist.splits prod.splits)
|
|
2614 |
|
|
2615 |
lemma not_trimmed_pos_MSLNil [simp]: "\<not>trimmed_pos (MS MSLNil f)"
|
|
2616 |
by (simp add: trimmed_pos_ms_iff)
|
|
2617 |
|
|
2618 |
lemma trimmed_pos_MSLCons [simp]: "trimmed_pos (MS (MSLCons x xs) f) = trimmed_pos (fst x)"
|
|
2619 |
by (simp add: trimmed_pos_ms_iff)
|
|
2620 |
|
|
2621 |
lemma drop_zero_ms:
|
|
2622 |
assumes "eventually (\<lambda>x. eval C x = 0) at_top"
|
|
2623 |
assumes "is_expansion (MS (MSLCons (C, e) xs) f) basis"
|
|
2624 |
shows "is_expansion (MS xs f) basis"
|
|
2625 |
using assms by (auto simp: is_expansion_ms_def intro: drop_zero_ms_aux)
|
|
2626 |
|
|
2627 |
lemma expansion_level_eq_Nil: "length [] = expansion_level TYPE(real)" by simp
|
|
2628 |
lemma expansion_level_eq_Cons:
|
|
2629 |
"length xs = expansion_level TYPE('a::multiseries) \<Longrightarrow>
|
|
2630 |
length (x # xs) = expansion_level TYPE('a ms)" by simp
|
|
2631 |
|
|
2632 |
lemma basis_wf_insert_ln:
|
|
2633 |
assumes "basis_wf [b]"
|
|
2634 |
shows "basis_wf [\<lambda>x. ln (b x)]" (is ?thesis1)
|
|
2635 |
"basis_wf [b, \<lambda>x. ln (b x)]" (is ?thesis2)
|
|
2636 |
"is_expansion (MS (MSLCons (1::real,1) MSLNil) (\<lambda>x. ln (b x))) [\<lambda>x. ln (b x)]" (is ?thesis3)
|
|
2637 |
proof -
|
|
2638 |
have "ln \<in> o(\<lambda>x. x :: real)"
|
|
2639 |
using ln_x_over_x_tendsto_0 by (intro smalloI_tendsto) auto
|
|
2640 |
with assms show ?thesis1 ?thesis2
|
|
2641 |
by (auto simp: basis_wf_Cons
|
|
2642 |
intro!: landau_o.small.compose[of _ _ _ "\<lambda>x. ln (b x)"] filterlim_compose[OF ln_at_top])
|
|
2643 |
from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
2644 |
hence ev: "eventually (\<lambda>x. b x > 1) at_top" by (simp add: filterlim_at_top_dense)
|
|
2645 |
have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
|
|
2646 |
by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
|
|
2647 |
also have "(\<lambda>x::real. x powr 1) \<in> o(\<lambda>a. a powr e)" if "e > 1" for e
|
|
2648 |
by (subst powr_smallo_iff) (auto simp: that filterlim_ident)
|
|
2649 |
finally show ?thesis3 using assms ev
|
|
2650 |
by (auto intro!: is_expansion_aux.intros landau_o.small.compose[of _ at_top _ "\<lambda>x. ln (b x)"]
|
|
2651 |
filterlim_compose[OF ln_at_top b_limit] is_expansion_real.intros
|
|
2652 |
elim!: eventually_mono)
|
|
2653 |
qed
|
|
2654 |
|
|
2655 |
lemma basis_wf_lift_modification:
|
|
2656 |
assumes "basis_wf (b' # b # bs)" "basis_wf (b # bs')"
|
|
2657 |
shows "basis_wf (b' # b # bs')"
|
|
2658 |
using assms by (simp add: basis_wf_many)
|
|
2659 |
|
|
2660 |
lemma default_basis_wf: "basis_wf [\<lambda>x. x]"
|
|
2661 |
by (simp add: basis_wf_Cons filterlim_ident)
|
|
2662 |
|
|
2663 |
|
|
2664 |
subsubsection \<open>Lifting expansions\<close>
|
|
2665 |
|
|
2666 |
definition is_lifting where
|
|
2667 |
"is_lifting f basis basis' \<longleftrightarrow>
|
|
2668 |
(\<forall>C. eval (f C) = eval C \<and> (is_expansion C basis \<longrightarrow> is_expansion (f C) basis') \<and>
|
|
2669 |
trimmed (f C) = trimmed C \<and> fst (dominant_term (f C)) = fst (dominant_term C))"
|
|
2670 |
|
|
2671 |
lemma is_liftingD:
|
|
2672 |
assumes "is_lifting f basis basis'"
|
|
2673 |
shows "eval (f C) = eval C" "is_expansion C basis \<Longrightarrow> is_expansion (f C) basis'"
|
|
2674 |
"trimmed (f C) \<longleftrightarrow> trimmed C" "fst (dominant_term (f C)) = fst (dominant_term C)"
|
|
2675 |
using assms [unfolded is_lifting_def] unfolding is_lifting_def by blast+
|
|
2676 |
|
|
2677 |
|
|
2678 |
definition lift_ms :: "'a :: multiseries \<Rightarrow> 'a ms" where
|
|
2679 |
"lift_ms C = MS (MSLCons (C, 0) MSLNil) (eval C)"
|
|
2680 |
|
|
2681 |
lemma is_expansion_lift:
|
|
2682 |
assumes "basis_wf (b # basis)" "is_expansion C basis"
|
|
2683 |
shows "is_expansion (lift_ms C) (b # basis)"
|
|
2684 |
proof -
|
|
2685 |
from assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
2686 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
2687 |
moreover from assms have "eval C \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e
|
|
2688 |
using that by (intro is_expansion_imp_smallo[OF _ assms(2)]) (simp_all add: basis_wf_Cons)
|
|
2689 |
ultimately show ?thesis using assms
|
|
2690 |
by (auto intro!: is_expansion_aux.intros simp: lift_ms_def is_expansion_length
|
|
2691 |
elim: eventually_mono)
|
|
2692 |
qed
|
|
2693 |
|
|
2694 |
lemma eval_lift_ms [simp]: "eval (lift_ms C) = eval C"
|
|
2695 |
by (simp add: lift_ms_def)
|
|
2696 |
|
|
2697 |
lemma is_lifting_lift:
|
|
2698 |
assumes "basis_wf (b # basis)"
|
|
2699 |
shows "is_lifting lift_ms basis (b # basis)"
|
|
2700 |
using is_expansion_lift[OF assms] unfolding is_lifting_def
|
|
2701 |
by (auto simp: lift_ms_def trimmed_ms_aux_MSLCons dominant_term_ms_aux_def case_prod_unfold)
|
|
2702 |
|
|
2703 |
|
|
2704 |
definition map_ms_aux ::
|
|
2705 |
"('a :: multiseries \<Rightarrow> 'b :: multiseries) \<Rightarrow>
|
|
2706 |
('a \<times> real) msllist \<Rightarrow> ('b \<times> real) msllist" where
|
|
2707 |
"map_ms_aux f xs = mslmap (\<lambda>(c,e). (f c, e)) xs"
|
|
2708 |
|
|
2709 |
lemma map_ms_aux_MSLNil [simp]: "map_ms_aux f MSLNil = MSLNil"
|
|
2710 |
by (simp add: map_ms_aux_def)
|
|
2711 |
|
|
2712 |
lemma map_ms_aux_MSLCons: "map_ms_aux f (MSLCons (C, e) xs) = MSLCons (f C, e) (map_ms_aux f xs)"
|
|
2713 |
by (simp add: map_ms_aux_def)
|
|
2714 |
|
|
2715 |
lemma hd_exp_map [simp]: "ms_aux_hd_exp (map_ms_aux f xs) = ms_aux_hd_exp xs"
|
|
2716 |
by (simp add: ms_aux_hd_exp_def map_ms_aux_def split: msllist.splits)
|
|
2717 |
|
|
2718 |
lemma map_ms_altdef: "map_ms f (MS xs g) = MS (map_ms_aux f xs) g"
|
|
2719 |
by (simp add: map_ms_aux_def map_prod_def)
|
|
2720 |
|
|
2721 |
lemma eval_map_ms [simp]: "eval (map_ms f F) = eval F"
|
|
2722 |
by (cases F) simp_all
|
|
2723 |
|
|
2724 |
lemma is_expansion_map_aux:
|
|
2725 |
fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
|
|
2726 |
assumes "is_expansion_aux xs g (b # basis)"
|
|
2727 |
assumes "\<And>C. is_expansion C basis \<Longrightarrow> is_expansion (f C) basis'"
|
|
2728 |
assumes "length basis' = expansion_level TYPE('b)"
|
|
2729 |
assumes "\<And>C. eval (f C) = eval C"
|
|
2730 |
shows "is_expansion_aux (map_ms_aux f xs) g (b # basis')"
|
|
2731 |
using assms(1)
|
|
2732 |
proof (coinduction arbitrary: xs g rule: is_expansion_aux_coinduct)
|
|
2733 |
case (is_expansion_aux xs g)
|
|
2734 |
show ?case
|
|
2735 |
proof (cases xs rule: ms_aux_cases)
|
|
2736 |
case MSLNil
|
|
2737 |
with is_expansion_aux show ?thesis
|
|
2738 |
by (auto simp: assms elim: is_expansion_aux.cases)
|
|
2739 |
next
|
|
2740 |
case (MSLCons c e xs')
|
|
2741 |
with is_expansion_aux show ?thesis
|
|
2742 |
by (auto elim!: is_expansion_aux_MSLCons simp: map_ms_aux_MSLCons assms)
|
|
2743 |
qed
|
|
2744 |
qed
|
|
2745 |
|
|
2746 |
lemma is_expansion_map:
|
|
2747 |
fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
|
|
2748 |
and F :: "'a ms"
|
|
2749 |
assumes "is_expansion G (b # basis)"
|
|
2750 |
assumes "\<And>C. is_expansion C basis \<Longrightarrow> is_expansion (f C) basis'"
|
|
2751 |
assumes "\<And>C. eval (f C) = eval C"
|
|
2752 |
assumes "length basis' = expansion_level TYPE('b)"
|
|
2753 |
shows "is_expansion (map_ms f G) (b # basis')"
|
|
2754 |
using assms by (cases G, simp only: map_ms_altdef) (auto intro!: is_expansion_map_aux)
|
|
2755 |
|
|
2756 |
lemma is_expansion_map_final:
|
|
2757 |
fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
|
|
2758 |
and F :: "'a ms"
|
|
2759 |
assumes "is_lifting f basis basis'"
|
|
2760 |
assumes "is_expansion G (b # basis)"
|
|
2761 |
assumes "length basis' = expansion_level TYPE('b)"
|
|
2762 |
shows "is_expansion (map_ms f G) (b # basis')"
|
|
2763 |
by (rule is_expansion_map[OF assms(2)]) (insert assms, simp_all add: is_lifting_def)
|
|
2764 |
|
|
2765 |
lemma fst_dominant_term_map_ms:
|
|
2766 |
"is_lifting f basis basis' \<Longrightarrow> fst (dominant_term (map_ms f C)) = fst (dominant_term C)"
|
|
2767 |
by (cases C)
|
|
2768 |
(simp add: dominant_term_ms_aux_def case_prod_unfold is_lifting_def split: msllist.splits)
|
|
2769 |
|
|
2770 |
lemma trimmed_map_ms:
|
|
2771 |
"is_lifting f basis basis' \<Longrightarrow> trimmed (map_ms f C) \<longleftrightarrow> trimmed C"
|
|
2772 |
by (cases C) (simp add: trimmed_ms_aux_def is_lifting_def split: msllist.splits)
|
|
2773 |
|
|
2774 |
lemma is_lifting_map:
|
|
2775 |
fixes f :: "'a :: multiseries \<Rightarrow> 'b :: multiseries"
|
|
2776 |
and F :: "'a ms"
|
|
2777 |
assumes "is_lifting f basis basis'"
|
|
2778 |
assumes "length basis' = expansion_level TYPE('b)"
|
|
2779 |
shows "is_lifting (map_ms f) (b # basis) (b # basis')"
|
|
2780 |
unfolding is_lifting_def
|
|
2781 |
by (intro allI conjI impI is_expansion_map_final[OF assms(1)])
|
|
2782 |
(insert assms, simp_all add: is_lifting_def fst_dominant_term_map_ms[OF assms(1)]
|
|
2783 |
trimmed_map_ms[OF assms(1)])
|
|
2784 |
|
|
2785 |
lemma is_lifting_id: "is_lifting (\<lambda>x. x) basis basis"
|
|
2786 |
by (simp add: is_lifting_def)
|
|
2787 |
|
|
2788 |
lemma is_lifting_trans:
|
|
2789 |
"is_lifting f basis basis' \<Longrightarrow> is_lifting g basis' basis'' \<Longrightarrow> is_lifting (\<lambda>x. g (f x)) basis basis''"
|
|
2790 |
by (auto simp: is_lifting_def)
|
|
2791 |
|
|
2792 |
lemma is_expansion_X: "is_expansion (MS (MSLCons (1 :: real, 1) MSLNil) (\<lambda>x. x)) [\<lambda>x. x]"
|
|
2793 |
proof -
|
|
2794 |
have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
|
|
2795 |
by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
|
|
2796 |
also have "(\<lambda>x::real. x powr 1) \<in> o(\<lambda>a. a powr e)" if "e > 1" for e
|
|
2797 |
by (subst powr_smallo_iff) (auto simp: that filterlim_ident)
|
|
2798 |
finally show ?thesis
|
|
2799 |
by (auto intro!: is_expansion_aux.intros is_expansion_real.intros
|
|
2800 |
eventually_mono[OF eventually_gt_at_top[of "0::real"]])
|
|
2801 |
qed
|
|
2802 |
|
|
2803 |
|
|
2804 |
inductive expands_to :: "(real \<Rightarrow> real) \<Rightarrow> 'a :: multiseries \<Rightarrow> basis \<Rightarrow> bool"
|
|
2805 |
(infix "(expands'_to)" 50) where
|
|
2806 |
"is_expansion F basis \<Longrightarrow> eventually (\<lambda>x. eval F x = f x) at_top \<Longrightarrow> (f expands_to F) basis"
|
|
2807 |
|
|
2808 |
lemma dominant_term_expands_to:
|
|
2809 |
assumes "basis_wf basis" "(f expands_to F) basis" "trimmed F"
|
|
2810 |
shows "f \<sim>[at_top] eval_monom (dominant_term F) basis"
|
|
2811 |
proof -
|
|
2812 |
from assms have "eval F \<sim>[at_top] f" by (intro asymp_equiv_refl_ev) (simp add: expands_to.simps)
|
|
2813 |
also from dominant_term[OF assms(1) _ assms(3)] assms(2)
|
|
2814 |
have "eval F \<sim>[at_top] eval_monom (dominant_term F) basis" by (simp add: expands_to.simps)
|
|
2815 |
finally show ?thesis by (subst asymp_equiv_sym) simp_all
|
|
2816 |
qed
|
|
2817 |
|
|
2818 |
lemma expands_to_cong:
|
|
2819 |
"(f expands_to F) basis \<Longrightarrow> eventually (\<lambda>x. f x = g x) at_top \<Longrightarrow> (g expands_to F) basis"
|
|
2820 |
by (auto simp: expands_to.simps elim: eventually_elim2)
|
|
2821 |
|
|
2822 |
lemma expands_to_cong':
|
|
2823 |
assumes "(f expands_to MS xs g) basis" "eventually (\<lambda>x. g x = g' x) at_top"
|
|
2824 |
shows "(f expands_to MS xs g') basis"
|
|
2825 |
proof -
|
|
2826 |
have "is_expansion_aux xs g' basis"
|
|
2827 |
by (rule is_expansion_aux_cong [OF _ assms(2)]) (insert assms(1), simp add: expands_to.simps)
|
|
2828 |
with assms show ?thesis
|
|
2829 |
by (auto simp: expands_to.simps elim: eventually_elim2)
|
|
2830 |
qed
|
|
2831 |
|
|
2832 |
lemma eval_expands_to: "(f expands_to F) basis \<Longrightarrow> eventually (\<lambda>x. eval F x = f x) at_top"
|
|
2833 |
by (simp add: expands_to.simps)
|
|
2834 |
|
|
2835 |
lemma expands_to_real_compose:
|
|
2836 |
assumes "(f expands_to (c::real)) bs"
|
|
2837 |
shows "((\<lambda>x. g (f x)) expands_to g c) bs"
|
|
2838 |
using assms by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
|
|
2839 |
|
|
2840 |
lemma expands_to_lift_compose:
|
|
2841 |
assumes "(f expands_to MS (MSLCons (C, e) xs) f') bs'"
|
|
2842 |
assumes "eventually (\<lambda>x. f' x - h x = 0) at_top" "e = 0"
|
|
2843 |
assumes "((\<lambda>x. g (h x)) expands_to G) bs" "basis_wf (b # bs)"
|
|
2844 |
shows "((\<lambda>x. g (f x)) expands_to lift_ms G) (b # bs)"
|
|
2845 |
proof -
|
|
2846 |
from assms have "\<forall>\<^sub>F x in at_top. f' x = f x" "\<forall>\<^sub>F x in at_top. eval G x = g (h x)"
|
|
2847 |
by (auto simp: expands_to.simps)
|
|
2848 |
with assms(2) have "\<forall>\<^sub>F x in at_top. eval G x = g (f x)"
|
|
2849 |
by eventually_elim simp
|
|
2850 |
with assms show ?thesis
|
|
2851 |
by (intro expands_to.intros is_expansion_lift) (auto simp: expands_to.simps)
|
|
2852 |
qed
|
|
2853 |
|
|
2854 |
lemma expands_to_zero:
|
|
2855 |
"basis_wf basis \<Longrightarrow> length basis = expansion_level TYPE('a) \<Longrightarrow>
|
|
2856 |
((\<lambda>_. 0) expands_to (zero_expansion :: 'a :: multiseries)) basis"
|
|
2857 |
by (auto simp add: expands_to.simps is_expansion_zero)
|
|
2858 |
|
|
2859 |
lemma expands_to_const:
|
|
2860 |
"basis_wf basis \<Longrightarrow> length basis = expansion_level TYPE('a) \<Longrightarrow>
|
|
2861 |
((\<lambda>_. c) expands_to (const_expansion c :: 'a :: multiseries)) basis"
|
|
2862 |
by (auto simp add: expands_to.simps is_expansion_const)
|
|
2863 |
|
|
2864 |
lemma expands_to_X: "((\<lambda>x. x) expands_to MS (MSLCons (1 :: real, 1) MSLNil) (\<lambda>x. x)) [\<lambda>x. x]"
|
|
2865 |
using is_expansion_X by (simp add: expands_to.simps)
|
|
2866 |
|
|
2867 |
lemma expands_to_uminus:
|
|
2868 |
"basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> ((\<lambda>x. -f x) expands_to -F) basis"
|
|
2869 |
by (auto simp: expands_to.simps is_expansion_uminus)
|
|
2870 |
|
|
2871 |
lemma expands_to_add:
|
|
2872 |
"basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
|
|
2873 |
((\<lambda>x. f x + g x) expands_to F + G) basis"
|
|
2874 |
by (auto simp: expands_to.simps is_expansion_add elim: eventually_elim2)
|
|
2875 |
|
|
2876 |
lemma expands_to_minus:
|
|
2877 |
"basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
|
|
2878 |
((\<lambda>x. f x - g x) expands_to F - G) basis"
|
|
2879 |
by (auto simp: expands_to.simps is_expansion_minus elim: eventually_elim2)
|
|
2880 |
|
|
2881 |
lemma expands_to_mult:
|
|
2882 |
"basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
|
|
2883 |
((\<lambda>x. f x * g x) expands_to F * G) basis"
|
|
2884 |
by (auto simp: expands_to.simps is_expansion_mult elim: eventually_elim2)
|
|
2885 |
|
|
2886 |
lemma expands_to_inverse:
|
|
2887 |
"trimmed F \<Longrightarrow> basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow>
|
|
2888 |
((\<lambda>x. inverse (f x)) expands_to inverse F) basis"
|
|
2889 |
by (auto simp: expands_to.simps is_expansion_inverse)
|
|
2890 |
|
|
2891 |
lemma expands_to_divide:
|
|
2892 |
"trimmed G \<Longrightarrow> basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> (g expands_to G) basis \<Longrightarrow>
|
|
2893 |
((\<lambda>x. f x / g x) expands_to F / G) basis"
|
|
2894 |
by (auto simp: expands_to.simps is_expansion_divide elim: eventually_elim2)
|
|
2895 |
|
|
2896 |
lemma expands_to_powr_0:
|
|
2897 |
"eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> g \<equiv> g \<Longrightarrow> basis_wf bs \<Longrightarrow>
|
|
2898 |
length bs = expansion_level TYPE('a) \<Longrightarrow>
|
|
2899 |
((\<lambda>x. f x powr g x) expands_to (zero_expansion :: 'a :: multiseries)) bs"
|
|
2900 |
by (erule (1) expands_to_cong[OF expands_to_zero]) simp_all
|
|
2901 |
|
|
2902 |
lemma expands_to_powr_const:
|
|
2903 |
"trimmed_pos F \<Longrightarrow> basis_wf basis \<Longrightarrow> (f expands_to F) basis \<Longrightarrow> p \<equiv> p \<Longrightarrow>
|
|
2904 |
((\<lambda>x. f x powr p) expands_to powr_expansion abort F p) basis"
|
|
2905 |
by (auto simp: expands_to.simps is_expansion_powr trimmed_pos_def elim: eventually_mono)
|
|
2906 |
|
|
2907 |
lemma expands_to_powr_const_0:
|
|
2908 |
"eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> basis_wf bs \<Longrightarrow>
|
|
2909 |
length bs = expansion_level TYPE('a :: multiseries) \<Longrightarrow>
|
|
2910 |
p \<equiv> p \<Longrightarrow> ((\<lambda>x. f x powr p) expands_to (zero_expansion :: 'a)) bs"
|
|
2911 |
by (rule expands_to_cong[OF expands_to_zero]) auto
|
|
2912 |
|
|
2913 |
|
|
2914 |
lemma expands_to_powr:
|
|
2915 |
assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'"
|
|
2916 |
assumes "((\<lambda>x. exp (ln (f x) * g x)) expands_to E) basis"
|
|
2917 |
shows "((\<lambda>x. f x powr g x) expands_to E) basis"
|
|
2918 |
using assms(4)
|
|
2919 |
proof (rule expands_to_cong)
|
|
2920 |
from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4)
|
|
2921 |
show "eventually (\<lambda>x. exp (ln (f x) * g x) = f x powr g x) at_top"
|
|
2922 |
by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2)
|
|
2923 |
qed
|
|
2924 |
|
|
2925 |
lemma expands_to_ln_powr:
|
|
2926 |
assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'"
|
|
2927 |
assumes "((\<lambda>x. ln (f x) * g x) expands_to E) basis"
|
|
2928 |
shows "((\<lambda>x. ln (f x powr g x)) expands_to E) basis"
|
|
2929 |
using assms(4)
|
|
2930 |
proof (rule expands_to_cong)
|
|
2931 |
from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4)
|
|
2932 |
show "eventually (\<lambda>x. ln (f x) * g x = ln (f x powr g x)) at_top"
|
|
2933 |
by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2)
|
|
2934 |
qed
|
|
2935 |
|
|
2936 |
lemma expands_to_exp_ln:
|
|
2937 |
assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
|
|
2938 |
shows "((\<lambda>x. exp (ln (f x))) expands_to F) basis"
|
|
2939 |
using assms(3)
|
|
2940 |
proof (rule expands_to_cong)
|
|
2941 |
from eval_pos_if_dominant_term_pos[of basis F] assms
|
|
2942 |
show "eventually (\<lambda>x. f x = exp (ln (f x))) at_top"
|
|
2943 |
by (auto simp: expands_to.simps trimmed_pos_def powr_def elim: eventually_elim2)
|
|
2944 |
qed
|
|
2945 |
|
|
2946 |
lemma expands_to_power:
|
|
2947 |
assumes "trimmed F" "basis_wf basis" "(f expands_to F) basis"
|
|
2948 |
shows "((\<lambda>x. f x ^ n) expands_to power_expansion abort F n) basis"
|
|
2949 |
using assms by (auto simp: expands_to.simps is_expansion_power elim: eventually_mono)
|
|
2950 |
|
|
2951 |
lemma expands_to_power_0:
|
|
2952 |
assumes "eventually (\<lambda>x. f x = 0) at_top" "basis_wf basis"
|
|
2953 |
"length basis = expansion_level TYPE('a :: multiseries)" "n \<equiv> n"
|
|
2954 |
shows "((\<lambda>x. f x ^ n) expands_to (const_expansion (0 ^ n) :: 'a)) basis"
|
|
2955 |
by (rule expands_to_cong[OF expands_to_const]) (insert assms, auto elim: eventually_mono)
|
|
2956 |
|
|
2957 |
lemma expands_to_0th_root:
|
|
2958 |
assumes "n = 0" "basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)" "f \<equiv> f"
|
|
2959 |
shows "((\<lambda>x. root n (f x)) expands_to (zero_expansion :: 'a)) basis"
|
|
2960 |
by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto)
|
|
2961 |
|
|
2962 |
lemma expands_to_root_0:
|
|
2963 |
assumes "n > 0" "eventually (\<lambda>x. f x = 0) at_top"
|
|
2964 |
"basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)"
|
|
2965 |
shows "((\<lambda>x. root n (f x)) expands_to (zero_expansion :: 'a)) basis"
|
|
2966 |
by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto elim: eventually_mono)
|
|
2967 |
|
|
2968 |
lemma expands_to_root:
|
|
2969 |
assumes "n > 0" "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
|
|
2970 |
shows "((\<lambda>x. root n (f x)) expands_to powr_expansion False F (inverse (real n))) basis"
|
|
2971 |
proof -
|
|
2972 |
have "((\<lambda>x. f x powr (inverse (real n))) expands_to
|
|
2973 |
powr_expansion False F (inverse (real n))) basis"
|
|
2974 |
using assms(2-) by (rule expands_to_powr_const)
|
|
2975 |
moreover have "eventually (\<lambda>x. f x powr (inverse (real n)) = root n (f x)) at_top"
|
|
2976 |
using eval_pos_if_dominant_term_pos[of basis F] assms
|
|
2977 |
by (auto simp: trimmed_pos_def expands_to.simps root_powr_inverse field_simps
|
|
2978 |
elim: eventually_elim2)
|
|
2979 |
ultimately show ?thesis by (rule expands_to_cong)
|
|
2980 |
qed
|
|
2981 |
|
|
2982 |
lemma expands_to_root_neg:
|
|
2983 |
assumes "n > 0" "trimmed_neg F" "basis_wf basis" "(f expands_to F) basis"
|
|
2984 |
shows "((\<lambda>x. root n (f x)) expands_to
|
|
2985 |
-powr_expansion False (-F) (inverse (real n))) basis"
|
|
2986 |
proof (rule expands_to_cong)
|
|
2987 |
show "((\<lambda>x. -root n (-f x)) expands_to
|
|
2988 |
-powr_expansion False (-F) (inverse (real n))) basis"
|
|
2989 |
using assms by (intro expands_to_uminus expands_to_root trimmed_pos_uminus) auto
|
|
2990 |
qed (simp_all add: real_root_minus)
|
|
2991 |
|
|
2992 |
lemma expands_to_sqrt:
|
|
2993 |
assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
|
|
2994 |
shows "((\<lambda>x. sqrt (f x)) expands_to powr_expansion False F (1/2)) basis"
|
|
2995 |
using expands_to_root[of 2 F basis f] assms by (simp add: sqrt_def)
|
|
2996 |
|
|
2997 |
lemma expands_to_exp_real:
|
|
2998 |
"(f expands_to (F :: real)) basis \<Longrightarrow> ((\<lambda>x. exp (f x)) expands_to exp F) basis"
|
|
2999 |
by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
|
|
3000 |
|
|
3001 |
lemma expands_to_exp_MSLNil:
|
|
3002 |
assumes "(f expands_to (MS MSLNil f' :: 'a :: multiseries ms)) bs" "basis_wf bs"
|
|
3003 |
shows "((\<lambda>x. exp (f x)) expands_to (const_expansion 1 :: 'a ms)) bs"
|
|
3004 |
proof -
|
|
3005 |
from assms have
|
|
3006 |
"eventually (\<lambda>x. f' x = 0) at_top" "eventually (\<lambda>x. f' x = f x) at_top"
|
|
3007 |
and [simp]: "length bs = Suc (expansion_level(TYPE('a)))"
|
|
3008 |
by (auto simp: expands_to.simps elim: is_expansion_aux.cases)
|
|
3009 |
from this(1,2) have "eventually (\<lambda>x. 1 = exp (f x)) at_top"
|
|
3010 |
by eventually_elim simp
|
|
3011 |
thus ?thesis by (auto simp: expands_to.simps intro!: is_expansion_const assms)
|
|
3012 |
qed
|
|
3013 |
|
|
3014 |
lemma expands_to_exp_zero:
|
|
3015 |
"(g expands_to MS xs f) basis \<Longrightarrow> eventually (\<lambda>x. f x = 0) at_top \<Longrightarrow> basis_wf basis \<Longrightarrow>
|
|
3016 |
length basis = expansion_level TYPE('a :: multiseries) \<Longrightarrow>
|
|
3017 |
((\<lambda>x. exp (g x)) expands_to (const_expansion 1 :: 'a)) basis"
|
|
3018 |
by (auto simp: expands_to.simps intro!: is_expansion_const elim: eventually_elim2)
|
|
3019 |
|
|
3020 |
lemma expands_to_sin_real:
|
|
3021 |
"(f expands_to (F :: real)) basis \<Longrightarrow> ((\<lambda>x. sin (f x)) expands_to sin F) basis"
|
|
3022 |
by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
|
|
3023 |
|
|
3024 |
lemma expands_to_cos_real:
|
|
3025 |
"(f expands_to (F :: real)) basis \<Longrightarrow> ((\<lambda>x. cos (f x)) expands_to cos F) basis"
|
|
3026 |
by (auto simp: expands_to.simps is_expansion_real.simps elim: eventually_mono)
|
|
3027 |
|
|
3028 |
lemma expands_to_sin_MSLNil:
|
|
3029 |
"(f expands_to MS (MSLNil:: ('a \<times> real) msllist) g) basis \<Longrightarrow> basis_wf basis \<Longrightarrow>
|
|
3030 |
((\<lambda>x. sin (f x)) expands_to MS (MSLNil :: ('a :: multiseries \<times> real) msllist) (\<lambda>x. 0)) basis"
|
|
3031 |
by (auto simp: expands_to.simps dominant_term_ms_aux_def intro!: is_expansion_aux.intros
|
|
3032 |
elim: eventually_elim2 is_expansion_aux.cases)
|
|
3033 |
|
|
3034 |
lemma expands_to_cos_MSLNil:
|
|
3035 |
"(f expands_to MS (MSLNil:: ('a :: multiseries \<times> real) msllist) g) basis \<Longrightarrow> basis_wf basis \<Longrightarrow>
|
|
3036 |
((\<lambda>x. cos (f x)) expands_to (const_expansion 1 :: 'a ms)) basis"
|
|
3037 |
by (auto simp: expands_to.simps dominant_term_ms_aux_def const_expansion_ms_def
|
|
3038 |
intro!: const_ms_aux elim: is_expansion_aux.cases eventually_elim2)
|
|
3039 |
|
|
3040 |
lemma sin_ms_aux':
|
|
3041 |
assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis"
|
|
3042 |
shows "is_expansion_aux (sin_ms_aux' xs) (\<lambda>x. sin (f x)) basis"
|
|
3043 |
unfolding sin_ms_aux'_def sin_conv_pre_sin power2_eq_square using assms
|
|
3044 |
by (intro times_ms_aux powser_ms_aux[unfolded o_def] convergent_powser_sin_series_stream)
|
|
3045 |
(auto simp: hd_exp_times add_neg_neg split: option.splits)
|
|
3046 |
|
|
3047 |
lemma cos_ms_aux':
|
|
3048 |
assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis"
|
|
3049 |
shows "is_expansion_aux (cos_ms_aux' xs) (\<lambda>x. cos (f x)) basis"
|
|
3050 |
unfolding cos_ms_aux'_def cos_conv_pre_cos power2_eq_square using assms
|
|
3051 |
by (intro times_ms_aux powser_ms_aux[unfolded o_def] convergent_powser_cos_series_stream)
|
|
3052 |
(auto simp: hd_exp_times add_neg_neg split: option.splits)
|
|
3053 |
|
|
3054 |
lemma expands_to_sin_ms_neg_exp:
|
|
3055 |
assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis"
|
|
3056 |
shows "((\<lambda>x. sin (f x)) expands_to MS (sin_ms_aux' (MSLCons (C, e) xs)) (\<lambda>x. sin (g x))) basis"
|
|
3057 |
using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
|
|
3058 |
intro!: sin_ms_aux' elim: eventually_mono)
|
|
3059 |
|
|
3060 |
lemma expands_to_cos_ms_neg_exp:
|
|
3061 |
assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis"
|
|
3062 |
shows "((\<lambda>x. cos (f x)) expands_to MS (cos_ms_aux' (MSLCons (C, e) xs)) (\<lambda>x. cos (g x))) basis"
|
|
3063 |
using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
|
|
3064 |
intro!: cos_ms_aux' elim: eventually_mono)
|
|
3065 |
|
|
3066 |
lemma is_expansion_sin_ms_zero_exp:
|
|
3067 |
fixes F :: "('a :: multiseries \<times> real) msllist"
|
|
3068 |
assumes basis: "basis_wf (b # basis)"
|
|
3069 |
assumes F: "is_expansion (MS (MSLCons (C, 0) xs) f) (b # basis)"
|
|
3070 |
assumes Sin: "((\<lambda>x. sin (eval C x)) expands_to (Sin :: 'a)) basis"
|
|
3071 |
assumes Cos: "((\<lambda>x. cos (eval C x)) expands_to (Cos :: 'a)) basis"
|
|
3072 |
shows "is_expansion
|
|
3073 |
(MS (plus_ms_aux (scale_shift_ms_aux (Sin, 0) (cos_ms_aux' xs))
|
|
3074 |
(scale_shift_ms_aux (Cos, 0) (sin_ms_aux' xs)))
|
|
3075 |
(\<lambda>x. sin (f x))) (b # basis)" (is "is_expansion (MS ?A _) _")
|
|
3076 |
proof -
|
|
3077 |
let ?g = "\<lambda>x. f x - eval C x * b x powr 0"
|
|
3078 |
let ?h = "\<lambda>x. eval Sin x * b x powr 0 * cos (?g x) + eval Cos x * b x powr 0 * sin (?g x)"
|
|
3079 |
from basis have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
3080 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
3081 |
from F have F': "is_expansion_aux xs ?g (b # basis)"
|
|
3082 |
"ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis"
|
|
3083 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
3084 |
have "is_expansion_aux ?A ?h (b # basis)"
|
|
3085 |
using basis Sin Cos unfolding F'(1)
|
|
3086 |
by (intro plus_ms_aux scale_shift_ms_aux cos_ms_aux' sin_ms_aux' F')
|
|
3087 |
(simp_all add: expands_to.simps)
|
|
3088 |
moreover from b_pos eval_expands_to[OF Sin] eval_expands_to[OF Cos]
|
|
3089 |
have "eventually (\<lambda>x. ?h x = sin (f x)) at_top"
|
|
3090 |
by eventually_elim (simp add: sin_add [symmetric])
|
|
3091 |
ultimately have "is_expansion_aux ?A (\<lambda>x. sin (f x)) (b # basis)"
|
|
3092 |
by (rule is_expansion_aux_cong)
|
|
3093 |
thus ?thesis by simp
|
|
3094 |
qed
|
|
3095 |
|
|
3096 |
lemma is_expansion_cos_ms_zero_exp:
|
|
3097 |
fixes F :: "('a :: multiseries \<times> real) msllist"
|
|
3098 |
assumes basis: "basis_wf (b # basis)"
|
|
3099 |
assumes F: "is_expansion (MS (MSLCons (C, 0) xs) f) (b # basis)"
|
|
3100 |
assumes Sin: "((\<lambda>x. sin (eval C x)) expands_to (Sin :: 'a)) basis"
|
|
3101 |
assumes Cos: "((\<lambda>x. cos (eval C x)) expands_to (Cos :: 'a)) basis"
|
|
3102 |
shows "is_expansion
|
|
3103 |
(MS (minus_ms_aux (scale_shift_ms_aux (Cos, 0) (cos_ms_aux' xs))
|
|
3104 |
(scale_shift_ms_aux (Sin, 0) (sin_ms_aux' xs)))
|
|
3105 |
(\<lambda>x. cos (f x))) (b # basis)" (is "is_expansion (MS ?A _) _")
|
|
3106 |
proof -
|
|
3107 |
let ?g = "\<lambda>x. f x - eval C x * b x powr 0"
|
|
3108 |
let ?h = "\<lambda>x. eval Cos x * b x powr 0 * cos (?g x) - eval Sin x * b x powr 0 * sin (?g x)"
|
|
3109 |
from basis have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
3110 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
3111 |
from F have F': "is_expansion_aux xs ?g (b # basis)"
|
|
3112 |
"ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis"
|
|
3113 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
3114 |
have "is_expansion_aux ?A ?h (b # basis)"
|
|
3115 |
using basis Sin Cos unfolding F'(1)
|
|
3116 |
by (intro minus_ms_aux scale_shift_ms_aux cos_ms_aux' sin_ms_aux' F')
|
|
3117 |
(simp_all add: expands_to.simps)
|
|
3118 |
moreover from b_pos eval_expands_to[OF Sin] eval_expands_to[OF Cos]
|
|
3119 |
have "eventually (\<lambda>x. ?h x = cos (f x)) at_top"
|
|
3120 |
by eventually_elim (simp add: cos_add [symmetric])
|
|
3121 |
ultimately have "is_expansion_aux ?A (\<lambda>x. cos (f x)) (b # basis)"
|
|
3122 |
by (rule is_expansion_aux_cong)
|
|
3123 |
thus ?thesis by simp
|
|
3124 |
qed
|
|
3125 |
|
|
3126 |
lemma expands_to_sin_ms_zero_exp:
|
|
3127 |
assumes e: "e = 0" and basis: "basis_wf (b # basis)"
|
|
3128 |
assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
3129 |
assumes Sin: "((\<lambda>x. sin (c x)) expands_to Sin) basis"
|
|
3130 |
assumes Cos: "((\<lambda>x. cos (c x)) expands_to Cos) basis"
|
|
3131 |
assumes C: "eval C = c"
|
|
3132 |
shows "((\<lambda>x. sin (f x)) expands_to
|
|
3133 |
MS (plus_ms_aux (scale_shift_ms_aux (Sin, 0) (cos_ms_aux' xs))
|
|
3134 |
(scale_shift_ms_aux (Cos, 0) (sin_ms_aux' xs)))
|
|
3135 |
(\<lambda>x. sin (g x))) (b # basis)"
|
|
3136 |
using is_expansion_sin_ms_zero_exp[of b basis C xs g Sin Cos] assms
|
|
3137 |
by (auto simp: expands_to.simps elim: eventually_mono)
|
|
3138 |
|
|
3139 |
lemma expands_to_cos_ms_zero_exp:
|
|
3140 |
assumes e: "e = 0" and basis: "basis_wf (b # basis)"
|
|
3141 |
assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
3142 |
assumes Sin: "((\<lambda>x. sin (c x)) expands_to Sin) basis"
|
|
3143 |
assumes Cos: "((\<lambda>x. cos (c x)) expands_to Cos) basis"
|
|
3144 |
assumes C: "eval C = c"
|
|
3145 |
shows "((\<lambda>x. cos (f x)) expands_to
|
|
3146 |
MS (minus_ms_aux (scale_shift_ms_aux (Cos, 0) (cos_ms_aux' xs))
|
|
3147 |
(scale_shift_ms_aux (Sin, 0) (sin_ms_aux' xs)))
|
|
3148 |
(\<lambda>x. cos (g x))) (b # basis)"
|
|
3149 |
using is_expansion_cos_ms_zero_exp[of b basis C xs g Sin Cos] assms
|
|
3150 |
by (auto simp: expands_to.simps elim: eventually_mono)
|
|
3151 |
|
|
3152 |
|
|
3153 |
|
|
3154 |
lemma expands_to_sgn_zero:
|
|
3155 |
assumes "eventually (\<lambda>x. f x = 0) at_top" "basis_wf basis"
|
|
3156 |
"length basis = expansion_level TYPE('a :: multiseries)"
|
|
3157 |
shows "((\<lambda>x. sgn (f x)) expands_to (zero_expansion :: 'a)) basis"
|
|
3158 |
by (rule expands_to_cong[OF expands_to_zero]) (insert assms, auto simp: sgn_eq_0_iff)
|
|
3159 |
|
|
3160 |
lemma expands_to_sgn_pos:
|
|
3161 |
assumes "trimmed_pos F" "(f expands_to F) basis" "basis_wf basis"
|
|
3162 |
"length basis = expansion_level TYPE('a :: multiseries)"
|
|
3163 |
shows "((\<lambda>x. sgn (f x)) expands_to (const_expansion 1 :: 'a)) basis"
|
|
3164 |
proof (rule expands_to_cong[OF expands_to_const])
|
|
3165 |
from trimmed_imp_eventually_sgn[of basis F] assms
|
|
3166 |
have "eventually (\<lambda>x. sgn (eval F x) = 1) at_top"
|
|
3167 |
by (simp add: expands_to.simps trimmed_pos_def)
|
|
3168 |
moreover from assms have "eventually (\<lambda>x. eval F x = f x) at_top"
|
|
3169 |
by (simp add: expands_to.simps)
|
|
3170 |
ultimately show "eventually (\<lambda>x. 1 = sgn (f x)) at_top" by eventually_elim simp
|
|
3171 |
qed (insert assms, auto)
|
|
3172 |
|
|
3173 |
lemma expands_to_sgn_neg:
|
|
3174 |
assumes "trimmed_neg F" "(f expands_to F) basis" "basis_wf basis"
|
|
3175 |
"length basis = expansion_level TYPE('a :: multiseries)"
|
|
3176 |
shows "((\<lambda>x. sgn (f x)) expands_to (const_expansion (-1) :: 'a)) basis"
|
|
3177 |
proof (rule expands_to_cong[OF expands_to_const])
|
|
3178 |
from trimmed_imp_eventually_sgn[of basis F] assms
|
|
3179 |
have "eventually (\<lambda>x. sgn (eval F x) = -1) at_top"
|
|
3180 |
by (simp add: expands_to.simps trimmed_neg_def)
|
|
3181 |
moreover from assms have "eventually (\<lambda>x. eval F x = f x) at_top"
|
|
3182 |
by (simp add: expands_to.simps)
|
|
3183 |
ultimately show "eventually (\<lambda>x. -1 = sgn (f x)) at_top" by eventually_elim simp
|
|
3184 |
qed (insert assms, auto)
|
|
3185 |
|
|
3186 |
|
|
3187 |
|
|
3188 |
lemma expands_to_abs_pos:
|
|
3189 |
assumes "trimmed_pos F" "basis_wf basis" "(f expands_to F) basis"
|
|
3190 |
shows "((\<lambda>x. abs (f x)) expands_to F) basis"
|
|
3191 |
using assms(3)
|
|
3192 |
proof (rule expands_to_cong)
|
|
3193 |
from trimmed_imp_eventually_sgn[of basis F] assms
|
|
3194 |
have "eventually (\<lambda>x. sgn (eval F x) = 1) at_top"
|
|
3195 |
by (simp add: expands_to.simps trimmed_pos_def)
|
|
3196 |
with assms(3) show "eventually (\<lambda>x. f x = abs (f x)) at_top"
|
|
3197 |
by (auto simp: expands_to.simps sgn_if split: if_splits elim: eventually_elim2)
|
|
3198 |
qed
|
|
3199 |
|
|
3200 |
lemma expands_to_abs_neg:
|
|
3201 |
assumes "trimmed_neg F" "basis_wf basis" "(f expands_to F) basis"
|
|
3202 |
shows "((\<lambda>x. abs (f x)) expands_to -F) basis"
|
|
3203 |
using expands_to_uminus[OF assms(2,3)]
|
|
3204 |
proof (rule expands_to_cong)
|
|
3205 |
from trimmed_imp_eventually_sgn[of basis F] assms
|
|
3206 |
have "eventually (\<lambda>x. sgn (eval F x) = -1) at_top"
|
|
3207 |
by (simp add: expands_to.simps trimmed_neg_def)
|
|
3208 |
with assms(3) show "eventually (\<lambda>x. -f x = abs (f x)) at_top"
|
|
3209 |
by (auto simp: expands_to.simps sgn_if split: if_splits elim: eventually_elim2)
|
|
3210 |
qed
|
|
3211 |
|
|
3212 |
lemma expands_to_imp_eventually_nz:
|
|
3213 |
assumes "basis_wf basis" "(f expands_to F) basis" "trimmed F"
|
|
3214 |
shows "eventually (\<lambda>x. f x \<noteq> 0) at_top"
|
|
3215 |
using trimmed_imp_eventually_nz[OF assms(1), of F] assms(2,3)
|
|
3216 |
by (auto simp: expands_to.simps elim: eventually_elim2)
|
|
3217 |
|
|
3218 |
lemma expands_to_imp_eventually_pos:
|
|
3219 |
assumes "basis_wf basis" "(f expands_to F) basis" "trimmed_pos F"
|
|
3220 |
shows "eventually (\<lambda>x. f x > 0) at_top"
|
|
3221 |
using eval_pos_if_dominant_term_pos[of basis F] assms
|
|
3222 |
by (auto simp: expands_to.simps trimmed_pos_def elim: eventually_elim2)
|
|
3223 |
|
|
3224 |
lemma expands_to_imp_eventually_neg:
|
|
3225 |
assumes "basis_wf basis" "(f expands_to F) basis" "trimmed_neg F"
|
|
3226 |
shows "eventually (\<lambda>x. f x < 0) at_top"
|
|
3227 |
using eval_pos_if_dominant_term_pos[of basis "-F"] assms
|
|
3228 |
by (auto simp: expands_to.simps trimmed_neg_def is_expansion_uminus elim: eventually_elim2)
|
|
3229 |
|
|
3230 |
lemma expands_to_imp_eventually_gt:
|
|
3231 |
assumes "basis_wf basis" "((\<lambda>x. f x - g x) expands_to F) basis" "trimmed_pos F"
|
|
3232 |
shows "eventually (\<lambda>x. f x > g x) at_top"
|
|
3233 |
using expands_to_imp_eventually_pos[OF assms] by simp
|
|
3234 |
|
|
3235 |
lemma expands_to_imp_eventually_lt:
|
|
3236 |
assumes "basis_wf basis" "((\<lambda>x. f x - g x) expands_to F) basis" "trimmed_neg F"
|
|
3237 |
shows "eventually (\<lambda>x. f x < g x) at_top"
|
|
3238 |
using expands_to_imp_eventually_neg[OF assms] by simp
|
|
3239 |
|
|
3240 |
lemma eventually_diff_zero_imp_eq:
|
|
3241 |
fixes f g :: "real \<Rightarrow> real"
|
|
3242 |
assumes "eventually (\<lambda>x. f x - g x = 0) at_top"
|
|
3243 |
shows "eventually (\<lambda>x. f x = g x) at_top"
|
|
3244 |
using assms by eventually_elim simp
|
|
3245 |
|
|
3246 |
lemma smallo_trimmed_imp_eventually_less:
|
|
3247 |
fixes f g :: "real \<Rightarrow> real"
|
|
3248 |
assumes "f \<in> o(g)" "(g expands_to G) bs" "basis_wf bs" "trimmed_pos G"
|
|
3249 |
shows "eventually (\<lambda>x. f x < g x) at_top"
|
|
3250 |
proof -
|
|
3251 |
from assms(2-4) have pos: "eventually (\<lambda>x. g x > 0) at_top"
|
|
3252 |
using expands_to_imp_eventually_pos by blast
|
|
3253 |
have "(1 / 2 :: real) > 0" by simp
|
|
3254 |
from landau_o.smallD[OF assms(1) this] and pos
|
|
3255 |
show ?thesis by eventually_elim (simp add: abs_if split: if_splits)
|
|
3256 |
qed
|
|
3257 |
|
|
3258 |
lemma smallo_trimmed_imp_eventually_greater:
|
|
3259 |
fixes f g :: "real \<Rightarrow> real"
|
|
3260 |
assumes "f \<in> o(g)" "(g expands_to G) bs" "basis_wf bs" "trimmed_neg G"
|
|
3261 |
shows "eventually (\<lambda>x. f x > g x) at_top"
|
|
3262 |
proof -
|
|
3263 |
from assms(2-4) have pos: "eventually (\<lambda>x. g x < 0) at_top"
|
|
3264 |
using expands_to_imp_eventually_neg by blast
|
|
3265 |
have "(1 / 2 :: real) > 0" by simp
|
|
3266 |
from landau_o.smallD[OF assms(1) this] and pos
|
|
3267 |
show ?thesis by eventually_elim (simp add: abs_if split: if_splits)
|
|
3268 |
qed
|
|
3269 |
|
|
3270 |
lemma expands_to_min_lt:
|
|
3271 |
assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x < g x) at_top"
|
|
3272 |
shows "((\<lambda>x. min (f x) (g x)) expands_to F) basis"
|
|
3273 |
using assms(1)
|
|
3274 |
by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono)
|
|
3275 |
|
|
3276 |
lemma expands_to_min_eq:
|
|
3277 |
assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x = g x) at_top"
|
|
3278 |
shows "((\<lambda>x. min (f x) (g x)) expands_to F) basis"
|
|
3279 |
using assms(1)
|
|
3280 |
by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono)
|
|
3281 |
|
|
3282 |
lemma expands_to_min_gt:
|
|
3283 |
assumes "(g expands_to G) basis" "eventually (\<lambda>x. f x > g x) at_top"
|
|
3284 |
shows "((\<lambda>x. min (f x) (g x)) expands_to G) basis"
|
|
3285 |
using assms(1)
|
|
3286 |
by (rule expands_to_cong) (insert assms(2), auto simp: min_def elim: eventually_mono)
|
|
3287 |
|
|
3288 |
lemma expands_to_max_gt:
|
|
3289 |
assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x > g x) at_top"
|
|
3290 |
shows "((\<lambda>x. max (f x) (g x)) expands_to F) basis"
|
|
3291 |
using assms(1)
|
|
3292 |
by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono)
|
|
3293 |
|
|
3294 |
lemma expands_to_max_eq:
|
|
3295 |
assumes "(f expands_to F) basis" "eventually (\<lambda>x. f x = g x) at_top"
|
|
3296 |
shows "((\<lambda>x. max (f x) (g x)) expands_to F) basis"
|
|
3297 |
using assms(1)
|
|
3298 |
by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono)
|
|
3299 |
|
|
3300 |
lemma expands_to_max_lt:
|
|
3301 |
assumes "(g expands_to G) basis" "eventually (\<lambda>x. f x < g x) at_top"
|
|
3302 |
shows "((\<lambda>x. max (f x) (g x)) expands_to G) basis"
|
|
3303 |
using assms(1)
|
|
3304 |
by (rule expands_to_cong) (insert assms(2), auto simp: max_def elim: eventually_mono)
|
|
3305 |
|
|
3306 |
|
|
3307 |
lemma expands_to_lift:
|
|
3308 |
"is_lifting f basis basis' \<Longrightarrow> (c expands_to C) basis \<Longrightarrow> (c expands_to (f C)) basis'"
|
|
3309 |
by (simp add: is_lifting_def expands_to.simps)
|
|
3310 |
|
|
3311 |
lemma expands_to_basic_powr:
|
|
3312 |
assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)"
|
|
3313 |
shows "((\<lambda>x. b x powr e) expands_to
|
|
3314 |
MS (MSLCons (const_expansion 1 :: 'a, e) MSLNil) (\<lambda>x. b x powr e)) (b # basis)"
|
|
3315 |
using assms by (auto simp: expands_to.simps basis_wf_Cons powr_smallo_iff
|
|
3316 |
intro!: is_expansion_aux.intros is_expansion_const powr_smallo_iff)
|
|
3317 |
|
|
3318 |
lemma expands_to_basic_inverse:
|
|
3319 |
assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)"
|
|
3320 |
shows "((\<lambda>x. inverse (b x)) expands_to
|
|
3321 |
MS (MSLCons (const_expansion 1 :: 'a, -1) MSLNil) (\<lambda>x. b x powr -1)) (b # basis)"
|
|
3322 |
proof -
|
|
3323 |
from assms have "eventually (\<lambda>x. b x > 0) at_top"
|
|
3324 |
by (simp add: basis_wf_Cons filterlim_at_top_dense)
|
|
3325 |
moreover have "(\<lambda>a. inverse (a powr 1)) \<in> o(\<lambda>a. a powr e')" if "e' > -1" for e' :: real
|
|
3326 |
using that by (simp add: landau_o.small.inverse_eq2 powr_add [symmetric] one_smallo_powr)
|
|
3327 |
ultimately show ?thesis using assms
|
|
3328 |
by (auto simp: expands_to.simps basis_wf_Cons powr_minus
|
|
3329 |
elim: eventually_mono
|
|
3330 |
intro!: is_expansion_aux.intros is_expansion_const
|
|
3331 |
landau_o.small.compose[of _ at_top _ b])
|
|
3332 |
qed
|
|
3333 |
|
|
3334 |
lemma expands_to_div':
|
|
3335 |
assumes "basis_wf basis" "(f expands_to F) basis" "((\<lambda>x. inverse (g x)) expands_to G) basis"
|
|
3336 |
shows "((\<lambda>x. f x / g x) expands_to F * G) basis"
|
|
3337 |
using expands_to_mult[OF assms] by (simp add: divide_simps)
|
|
3338 |
|
|
3339 |
lemma expands_to_basic:
|
|
3340 |
assumes "basis_wf (b # basis)" "length basis = expansion_level TYPE('a::multiseries)"
|
|
3341 |
shows "(b expands_to MS (MSLCons (const_expansion 1 :: 'a, 1) MSLNil) b) (b # basis)"
|
|
3342 |
proof -
|
|
3343 |
from assms have "eventually (\<lambda>x. b x > 0) at_top"
|
|
3344 |
by (simp add: basis_wf_Cons filterlim_at_top_dense)
|
|
3345 |
moreover {
|
|
3346 |
fix e' :: real assume e': "e' > 1"
|
|
3347 |
have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
|
|
3348 |
by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
|
|
3349 |
also have "(\<lambda>x. x powr 1) \<in> o(\<lambda>x. x powr e')"
|
|
3350 |
using e' by (subst powr_smallo_iff) (auto simp: filterlim_ident)
|
|
3351 |
finally have "(\<lambda>x. x) \<in> o(\<lambda>x. x powr e')" .
|
|
3352 |
}
|
|
3353 |
ultimately show ?thesis using assms
|
|
3354 |
by (auto simp: expands_to.simps basis_wf_Cons elim: eventually_mono
|
|
3355 |
intro!: is_expansion_aux.intros is_expansion_const
|
|
3356 |
landau_o.small.compose[of _ at_top _ b])
|
|
3357 |
qed
|
|
3358 |
|
|
3359 |
lemma expands_to_lift':
|
|
3360 |
assumes "basis_wf (b # basis)" "(f expands_to MS xs g) basis"
|
|
3361 |
shows "(f expands_to (MS (MSLCons (MS xs g, 0) MSLNil) g)) (b # basis)"
|
|
3362 |
proof -
|
|
3363 |
have "is_lifting lift_ms basis (b # basis)" by (rule is_lifting_lift) fact+
|
|
3364 |
from expands_to_lift[OF this assms(2)] show ?thesis by (simp add: lift_ms_def)
|
|
3365 |
qed
|
|
3366 |
|
|
3367 |
lemma expands_to_lift'':
|
|
3368 |
assumes "basis_wf (b # basis)" "(f expands_to F) basis"
|
|
3369 |
shows "(f expands_to (MS (MSLCons (F, 0) MSLNil) (eval F))) (b # basis)"
|
|
3370 |
proof -
|
|
3371 |
have "is_lifting lift_ms basis (b # basis)" by (rule is_lifting_lift) fact+
|
|
3372 |
from expands_to_lift[OF this assms(2)] show ?thesis by (simp add: lift_ms_def)
|
|
3373 |
qed
|
|
3374 |
|
|
3375 |
lemma expands_to_drop_zero:
|
|
3376 |
assumes "eventually (\<lambda>x. eval C x = 0) at_top"
|
|
3377 |
assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) basis"
|
|
3378 |
shows "(f expands_to (MS xs g)) basis"
|
|
3379 |
using assms drop_zero_ms[of C e xs g basis] by (simp add: expands_to.simps)
|
|
3380 |
|
|
3381 |
lemma expands_to_hd'':
|
|
3382 |
assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) (b # basis)"
|
|
3383 |
shows "(eval C expands_to C) basis"
|
|
3384 |
using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons)
|
|
3385 |
|
|
3386 |
lemma expands_to_hd:
|
|
3387 |
assumes "(f expands_to (MS (MSLCons (MS ys h, e) xs) g)) (b # basis)"
|
|
3388 |
shows "(h expands_to MS ys h) basis"
|
|
3389 |
using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons)
|
|
3390 |
|
|
3391 |
lemma expands_to_hd':
|
|
3392 |
assumes "(f expands_to (MS (MSLCons (c :: real, e) xs) g)) (b # basis)"
|
|
3393 |
shows "((\<lambda>_. c) expands_to c) basis"
|
|
3394 |
using assms by (auto simp: expands_to.simps elim: is_expansion_aux_MSLCons)
|
|
3395 |
|
|
3396 |
lemma expands_to_trim_cong:
|
|
3397 |
assumes "(f expands_to (MS (MSLCons (C, e) xs) g)) (b # basis)"
|
|
3398 |
assumes "(eval C expands_to C') basis"
|
|
3399 |
shows "(f expands_to (MS (MSLCons (C', e) xs) g)) (b # basis)"
|
|
3400 |
proof -
|
|
3401 |
from assms(1) have "is_expansion_aux xs (\<lambda>x. g x - eval C x * b x powr e) (b # basis)"
|
|
3402 |
by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons)
|
|
3403 |
hence "is_expansion_aux xs (\<lambda>x. g x - eval C' x * b x powr e) (b # basis)"
|
|
3404 |
by (rule is_expansion_aux_cong)
|
|
3405 |
(insert assms(2), auto simp: expands_to.simps elim: eventually_mono)
|
|
3406 |
with assms show ?thesis
|
|
3407 |
by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons intro!: is_expansion_aux.intros)
|
|
3408 |
qed
|
|
3409 |
|
|
3410 |
lemma is_expansion_aux_expands_to:
|
|
3411 |
assumes "(f expands_to MS xs g) basis"
|
|
3412 |
shows "is_expansion_aux xs f basis"
|
|
3413 |
proof -
|
|
3414 |
from assms have "is_expansion_aux xs g basis" "eventually (\<lambda>x. g x = f x) at_top"
|
|
3415 |
by (simp_all add: expands_to.simps)
|
|
3416 |
thus ?thesis by (rule is_expansion_aux_cong)
|
|
3417 |
qed
|
|
3418 |
|
|
3419 |
lemma ln_series_stream_aux_code:
|
|
3420 |
"ln_series_stream_aux True c = MSSCons (- inverse c) (ln_series_stream_aux False (c + 1))"
|
|
3421 |
"ln_series_stream_aux False c = MSSCons (inverse c) (ln_series_stream_aux True (c + 1))"
|
|
3422 |
by (subst ln_series_stream_aux.code, simp)+
|
|
3423 |
|
|
3424 |
definition powser_ms_aux' where
|
|
3425 |
"powser_ms_aux' cs xs = powser_ms_aux (msllist_of_msstream cs) xs"
|
|
3426 |
|
|
3427 |
lemma ln_ms_aux:
|
|
3428 |
fixes C L :: "'a :: multiseries"
|
|
3429 |
assumes trimmed: "trimmed_pos C" and basis: "basis_wf (b # basis)"
|
|
3430 |
assumes F: "is_expansion_aux (MSLCons (C, e) xs) f (b # basis)"
|
|
3431 |
assumes L: "((\<lambda>x. ln (eval C x) + e * ln (b x)) expands_to L) basis"
|
|
3432 |
shows "is_expansion_aux
|
|
3433 |
(MSLCons (L, 0) (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs)
|
|
3434 |
(powser_ms_aux' (ln_series_stream_aux False 1)
|
|
3435 |
(scale_shift_ms_aux (inverse C, - e) xs))))
|
|
3436 |
(\<lambda>x. ln (f x)) (b # basis)"
|
|
3437 |
(is "is_expansion_aux ?G _ _")
|
|
3438 |
proof -
|
|
3439 |
from F have F':
|
|
3440 |
"is_expansion_aux xs (\<lambda>x. f x - eval C x * b x powr e) (b # basis)"
|
|
3441 |
"is_expansion C basis" "\<forall>e'>e. f \<in> o(\<lambda>x. b x powr e')" "ms_exp_gt e (ms_aux_hd_exp xs)"
|
|
3442 |
by (auto elim!: is_expansion_aux_MSLCons)
|
|
3443 |
from basis have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
3444 |
hence b_pos: "eventually (\<lambda>x. b x > 1) at_top" by (simp add: filterlim_at_top_dense)
|
|
3445 |
from eval_pos_if_dominant_term_pos[of basis C] trimmed F' basis
|
|
3446 |
have C_pos: "eventually (\<lambda>x. eval C x > 0) at_top"
|
|
3447 |
by (auto simp: basis_wf_Cons trimmed_pos_def)
|
|
3448 |
from eval_pos_if_dominant_term_pos'[OF basis _ F] trimmed
|
|
3449 |
have f_pos: "eventually (\<lambda>x. f x > 0) at_top"
|
|
3450 |
by (simp add: trimmed_pos_def trimmed_ms_aux_def dominant_term_ms_aux_def case_prod_unfold)
|
|
3451 |
|
|
3452 |
from F' have "ms_exp_gt 0 (map_option ((+) (-e)) (ms_aux_hd_exp xs))"
|
|
3453 |
by (cases "ms_aux_hd_exp xs") simp_all
|
|
3454 |
hence "is_expansion_aux (powser_ms_aux' ln_series_stream (scale_shift_ms_aux (inverse C, -e) xs))
|
|
3455 |
((\<lambda>x. ln (1 + x)) \<circ> (\<lambda>x. eval (inverse C) x * b x powr -e *
|
|
3456 |
(f x - eval C x * b x powr e))) (b # basis)" (is "is_expansion_aux _ ?g _")
|
|
3457 |
unfolding powser_ms_aux'_def using powser_ms_aux' basis F' trimmed
|
|
3458 |
by (intro powser_ms_aux' convergent_powser'_ln scale_shift_ms_aux is_expansion_inverse F')
|
|
3459 |
(simp_all add: trimmed_pos_def hd_exp_basis basis_wf_Cons)
|
|
3460 |
moreover have "eventually (\<lambda>x. ?g x = ln (f x) - eval L x * b x powr 0) at_top"
|
|
3461 |
using b_pos C_pos f_pos eval_expands_to[OF L]
|
|
3462 |
by eventually_elim
|
|
3463 |
(simp add: powr_minus algebra_simps ln_mult ln_inverse expands_to.simps ln_powr)
|
|
3464 |
ultimately
|
|
3465 |
have "is_expansion_aux (powser_ms_aux' ln_series_stream (scale_shift_ms_aux (inverse C, -e) xs))
|
|
3466 |
(\<lambda>x. ln (f x) - eval L x * b x powr 0) (b # basis)"
|
|
3467 |
by (rule is_expansion_aux_cong)
|
|
3468 |
hence *: "is_expansion_aux (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs)
|
|
3469 |
(powser_ms_aux' (ln_series_stream_aux False 1) (scale_shift_ms_aux (inverse C, - e) xs)))
|
|
3470 |
(\<lambda>x. ln (f x) - eval L x * b x powr 0) (b # basis)"
|
|
3471 |
unfolding ln_series_stream_def powser_ms_aux'_def powser_ms_aux_MSLCons msllist_of_msstream_MSSCons
|
|
3472 |
by (rule drop_zero_ms_aux [rotated]) simp_all
|
|
3473 |
moreover from F' have exp: "ms_exp_gt 0 (map_option ((+) (-e)) (ms_aux_hd_exp xs))"
|
|
3474 |
by (cases "ms_aux_hd_exp xs") simp_all
|
|
3475 |
moreover have "(\<lambda>x. ln (f x)) \<in> o(\<lambda>x. b x powr e')" if "e' > 0" for e'
|
|
3476 |
proof -
|
|
3477 |
from is_expansion_aux_imp_smallo[OF *, of e'] that exp
|
|
3478 |
have "(\<lambda>x. ln (f x) - eval L x * b x powr 0) \<in> o(\<lambda>x. b x powr e')"
|
|
3479 |
by (cases "ms_aux_hd_exp xs")
|
|
3480 |
(simp_all add: hd_exp_times hd_exp_powser hd_exp_basis powser_ms_aux'_def)
|
|
3481 |
moreover {
|
|
3482 |
from L F' basis that have "eval L \<in> o(\<lambda>x. b x powr e')"
|
|
3483 |
by (intro is_expansion_imp_smallo[of basis]) (simp_all add: basis_wf_Cons expands_to.simps)
|
|
3484 |
also have "eval L \<in> o(\<lambda>x. b x powr e') \<longleftrightarrow> (\<lambda>x. eval L x * b x powr 0) \<in> o(\<lambda>x. b x powr e')"
|
|
3485 |
using b_pos by (intro landau_o.small.in_cong) (auto elim: eventually_mono)
|
|
3486 |
finally have \<dots> .
|
|
3487 |
} ultimately have "(\<lambda>x. ln (f x) - eval L x * b x powr 0 + eval L x * b x powr 0) \<in>
|
|
3488 |
o(\<lambda>x. b x powr e')" by (rule sum_in_smallo)
|
|
3489 |
thus ?thesis by simp
|
|
3490 |
qed
|
|
3491 |
ultimately show "is_expansion_aux ?G (\<lambda>x. ln (f x)) (b # basis)" using L
|
|
3492 |
by (intro is_expansion_aux.intros)
|
|
3493 |
(auto simp: expands_to.simps hd_exp_times hd_exp_powser hd_exp_basis
|
|
3494 |
powser_ms_aux'_def split: option.splits)
|
|
3495 |
qed
|
|
3496 |
|
|
3497 |
lemma expands_to_ln:
|
|
3498 |
fixes C L :: "'a :: multiseries"
|
|
3499 |
assumes trimmed: "trimmed_pos (MS (MSLCons (C, e) xs) g)" and basis: "basis_wf (b # basis)"
|
|
3500 |
assumes F: "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
3501 |
assumes L: "((\<lambda>x. ln (h x) + e * ln (b x)) expands_to L) basis"
|
|
3502 |
and h: "eval C = h"
|
|
3503 |
shows "((\<lambda>x. ln (f x)) expands_to MS
|
|
3504 |
(MSLCons (L, 0) (times_ms_aux (scale_shift_ms_aux (inverse C, - e) xs)
|
|
3505 |
(powser_ms_aux' (ln_series_stream_aux False 1)
|
|
3506 |
(scale_shift_ms_aux (inverse C, - e) xs)))) (\<lambda>x. ln (f x))) (b # basis)"
|
|
3507 |
using is_expansion_aux_expands_to[OF F] assms by (auto simp: expands_to.simps intro!: ln_ms_aux)
|
|
3508 |
|
|
3509 |
lemma trimmed_lifting:
|
|
3510 |
assumes "is_lifting f basis basis'"
|
|
3511 |
assumes "trimmed F"
|
|
3512 |
shows "trimmed (f F)"
|
|
3513 |
using assms by (simp add: is_lifting_def)
|
|
3514 |
|
|
3515 |
lemma trimmed_pos_lifting:
|
|
3516 |
assumes "is_lifting f basis basis'"
|
|
3517 |
assumes "trimmed_pos F"
|
|
3518 |
shows "trimmed_pos (f F)"
|
|
3519 |
using assms by (simp add: is_lifting_def trimmed_pos_def)
|
|
3520 |
|
|
3521 |
lemma expands_to_ln_aux_0:
|
|
3522 |
assumes e: "e = 0"
|
|
3523 |
assumes L1: "((\<lambda>x. ln (g x)) expands_to L) basis"
|
|
3524 |
shows "((\<lambda>x. ln (g x) + e * ln (b x)) expands_to L) basis"
|
|
3525 |
using assms by simp
|
|
3526 |
|
|
3527 |
lemma expands_to_ln_aux_1:
|
|
3528 |
assumes e: "e = 1"
|
|
3529 |
assumes basis: "basis_wf (b # basis)"
|
|
3530 |
assumes L1: "((\<lambda>x. ln (g x)) expands_to L1) basis"
|
|
3531 |
assumes L2: "((\<lambda>x. ln (b x)) expands_to L2) basis"
|
|
3532 |
shows "((\<lambda>x. ln (g x) + e * ln (b x)) expands_to L1 + L2) basis"
|
|
3533 |
using assms by (intro expands_to_add) (simp_all add: basis_wf_Cons)
|
|
3534 |
|
|
3535 |
lemma expands_to_ln_eventually_1:
|
|
3536 |
assumes "basis_wf basis" "length basis = expansion_level TYPE('a)"
|
|
3537 |
assumes "eventually (\<lambda>x. f x - 1 = 0) at_top"
|
|
3538 |
shows "((\<lambda>x. ln (f x)) expands_to (zero_expansion :: 'a :: multiseries)) basis"
|
|
3539 |
by (rule expands_to_cong [OF expands_to_zero]) (insert assms, auto elim: eventually_mono)
|
|
3540 |
|
|
3541 |
lemma expands_to_ln_aux:
|
|
3542 |
assumes basis: "basis_wf (b # basis)"
|
|
3543 |
assumes L1: "((\<lambda>x. ln (g x)) expands_to L1) basis"
|
|
3544 |
assumes L2: "((\<lambda>x. ln (b x)) expands_to L2) basis"
|
|
3545 |
shows "((\<lambda>x. ln (g x) + e * ln (b x)) expands_to L1 + scale_ms e L2) basis"
|
|
3546 |
proof -
|
|
3547 |
from L1 have "length basis = expansion_level TYPE('a)"
|
|
3548 |
by (auto simp: expands_to.simps is_expansion_length)
|
|
3549 |
with assms show ?thesis unfolding scale_ms_def
|
|
3550 |
by (intro expands_to_add assms expands_to_mult expands_to_const)
|
|
3551 |
(simp_all add: basis_wf_Cons)
|
|
3552 |
qed
|
|
3553 |
|
|
3554 |
lemma expands_to_ln_to_expands_to_ln_eval:
|
|
3555 |
assumes "((\<lambda>x. ln (f x) + F x) expands_to L) basis"
|
|
3556 |
shows "((\<lambda>x. ln (eval (MS xs f) x) + F x) expands_to L) basis"
|
|
3557 |
using assms by simp
|
|
3558 |
|
|
3559 |
lemma expands_to_ln_const:
|
|
3560 |
"((\<lambda>x. ln (eval (C :: real) x)) expands_to ln C) []"
|
|
3561 |
by (simp add: expands_to.simps is_expansion_real.intros)
|
|
3562 |
|
|
3563 |
lemma expands_to_meta_eq_cong:
|
|
3564 |
assumes "(f expands_to F) basis" "F \<equiv> G"
|
|
3565 |
shows "(f expands_to G) basis"
|
|
3566 |
using assms by simp
|
|
3567 |
|
|
3568 |
lemma expands_to_meta_eq_cong':
|
|
3569 |
assumes "(g expands_to F) basis" "f \<equiv> g"
|
|
3570 |
shows "(f expands_to F) basis"
|
|
3571 |
using assms by simp
|
|
3572 |
|
|
3573 |
|
|
3574 |
lemma uminus_ms_aux_MSLCons_eval:
|
|
3575 |
"uminus_ms_aux (MSLCons (c, e) xs) = MSLCons (-c, e) (uminus_ms_aux xs)"
|
|
3576 |
by (simp add: uminus_ms_aux_MSLCons)
|
|
3577 |
|
|
3578 |
lemma scale_shift_ms_aux_MSLCons_eval:
|
|
3579 |
"scale_shift_ms_aux (c, e) (MSLCons (c', e') xs) =
|
|
3580 |
MSLCons (c * c', e + e') (scale_shift_ms_aux (c, e) xs)"
|
|
3581 |
by (simp add: scale_shift_ms_aux_MSLCons)
|
|
3582 |
|
|
3583 |
lemma times_ms_aux_MSLCons_eval: "times_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) =
|
|
3584 |
MSLCons (c1 * c2, e1 + e2) (plus_ms_aux (scale_shift_ms_aux (c1, e1) ys)
|
|
3585 |
(times_ms_aux xs (MSLCons (c2, e2) ys)))"
|
|
3586 |
by (simp add: times_ms_aux_MSLCons)
|
|
3587 |
|
|
3588 |
lemma plus_ms_aux_MSLCons_eval:
|
|
3589 |
"plus_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) =
|
|
3590 |
CMP_BRANCH (COMPARE e1 e2)
|
|
3591 |
(MSLCons (c2, e2) (plus_ms_aux (MSLCons (c1, e1) xs) ys))
|
|
3592 |
(MSLCons (c1 + c2, e1) (plus_ms_aux xs ys))
|
|
3593 |
(MSLCons (c1, e1) (plus_ms_aux xs (MSLCons (c2, e2) ys)))"
|
|
3594 |
by (simp add: CMP_BRANCH_def COMPARE_def plus_ms_aux_MSLCons)
|
|
3595 |
|
|
3596 |
lemma inverse_ms_aux_MSLCons: "inverse_ms_aux (MSLCons (C, e) xs) =
|
|
3597 |
scale_shift_ms_aux (inverse C, - e)
|
|
3598 |
(powser_ms_aux' (mssalternate 1 (- 1))
|
|
3599 |
(scale_shift_ms_aux (inverse C, - e)
|
|
3600 |
xs))"
|
|
3601 |
by (simp add: Let_def inverse_ms_aux.simps powser_ms_aux'_def)
|
|
3602 |
|
|
3603 |
lemma powr_ms_aux_MSLCons: "powr_ms_aux abort (MSLCons (C, e) xs) p =
|
|
3604 |
scale_shift_ms_aux (powr_expansion abort C p, e * p)
|
|
3605 |
(powser_ms_aux (gbinomial_series abort p)
|
|
3606 |
(scale_shift_ms_aux (inverse C, - e) xs))"
|
|
3607 |
by simp
|
|
3608 |
|
|
3609 |
lemma power_ms_aux_MSLCons: "power_ms_aux abort (MSLCons (C, e) xs) n =
|
|
3610 |
scale_shift_ms_aux (power_expansion abort C n, e * real n)
|
|
3611 |
(powser_ms_aux (gbinomial_series abort (real n))
|
|
3612 |
(scale_shift_ms_aux (inverse C, - e) xs))"
|
|
3613 |
by simp
|
|
3614 |
|
|
3615 |
lemma minus_ms_aux_MSLCons_eval:
|
|
3616 |
"minus_ms_aux (MSLCons (c1, e1) xs) (MSLCons (c2, e2) ys) =
|
|
3617 |
CMP_BRANCH (COMPARE e1 e2)
|
|
3618 |
(MSLCons (-c2, e2) (minus_ms_aux (MSLCons (c1, e1) xs) ys))
|
|
3619 |
(MSLCons (c1 - c2, e1) (minus_ms_aux xs ys))
|
|
3620 |
(MSLCons (c1, e1) (minus_ms_aux xs (MSLCons (c2, e2) ys)))"
|
|
3621 |
by (simp add: minus_ms_aux_def plus_ms_aux_MSLCons_eval uminus_ms_aux_MSLCons minus_eq_plus_uminus)
|
|
3622 |
|
|
3623 |
lemma minus_ms_altdef: "MS xs f - MS ys g = MS (minus_ms_aux xs ys) (\<lambda>x. f x - g x)"
|
|
3624 |
by (simp add: minus_ms_def minus_ms_aux_def)
|
|
3625 |
|
|
3626 |
lemma const_expansion_ms_eval: "const_expansion c = MS (MSLCons (const_expansion c, 0) MSLNil) (\<lambda>_. c)"
|
|
3627 |
by (simp add: const_expansion_ms_def const_ms_aux_def)
|
|
3628 |
|
|
3629 |
lemma powser_ms_aux'_MSSCons:
|
|
3630 |
"powser_ms_aux' (MSSCons c cs) xs =
|
|
3631 |
MSLCons (const_expansion c, 0) (times_ms_aux xs (powser_ms_aux' cs xs))"
|
|
3632 |
by (simp add: powser_ms_aux'_def powser_ms_aux_MSLCons)
|
|
3633 |
|
|
3634 |
lemma sin_ms_aux'_altdef:
|
|
3635 |
"sin_ms_aux' xs = times_ms_aux xs (powser_ms_aux' sin_series_stream (times_ms_aux xs xs))"
|
|
3636 |
by (simp add: sin_ms_aux'_def powser_ms_aux'_def)
|
|
3637 |
|
|
3638 |
lemma cos_ms_aux'_altdef:
|
|
3639 |
"cos_ms_aux' xs = powser_ms_aux' cos_series_stream (times_ms_aux xs xs)"
|
|
3640 |
by (simp add: cos_ms_aux'_def powser_ms_aux'_def)
|
|
3641 |
|
|
3642 |
lemma sin_series_stream_aux_code:
|
|
3643 |
"sin_series_stream_aux True acc m =
|
|
3644 |
MSSCons (-inverse acc) (sin_series_stream_aux False (acc * (m + 1) * (m + 2)) (m + 2))"
|
|
3645 |
"sin_series_stream_aux False acc m =
|
|
3646 |
MSSCons (inverse acc) (sin_series_stream_aux True (acc * (m + 1) * (m + 2)) (m + 2))"
|
|
3647 |
by (subst sin_series_stream_aux.code; simp)+
|
|
3648 |
|
|
3649 |
lemma arctan_series_stream_aux_code:
|
|
3650 |
"arctan_series_stream_aux True n = MSSCons (-inverse n) (arctan_series_stream_aux False (n + 2))"
|
|
3651 |
"arctan_series_stream_aux False n = MSSCons (inverse n) (arctan_series_stream_aux True (n + 2))"
|
|
3652 |
by (subst arctan_series_stream_aux.code; simp)+
|
|
3653 |
|
|
3654 |
|
|
3655 |
subsubsection \<open>Composition with an asymptotic power series\<close>
|
|
3656 |
|
|
3657 |
context
|
|
3658 |
fixes h :: "real \<Rightarrow> real" and F :: "real filter"
|
|
3659 |
begin
|
|
3660 |
|
|
3661 |
coinductive asymp_powser ::
|
|
3662 |
"(real \<Rightarrow> real) \<Rightarrow> real msstream \<Rightarrow> bool" where
|
|
3663 |
"asymp_powser f cs \<Longrightarrow> f' \<in> O[F](\<lambda>_. 1) \<Longrightarrow>
|
|
3664 |
eventually (\<lambda>x. f' x = c + f x * h x) F \<Longrightarrow> asymp_powser f' (MSSCons c cs)"
|
|
3665 |
|
|
3666 |
lemma asymp_powser_coinduct [case_names asymp_powser]:
|
|
3667 |
"P x1 x2 \<Longrightarrow> (\<And>x1 x2. P x1 x2 \<Longrightarrow> \<exists>f cs c.
|
|
3668 |
x2 = MSSCons c cs \<and> asymp_powser f cs \<and>
|
|
3669 |
x1 \<in> O[F](\<lambda>_. 1) \<and> (\<forall>\<^sub>F x in F. x1 x = c + f x * h x)) \<Longrightarrow> asymp_powser x1 x2"
|
|
3670 |
using asymp_powser.coinduct[of P x1 x2] by blast
|
|
3671 |
|
|
3672 |
lemma asymp_powser_coinduct' [case_names asymp_powser]:
|
|
3673 |
"P x1 x2 \<Longrightarrow> (\<And>x1 x2. P x1 x2 \<Longrightarrow> \<exists>f cs c.
|
|
3674 |
x2 = MSSCons c cs \<and> P f cs \<and>
|
|
3675 |
x1 \<in> O[F](\<lambda>_. 1) \<and> (\<forall>\<^sub>F x in F. x1 x = c + f x * h x)) \<Longrightarrow> asymp_powser x1 x2"
|
|
3676 |
using asymp_powser.coinduct[of P x1 x2] by blast
|
|
3677 |
|
|
3678 |
lemma asymp_powser_transfer:
|
|
3679 |
assumes "asymp_powser f cs" "eventually (\<lambda>x. f x = f' x) F"
|
|
3680 |
shows "asymp_powser f' cs"
|
|
3681 |
using assms(1)
|
|
3682 |
proof (cases rule: asymp_powser.cases)
|
|
3683 |
case (1 f cs' c)
|
|
3684 |
have "asymp_powser f' (MSSCons c cs')"
|
|
3685 |
by (rule asymp_powser.intros)
|
|
3686 |
(insert 1 assms(2), auto elim: eventually_elim2 dest: landau_o.big.in_cong)
|
|
3687 |
thus ?thesis by (simp add: 1)
|
|
3688 |
qed
|
|
3689 |
|
|
3690 |
lemma sum_bigo_imp_asymp_powser:
|
|
3691 |
assumes filterlim_h: "filterlim h (at 0) F"
|
|
3692 |
assumes "(\<And>n. (\<lambda>x. f x - (\<Sum>k<n. mssnth cs k * h x ^ k)) \<in> O[F](\<lambda>x. h x ^ n))"
|
|
3693 |
shows "asymp_powser f cs"
|
|
3694 |
using assms(2)
|
|
3695 |
proof (coinduction arbitrary: f cs rule: asymp_powser_coinduct')
|
|
3696 |
case (asymp_powser f cs)
|
|
3697 |
from filterlim_h have h_nz:"eventually (\<lambda>x. h x \<noteq> 0) F"
|
|
3698 |
by (auto simp: filterlim_at)
|
|
3699 |
show ?case
|
|
3700 |
proof (cases cs)
|
|
3701 |
case (MSSCons c cs')
|
|
3702 |
define f' where "f' = (\<lambda>x. (f x - c) / h x)"
|
|
3703 |
have "(\<lambda>x. f' x - (\<Sum>k<n. mssnth cs' k * h x ^ k)) \<in> O[F](\<lambda>x. h x ^ n)" for n
|
|
3704 |
proof -
|
|
3705 |
have "(\<lambda>x. h x * (f' x - (\<Sum>i<n. mssnth cs' i * h x ^ i))) \<in>
|
|
3706 |
\<Theta>[F](\<lambda>x. f x - c - h x * (\<Sum>i<n. mssnth cs' i * h x ^ i))"
|
|
3707 |
using h_nz by (intro bigthetaI_cong) (auto elim!: eventually_mono simp: f'_def field_simps)
|
|
3708 |
also from spec[OF asymp_powser, of "Suc n"]
|
|
3709 |
have "(\<lambda>x. f x - c - h x * (\<Sum>i<n. mssnth cs' i * h x ^ i)) \<in> O[F](\<lambda>x. h x * h x ^ n)"
|
|
3710 |
unfolding sum_lessThan_Suc_shift
|
|
3711 |
by (simp add: MSSCons algebra_simps sum_distrib_left sum_distrib_right)
|
|
3712 |
finally show ?thesis
|
|
3713 |
by (subst (asm) landau_o.big.mult_cancel_left) (insert h_nz, auto)
|
|
3714 |
qed
|
|
3715 |
moreover from h_nz have "\<forall>\<^sub>F x in F. f x = c + f' x * h x"
|
|
3716 |
by eventually_elim (simp add: f'_def)
|
|
3717 |
ultimately show ?thesis using spec[OF asymp_powser, of 0]
|
|
3718 |
by (auto simp: MSSCons intro!: exI[of _ f'])
|
|
3719 |
qed
|
|
3720 |
qed
|
|
3721 |
|
|
3722 |
end
|
|
3723 |
|
|
3724 |
|
|
3725 |
lemma asymp_powser_o:
|
|
3726 |
assumes "asymp_powser h F f cs" assumes "filterlim g F G"
|
|
3727 |
shows "asymp_powser (h \<circ> g) G (f \<circ> g) cs"
|
|
3728 |
using assms(1)
|
|
3729 |
proof (coinduction arbitrary: f cs rule: asymp_powser_coinduct')
|
|
3730 |
case (asymp_powser f cs)
|
|
3731 |
thus ?case
|
|
3732 |
proof (cases rule: asymp_powser.cases)
|
|
3733 |
case (1 f' cs' c)
|
|
3734 |
from assms(2) have "filtermap g G \<le> F" by (simp add: filterlim_def)
|
|
3735 |
moreover have "eventually (\<lambda>x. f x = c + f' x * h x) F" by fact
|
|
3736 |
ultimately have "eventually (\<lambda>x. f x = c + f' x * h x) (filtermap g G)" by (rule filter_leD)
|
|
3737 |
hence "eventually (\<lambda>x. f (g x) = c + f' (g x) * h (g x)) G" by (simp add: eventually_filtermap)
|
|
3738 |
moreover from 1 have "f \<circ> g \<in> O[G](\<lambda>_. 1)" unfolding o_def
|
|
3739 |
by (intro landau_o.big.compose[OF _ assms(2)]) auto
|
|
3740 |
ultimately show ?thesis using 1 by force
|
|
3741 |
qed
|
|
3742 |
qed
|
|
3743 |
|
|
3744 |
lemma asymp_powser_compose:
|
|
3745 |
assumes "asymp_powser h F f cs" assumes "filterlim g F G"
|
|
3746 |
shows "asymp_powser (\<lambda>x. h (g x)) G (\<lambda>x. f (g x)) cs"
|
|
3747 |
using asymp_powser_o[OF assms] by (simp add: o_def)
|
|
3748 |
|
|
3749 |
lemma hd_exp_powser': "ms_aux_hd_exp (powser_ms_aux' cs f) = Some 0"
|
|
3750 |
by (simp add: powser_ms_aux'_def hd_exp_powser)
|
|
3751 |
|
|
3752 |
lemma powser_ms_aux_asymp_powser:
|
|
3753 |
assumes "asymp_powser h at_top f cs" and basis: "basis_wf bs"
|
|
3754 |
assumes "is_expansion_aux xs h bs" "ms_exp_gt 0 (ms_aux_hd_exp xs)"
|
|
3755 |
shows "is_expansion_aux (powser_ms_aux' cs xs) f bs"
|
|
3756 |
using assms(1)
|
|
3757 |
proof (coinduction arbitrary: cs f rule: is_expansion_aux_coinduct_upto)
|
|
3758 |
show "basis_wf bs" by fact
|
|
3759 |
next
|
|
3760 |
case (B cs f)
|
|
3761 |
thus ?case
|
|
3762 |
proof (cases rule: asymp_powser.cases)
|
|
3763 |
case (1 f' cs' c)
|
|
3764 |
from assms(3) obtain b bs' where [simp]: "bs = b # bs'"
|
|
3765 |
by (cases bs) (auto dest: is_expansion_aux_basis_nonempty)
|
|
3766 |
from basis have b_at_top: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
3767 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (auto simp: filterlim_at_top_dense)
|
|
3768 |
|
|
3769 |
have A: "f \<in> o(\<lambda>x. b x powr e)" if "e > 0" for e :: real
|
|
3770 |
proof -
|
|
3771 |
have "f \<in> O(\<lambda>_. 1)" by fact
|
|
3772 |
also have "(\<lambda>_::real. 1) \<in> o(\<lambda>x. b x powr e)"
|
|
3773 |
by (rule landau_o.small.compose[OF _ b_at_top]) (insert that, simp_all add: one_smallo_powr)
|
|
3774 |
finally show ?thesis .
|
|
3775 |
qed
|
|
3776 |
have "ms_closure (\<lambda>xsa'' fa'' basis''.
|
|
3777 |
\<exists>cs. xsa'' = powser_ms_aux' cs xs \<and> basis'' = b # bs' \<and> asymp_powser h at_top fa'' cs)
|
|
3778 |
(times_ms_aux xs (powser_ms_aux' cs' xs)) (\<lambda>x. h x * f' x) bs"
|
|
3779 |
(is "ms_closure ?Q ?ys _ _")
|
|
3780 |
by (rule ms_closure_mult[OF ms_closure_embed' ms_closure_embed])
|
|
3781 |
(insert assms 1, auto intro!: exI[of _ cs'])
|
|
3782 |
moreover have "eventually (\<lambda>x. h x * f' x = f x - c * b x powr 0) at_top"
|
|
3783 |
using b_pos and \<open>\<forall>\<^sub>F x in at_top. f x = c + f' x * h x\<close>
|
|
3784 |
by eventually_elim simp
|
|
3785 |
ultimately have "ms_closure ?Q ?ys (\<lambda>x. f x - c * b x powr 0) bs"
|
|
3786 |
by (rule ms_closure_cong)
|
|
3787 |
with 1 assms A show ?thesis
|
|
3788 |
using is_expansion_aux_expansion_level[OF assms(3)]
|
|
3789 |
by (auto simp: powser_ms_aux'_MSSCons basis_wf_Cons hd_exp_times hd_exp_powser'
|
|
3790 |
intro!: is_expansion_const ms_closure_add ms_closure_mult split: option.splits)
|
|
3791 |
qed
|
|
3792 |
qed
|
|
3793 |
|
|
3794 |
|
|
3795 |
subsubsection \<open>Arc tangent\<close>
|
|
3796 |
|
|
3797 |
definition arctan_ms_aux where
|
|
3798 |
"arctan_ms_aux xs = times_ms_aux xs (powser_ms_aux' arctan_series_stream (times_ms_aux xs xs))"
|
|
3799 |
|
|
3800 |
lemma arctan_ms_aux_0:
|
|
3801 |
assumes "ms_exp_gt 0 (ms_aux_hd_exp xs)" "basis_wf basis" "is_expansion_aux xs f basis"
|
|
3802 |
shows "is_expansion_aux (arctan_ms_aux xs) (\<lambda>x. arctan (f x)) basis"
|
|
3803 |
proof (rule is_expansion_aux_cong)
|
|
3804 |
let ?g = "powser (msllist_of_msstream arctan_series_stream)"
|
|
3805 |
show "is_expansion_aux (arctan_ms_aux xs)
|
|
3806 |
(\<lambda>x. f x * (?g \<circ> (\<lambda>x. f x * f x)) x) basis"
|
|
3807 |
unfolding arctan_ms_aux_def powser_ms_aux'_def using assms
|
|
3808 |
by (intro times_ms_aux powser_ms_aux powser_ms_aux convergent_powser_arctan)
|
|
3809 |
(auto simp: hd_exp_times add_neg_neg split: option.splits)
|
|
3810 |
|
|
3811 |
have "f \<in> o(\<lambda>x. hd basis x powr 0)" using assms
|
|
3812 |
by (intro is_expansion_aux_imp_smallo[OF assms(3)]) auto
|
|
3813 |
also have "(\<lambda>x. hd basis x powr 0) \<in> O(\<lambda>_. 1)"
|
|
3814 |
by (intro bigoI[of _ 1]) auto
|
|
3815 |
finally have "filterlim f (nhds 0) at_top"
|
|
3816 |
by (auto dest!: smalloD_tendsto)
|
|
3817 |
from order_tendstoD(2)[OF tendsto_norm [OF this], of 1]
|
|
3818 |
have "eventually (\<lambda>x. norm (f x) < 1) at_top" by simp
|
|
3819 |
thus "eventually (\<lambda>x. f x * (?g \<circ> (\<lambda>x. f x * f x)) x = arctan (f x)) at_top"
|
|
3820 |
by eventually_elim (simp add: arctan_conv_pre_arctan power2_eq_square)
|
|
3821 |
qed
|
|
3822 |
|
|
3823 |
lemma arctan_ms_aux_inf:
|
|
3824 |
assumes "\<exists>e>0. ms_aux_hd_exp xs = Some e" "fst (dominant_term_ms_aux xs) > 0" "trimmed_ms_aux xs"
|
|
3825 |
"basis_wf basis" "is_expansion_aux xs f basis"
|
|
3826 |
shows "is_expansion_aux (minus_ms_aux (const_ms_aux (pi / 2))
|
|
3827 |
(arctan_ms_aux (inverse_ms_aux xs))) (\<lambda>x. arctan (f x)) basis"
|
|
3828 |
(is "is_expansion_aux ?xs' _ _")
|
|
3829 |
proof (rule is_expansion_aux_cong)
|
|
3830 |
from \<open>trimmed_ms_aux xs\<close> have [simp]: "xs \<noteq> MSLNil" by auto
|
|
3831 |
thus "is_expansion_aux ?xs' (\<lambda>x. pi / 2 - arctan (inverse (f x))) basis" using assms
|
|
3832 |
by (intro minus_ms_aux arctan_ms_aux_0 inverse_ms_aux const_ms_aux)
|
|
3833 |
(auto simp: hd_exp_inverse is_expansion_aux_expansion_level)
|
|
3834 |
next
|
|
3835 |
from assms(2-5) have "eventually (\<lambda>x. f x > 0) at_top"
|
|
3836 |
by (intro eval_pos_if_dominant_term_pos') simp_all
|
|
3837 |
thus "eventually (\<lambda>x. ((\<lambda>x. pi/2 - arctan (inverse (f x)))) x = arctan (f x)) at_top"
|
|
3838 |
unfolding o_def by eventually_elim (subst arctan_inverse_pos, simp_all)
|
|
3839 |
qed
|
|
3840 |
|
|
3841 |
lemma arctan_ms_aux_neg_inf:
|
|
3842 |
assumes "\<exists>e>0. ms_aux_hd_exp xs = Some e" "fst (dominant_term_ms_aux xs) < 0" "trimmed_ms_aux xs"
|
|
3843 |
"basis_wf basis" "is_expansion_aux xs f basis"
|
|
3844 |
shows "is_expansion_aux (minus_ms_aux (const_ms_aux (-pi / 2))
|
|
3845 |
(arctan_ms_aux (inverse_ms_aux xs))) (\<lambda>x. arctan (f x)) basis"
|
|
3846 |
(is "is_expansion_aux ?xs' _ _")
|
|
3847 |
proof (rule is_expansion_aux_cong)
|
|
3848 |
from \<open>trimmed_ms_aux xs\<close> have [simp]: "xs \<noteq> MSLNil" by auto
|
|
3849 |
thus "is_expansion_aux ?xs' (\<lambda>x. -pi / 2 - arctan (inverse (f x))) basis" using assms
|
|
3850 |
by (intro minus_ms_aux arctan_ms_aux_0 inverse_ms_aux const_ms_aux)
|
|
3851 |
(auto simp: hd_exp_inverse is_expansion_aux_expansion_level)
|
|
3852 |
next
|
|
3853 |
from assms(2-5) have "eventually (\<lambda>x. f x < 0) at_top"
|
|
3854 |
by (intro eval_neg_if_dominant_term_neg') simp_all
|
|
3855 |
thus "eventually (\<lambda>x. ((\<lambda>x. -pi/2 - arctan (inverse (f x)))) x = arctan (f x)) at_top"
|
|
3856 |
unfolding o_def by eventually_elim (subst arctan_inverse_neg, simp_all)
|
|
3857 |
qed
|
|
3858 |
|
|
3859 |
lemma expands_to_arctan_zero:
|
|
3860 |
assumes "((\<lambda>_::real. 0::real) expands_to C) bs" "eventually (\<lambda>x. f x = 0) at_top"
|
|
3861 |
shows "((\<lambda>x::real. arctan (f x)) expands_to C) bs"
|
|
3862 |
using assms by (auto simp: expands_to.simps elim: eventually_elim2)
|
|
3863 |
|
|
3864 |
lemma expands_to_arctan_ms_neg_exp:
|
|
3865 |
assumes "e < 0" "basis_wf basis" "(f expands_to MS (MSLCons (C, e) xs) g) basis"
|
|
3866 |
shows "((\<lambda>x. arctan (f x)) expands_to
|
|
3867 |
MS (arctan_ms_aux (MSLCons (C, e) xs)) (\<lambda>x. arctan (g x))) basis"
|
|
3868 |
using assms by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
|
|
3869 |
intro!: arctan_ms_aux_0 elim: eventually_mono)
|
|
3870 |
|
|
3871 |
lemma expands_to_arctan_ms_pos_exp_pos:
|
|
3872 |
assumes "e > 0" "trimmed_pos (MS (MSLCons (C, e) xs) g)" "basis_wf basis"
|
|
3873 |
"(f expands_to MS (MSLCons (C, e) xs) g) basis"
|
|
3874 |
shows "((\<lambda>x. arctan (f x)) expands_to MS (minus_ms_aux (const_ms_aux (pi / 2))
|
|
3875 |
(arctan_ms_aux (inverse_ms_aux (MSLCons (C, e) xs))))
|
|
3876 |
(\<lambda>x. arctan (g x))) basis"
|
|
3877 |
using arctan_ms_aux_inf[of "MSLCons (C, e) xs" basis g] assms
|
|
3878 |
by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
|
|
3879 |
dominant_term_ms_aux_MSLCons trimmed_pos_def elim: eventually_mono)
|
|
3880 |
|
|
3881 |
lemma expands_to_arctan_ms_pos_exp_neg:
|
|
3882 |
assumes "e > 0" "trimmed_neg (MS (MSLCons (C, e) xs) g)" "basis_wf basis"
|
|
3883 |
"(f expands_to MS (MSLCons (C, e) xs) g) basis"
|
|
3884 |
shows "((\<lambda>x. arctan (f x)) expands_to MS (minus_ms_aux (const_ms_aux (-pi/2))
|
|
3885 |
(arctan_ms_aux (inverse_ms_aux (MSLCons (C, e) xs))))
|
|
3886 |
(\<lambda>x. arctan (g x))) basis"
|
|
3887 |
using arctan_ms_aux_neg_inf[of "MSLCons (C, e) xs" basis g] assms
|
|
3888 |
by (auto simp: expands_to.simps snd_dominant_term_ms_aux_MSLCons o_def
|
|
3889 |
dominant_term_ms_aux_MSLCons trimmed_neg_def elim: eventually_mono)
|
|
3890 |
|
|
3891 |
|
|
3892 |
subsubsection \<open>Exponential function\<close>
|
|
3893 |
|
|
3894 |
(* TODO: Move *)
|
|
3895 |
lemma ln_smallo_powr:
|
|
3896 |
assumes "(e::real) > 0"
|
|
3897 |
shows "(\<lambda>x. ln x) \<in> o(\<lambda>x. x powr e)"
|
|
3898 |
proof -
|
|
3899 |
have *: "ln \<in> o(\<lambda>x::real. x)"
|
|
3900 |
using ln_x_over_x_tendsto_0 by (intro smalloI_tendsto) auto
|
|
3901 |
have "(\<lambda>x. ln x) \<in> \<Theta>(\<lambda>x::real. ln x powr 1)"
|
|
3902 |
by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 1]]) auto
|
|
3903 |
also have "(\<lambda>x::real. ln x powr 1) \<in> o(\<lambda>x. x powr e)"
|
|
3904 |
by (intro ln_smallo_imp_flat filterlim_ident ln_at_top assms
|
|
3905 |
landau_o.small.compose[of _ at_top _ ln] *)
|
|
3906 |
finally show ?thesis .
|
|
3907 |
qed
|
|
3908 |
|
|
3909 |
lemma basis_wf_insert_exp_pos:
|
|
3910 |
assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
3911 |
"basis_wf (b # basis)" "trimmed (MS (MSLCons (C, e) xs) g)" "e - 0 > 0"
|
|
3912 |
shows "(\<lambda>x. ln (b x)) \<in> o(f)"
|
|
3913 |
proof -
|
|
3914 |
from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
3915 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" by (simp add: filterlim_at_top_dense)
|
|
3916 |
|
|
3917 |
from assms have "is_expansion_aux xs (\<lambda>x. g x - eval C x * b x powr e) (b # basis)"
|
|
3918 |
"ms_exp_gt e (ms_aux_hd_exp xs)"
|
|
3919 |
by (auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps)
|
|
3920 |
from is_expansion_aux_imp_smallo''[OF this] obtain e' where e':
|
|
3921 |
"e' < e" "(\<lambda>x. g x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr e')" unfolding list.sel .
|
|
3922 |
|
|
3923 |
define eps where "eps = e/2"
|
|
3924 |
have "(\<lambda>x. ln (b x)) \<in> o(\<lambda>x. b x powr (e - eps))" unfolding eps_def
|
|
3925 |
by (rule landau_o.small.compose[OF _ b_limit])
|
|
3926 |
(insert e'(1) \<open>e - 0 > 0\<close>, simp add: ln_smallo_powr)
|
|
3927 |
also from assms e'(1) have "eval C \<in> \<omega>(\<lambda>x. b x powr -eps)" unfolding eps_def
|
|
3928 |
by (intro is_expansion_imp_smallomega[of basis])
|
|
3929 |
(auto simp: basis_wf_Cons expands_to.simps trimmed_pos_def trimmed_ms_aux_MSLCons
|
|
3930 |
elim!: is_expansion_aux_MSLCons)
|
|
3931 |
hence "(\<lambda>x. b x powr -eps * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
|
|
3932 |
by (intro landau_o.small_big_mult) (simp_all add: smallomega_iff_smallo)
|
|
3933 |
hence "(\<lambda>x. b x powr (e - eps)) \<in> o(\<lambda>x. eval C x * b x powr e)"
|
|
3934 |
by (simp add: powr_add [symmetric] field_simps)
|
|
3935 |
finally have "(\<lambda>x. ln (b x)) \<in> o(\<lambda>x. eval C x * b x powr e)" .
|
|
3936 |
also {
|
|
3937 |
from e' have "(\<lambda>x. g x - eval C x * b x powr e) \<in> o(\<lambda>x. b x powr (e' - e) * b x powr e)"
|
|
3938 |
by (simp add: powr_add [symmetric])
|
|
3939 |
also from assms e'(1) have "eval C \<in> \<omega>(\<lambda>x. b x powr (e' - e))" unfolding eps_def
|
|
3940 |
by (intro is_expansion_imp_smallomega[of basis])
|
|
3941 |
(auto simp: basis_wf_Cons expands_to.simps trimmed_pos_def trimmed_ms_aux_MSLCons
|
|
3942 |
elim!: is_expansion_aux_MSLCons)
|
|
3943 |
hence "(\<lambda>x. b x powr (e' - e) * b x powr e) \<in> o(\<lambda>x. eval C x * b x powr e)"
|
|
3944 |
by (intro landau_o.small_big_mult is_expansion_imp_smallo)
|
|
3945 |
(simp_all add: smallomega_iff_smallo)
|
|
3946 |
finally have "o(\<lambda>x. eval C x * b x powr e) =
|
|
3947 |
o(\<lambda>x. g x - eval C x * b x powr e + eval C x * b x powr e)"
|
|
3948 |
by (intro landau_o.small.plus_absorb1 [symmetric]) simp_all
|
|
3949 |
}
|
|
3950 |
also have "o(\<lambda>x. g x - eval C x * b x powr e + eval C x * b x powr e) = o(g)" by simp
|
|
3951 |
also from assms have "g \<in> \<Theta>(f)" by (intro bigthetaI_cong) (simp add: expands_to.simps)
|
|
3952 |
finally show "(\<lambda>x. ln (b x)) \<in> o(f)" .
|
|
3953 |
qed
|
|
3954 |
|
|
3955 |
lemma basis_wf_insert_exp_uminus:
|
|
3956 |
"(\<lambda>x. ln (b x)) \<in> o(f) \<Longrightarrow> (\<lambda>x. ln (b x)) \<in> o(\<lambda>x. -f x :: real)"
|
|
3957 |
by simp
|
|
3958 |
|
|
3959 |
lemma basis_wf_insert_exp_uminus':
|
|
3960 |
"f \<in> o(\<lambda>x. ln (b x)) \<Longrightarrow> (\<lambda>x. -f x) \<in> o(\<lambda>x. ln (b x))"
|
|
3961 |
by simp
|
|
3962 |
|
|
3963 |
lemma expands_to_exp_insert_pos:
|
|
3964 |
fixes C :: "'a :: multiseries"
|
|
3965 |
assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
3966 |
"basis_wf (b # basis)" "trimmed_pos (MS (MSLCons (C, e) xs) g)"
|
|
3967 |
"(\<lambda>x. ln (b x)) \<in> o(f)"
|
|
3968 |
shows "filterlim (\<lambda>x. exp (f x)) at_top at_top"
|
|
3969 |
"((\<lambda>x. ln (exp (f x))) expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
3970 |
"((\<lambda>x. exp (f x)) expands_to
|
|
3971 |
MS (MSLCons (const_expansion 1 :: 'a ms, 1) MSLNil) (\<lambda>x. exp (g x)))
|
|
3972 |
((\<lambda>x. exp (f x)) # b # basis)" (is ?thesis3)
|
|
3973 |
proof -
|
|
3974 |
have "ln \<in> \<omega>(\<lambda>x::real. 1)"
|
|
3975 |
by (intro smallomegaI_filterlim_at_top_norm)
|
|
3976 |
(auto intro!: filterlim_compose[OF filterlim_abs_real] ln_at_top)
|
|
3977 |
with assms have "(\<lambda>x. 1) \<in> o(\<lambda>x. ln (b x))"
|
|
3978 |
by (intro landau_o.small.compose[of _ at_top _ b])
|
|
3979 |
(simp_all add: basis_wf_Cons smallomega_iff_smallo)
|
|
3980 |
also note \<open>(\<lambda>x. ln (b x)) \<in> o(f)\<close>
|
|
3981 |
finally have f: "f \<in> \<omega>(\<lambda>_. 1)" by (simp add: smallomega_iff_smallo)
|
|
3982 |
hence f: "filterlim (\<lambda>x. abs (f x)) at_top at_top"
|
|
3983 |
by (auto dest: smallomegaD_filterlim_at_top_norm)
|
|
3984 |
|
|
3985 |
define c where "c = fst (dominant_term C)"
|
|
3986 |
define es where "es = snd (dominant_term C)"
|
|
3987 |
from trimmed_imp_eventually_sgn[OF assms(2), of "MS (MSLCons (C, e) xs) g"] assms
|
|
3988 |
have "\<forall>\<^sub>F x in at_top. abs (f x) = f x"
|
|
3989 |
by (auto simp: dominant_term_ms_aux_MSLCons trimmed_pos_def expands_to.simps sgn_if
|
|
3990 |
split: if_splits elim: eventually_elim2)
|
|
3991 |
from filterlim_cong[OF refl refl this] f
|
|
3992 |
show *: "filterlim (\<lambda>x. exp (f x)) at_top at_top"
|
|
3993 |
by (auto intro: filterlim_compose[OF exp_at_top])
|
|
3994 |
|
|
3995 |
{
|
|
3996 |
fix e' :: real assume e': "e' > 1"
|
|
3997 |
from assms have "(\<lambda>x. exp (g x)) \<in> \<Theta>(\<lambda>x. exp (f x) powr 1)"
|
|
3998 |
by (intro bigthetaI_cong) (auto elim: eventually_mono simp: expands_to.simps)
|
|
3999 |
also from e' * have "(\<lambda>x. exp (f x) powr 1) \<in> o(\<lambda>x. exp (f x) powr e')"
|
|
4000 |
by (subst powr_smallo_iff) (insert *, simp_all)
|
|
4001 |
finally have "(\<lambda>x. exp (g x)) \<in> o(\<lambda>x. exp (f x) powr e')" .
|
|
4002 |
}
|
|
4003 |
with assms show ?thesis3
|
|
4004 |
by (auto intro!: is_expansion_aux.intros is_expansion_const
|
|
4005 |
simp: expands_to.simps is_expansion_aux_expansion_level
|
|
4006 |
dest: is_expansion_aux_expansion_level)
|
|
4007 |
qed (insert assms, simp_all)
|
|
4008 |
|
|
4009 |
lemma expands_to_exp_insert_neg:
|
|
4010 |
fixes C :: "'a :: multiseries"
|
|
4011 |
assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
4012 |
"basis_wf (b # basis)" "trimmed_neg (MS (MSLCons (C, e) xs) g)"
|
|
4013 |
"(\<lambda>x. ln (b x)) \<in> o(f)"
|
|
4014 |
shows "filterlim (\<lambda>x. exp (-f x)) at_top at_top" (is ?thesis1)
|
|
4015 |
"((\<lambda>x. ln (exp (-f x))) expands_to MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\<lambda>x. - g x))
|
|
4016 |
(b # basis)"
|
|
4017 |
"trimmed_pos (MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\<lambda>x. - g x))"
|
|
4018 |
"((\<lambda>x. exp (f x)) expands_to
|
|
4019 |
MS (MSLCons (const_expansion 1 :: 'a ms, -1) MSLNil) (\<lambda>x. exp (g x)))
|
|
4020 |
((\<lambda>x. exp (-f x)) # b # basis)" (is ?thesis3)
|
|
4021 |
proof -
|
|
4022 |
have "ln \<in> \<omega>(\<lambda>x::real. 1)"
|
|
4023 |
by (intro smallomegaI_filterlim_at_top_norm)
|
|
4024 |
(auto intro!: filterlim_compose[OF filterlim_abs_real] ln_at_top)
|
|
4025 |
with assms have "(\<lambda>x. 1) \<in> o(\<lambda>x. ln (b x))"
|
|
4026 |
by (intro landau_o.small.compose[of _ at_top _ b])
|
|
4027 |
(simp_all add: basis_wf_Cons smallomega_iff_smallo)
|
|
4028 |
also note \<open>(\<lambda>x. ln (b x)) \<in> o(f)\<close>
|
|
4029 |
finally have f: "f \<in> \<omega>(\<lambda>_. 1)" by (simp add: smallomega_iff_smallo)
|
|
4030 |
hence f: "filterlim (\<lambda>x. abs (f x)) at_top at_top"
|
|
4031 |
by (auto dest: smallomegaD_filterlim_at_top_norm)
|
|
4032 |
from trimmed_imp_eventually_sgn[OF assms(2), of "MS (MSLCons (C, e) xs) g"] assms
|
|
4033 |
have "\<forall>\<^sub>F x in at_top. abs (f x) = -f x"
|
|
4034 |
by (auto simp: dominant_term_ms_aux_MSLCons trimmed_neg_def expands_to.simps sgn_if
|
|
4035 |
split: if_splits elim: eventually_elim2)
|
|
4036 |
from filterlim_cong[OF refl refl this] f
|
|
4037 |
show *: "filterlim (\<lambda>x. exp (-f x)) at_top at_top"
|
|
4038 |
by (auto intro: filterlim_compose[OF exp_at_top])
|
|
4039 |
|
|
4040 |
have "((\<lambda>x. -f x) expands_to -MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
4041 |
by (intro expands_to_uminus assms)
|
|
4042 |
thus "((\<lambda>x. ln (exp (-f x))) expands_to MS (MSLCons (-C, e) (uminus_ms_aux xs)) (\<lambda>x. - g x))
|
|
4043 |
(b # basis)"
|
|
4044 |
by (simp add: uminus_ms_aux_MSLCons)
|
|
4045 |
|
|
4046 |
{
|
|
4047 |
fix e' :: real assume e': "e' > -1"
|
|
4048 |
from assms have "(\<lambda>x. exp (g x)) \<in> \<Theta>(\<lambda>x. exp (-f x) powr -1)"
|
|
4049 |
by (intro bigthetaI_cong)
|
|
4050 |
(auto elim: eventually_mono simp: expands_to.simps powr_minus exp_minus)
|
|
4051 |
also from e' * have "(\<lambda>x. exp (-f x) powr -1) \<in> o(\<lambda>x. exp (-f x) powr e')"
|
|
4052 |
by (subst powr_smallo_iff) (insert *, simp_all)
|
|
4053 |
finally have "(\<lambda>x. exp (g x)) \<in> o(\<lambda>x. exp (-f x) powr e')" .
|
|
4054 |
}
|
|
4055 |
with assms show ?thesis3
|
|
4056 |
by (auto intro!: is_expansion_aux.intros is_expansion_const
|
|
4057 |
simp: expands_to.simps is_expansion_aux_expansion_level powr_minus exp_minus
|
|
4058 |
dest: is_expansion_aux_expansion_level)
|
|
4059 |
qed (insert assms, simp_all add: trimmed_pos_def trimmed_neg_def
|
|
4060 |
trimmed_ms_aux_MSLCons dominant_term_ms_aux_MSLCons)
|
|
4061 |
|
|
4062 |
lemma is_expansion_aux_exp_neg:
|
|
4063 |
assumes "is_expansion_aux F f basis" "basis_wf basis" "ms_exp_gt 0 (ms_aux_hd_exp F)"
|
|
4064 |
shows "is_expansion_aux (powser_ms_aux' exp_series_stream F) (\<lambda>x. exp (f x)) basis"
|
|
4065 |
using powser_ms_aux'[OF convergent_powser'_exp assms(2,1,3)]
|
|
4066 |
by (simp add: o_def powser_ms_aux'_def)
|
|
4067 |
|
|
4068 |
lemma is_expansion_exp_neg:
|
|
4069 |
assumes "is_expansion (MS (MSLCons (C, e) xs) f) basis" "basis_wf basis" "e < 0"
|
|
4070 |
shows "is_expansion (MS (powser_ms_aux' exp_series_stream (MSLCons (C, e) xs))
|
|
4071 |
(\<lambda>x. exp (f x))) basis"
|
|
4072 |
using is_expansion_aux_exp_neg[of "MSLCons (C, e) xs" f basis] assms by simp
|
|
4073 |
|
|
4074 |
lemma expands_to_exp_neg:
|
|
4075 |
assumes "(f expands_to MS (MSLCons (C, e) xs) g) basis" "basis_wf basis" "e - 0 < 0"
|
|
4076 |
shows "((\<lambda>x. exp (f x)) expands_to MS (powser_ms_aux' exp_series_stream (MSLCons (C, e) xs))
|
|
4077 |
(\<lambda>x. exp (g x))) basis"
|
|
4078 |
using assms is_expansion_exp_neg[of C e xs g basis] by (auto simp: expands_to.simps)
|
|
4079 |
|
|
4080 |
lemma basis_wf_insert_exp_near:
|
|
4081 |
assumes "basis_wf (b # basis)" "basis_wf ((\<lambda>x. exp (f x)) # basis)" "f \<in> o(\<lambda>x. ln (b x))"
|
|
4082 |
shows "basis_wf (b # (\<lambda>x. exp (f x)) # basis)"
|
|
4083 |
using assms by (simp_all add: basis_wf_Cons)
|
|
4084 |
|
|
4085 |
|
|
4086 |
definition scale_ms_aux' where "scale_ms_aux' c F = scale_shift_ms_aux (c, 0) F"
|
|
4087 |
definition shift_ms_aux where "shift_ms_aux e F = scale_shift_ms_aux (const_expansion 1, e) F"
|
|
4088 |
|
|
4089 |
lemma scale_ms_1 [simp]: "scale_ms 1 F = F"
|
|
4090 |
by (simp add: scale_ms_def times_const_expansion_1)
|
|
4091 |
|
|
4092 |
lemma scale_ms_aux_1 [simp]: "scale_ms_aux 1 xs = xs"
|
|
4093 |
by (simp add: scale_ms_aux_def scale_shift_ms_aux_def map_prod_def msllist.map_ident
|
|
4094 |
case_prod_unfold times_const_expansion_1)
|
|
4095 |
|
|
4096 |
lemma scale_ms_aux'_1 [simp]: "scale_ms_aux' (const_expansion 1) xs = xs"
|
|
4097 |
by (simp add: scale_ms_aux'_def scale_shift_ms_aux_def map_prod_def msllist.map_ident
|
|
4098 |
case_prod_unfold times_const_expansion_1)
|
|
4099 |
|
|
4100 |
lemma shift_ms_aux_0 [simp]: "shift_ms_aux 0 xs = xs"
|
|
4101 |
by (simp add: shift_ms_aux_def scale_shift_ms_aux_def map_prod_def case_prod_unfold
|
|
4102 |
times_const_expansion_1 msllist.map_ident)
|
|
4103 |
|
|
4104 |
lemma scale_ms_aux'_MSLNil [simp]: "scale_ms_aux' C MSLNil = MSLNil"
|
|
4105 |
by (simp add: scale_ms_aux'_def)
|
|
4106 |
|
|
4107 |
lemma scale_ms_aux'_MSLCons [simp]:
|
|
4108 |
"scale_ms_aux' C (MSLCons (C', e) xs) = MSLCons (C * C', e) (scale_ms_aux' C xs)"
|
|
4109 |
by (simp add: scale_ms_aux'_def scale_shift_ms_aux_MSLCons)
|
|
4110 |
|
|
4111 |
lemma shift_ms_aux_MSLNil [simp]: "shift_ms_aux e MSLNil = MSLNil"
|
|
4112 |
by (simp add: shift_ms_aux_def)
|
|
4113 |
|
|
4114 |
lemma shift_ms_aux_MSLCons [simp]:
|
|
4115 |
"shift_ms_aux e (MSLCons (C, e') xs) = MSLCons (C, e' + e) (shift_ms_aux e xs)"
|
|
4116 |
by (simp add: shift_ms_aux_def scale_shift_ms_aux_MSLCons times_const_expansion_1)
|
|
4117 |
|
|
4118 |
lemma scale_ms_aux:
|
|
4119 |
fixes xs :: "('a :: multiseries \<times> real) msllist"
|
|
4120 |
assumes "is_expansion_aux xs f basis" "basis_wf basis"
|
|
4121 |
shows "is_expansion_aux (scale_ms_aux c xs) (\<lambda>x. c * f x) basis"
|
|
4122 |
proof (cases basis)
|
|
4123 |
case (Cons b basis')
|
|
4124 |
show ?thesis
|
|
4125 |
proof (rule is_expansion_aux_cong)
|
|
4126 |
show "is_expansion_aux (scale_ms_aux c xs)
|
|
4127 |
(\<lambda>x. eval (const_expansion c :: 'a) x * b x powr 0 * f x) basis"
|
|
4128 |
using assms unfolding scale_ms_aux_def Cons
|
|
4129 |
by (intro scale_shift_ms_aux is_expansion_const)
|
|
4130 |
(auto simp: basis_wf_Cons dest: is_expansion_aux_expansion_level)
|
|
4131 |
next
|
|
4132 |
from assms have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4133 |
by (simp add: basis_wf_Cons Cons filterlim_at_top_dense)
|
|
4134 |
thus "eventually (\<lambda>x. eval (const_expansion c :: 'a) x * b x powr 0 * f x = c * f x) at_top"
|
|
4135 |
by eventually_elim simp
|
|
4136 |
qed
|
|
4137 |
qed (insert assms, auto dest: is_expansion_aux_basis_nonempty)
|
|
4138 |
|
|
4139 |
lemma scale_ms_aux':
|
|
4140 |
assumes "is_expansion_aux xs f (b # basis)" "is_expansion C basis" "basis_wf (b # basis)"
|
|
4141 |
shows "is_expansion_aux (scale_ms_aux' C xs) (\<lambda>x. eval C x * f x) (b # basis)"
|
|
4142 |
proof (rule is_expansion_aux_cong)
|
|
4143 |
show "is_expansion_aux (scale_ms_aux' C xs) (\<lambda>x. eval C x * b x powr 0 * f x) (b # basis)"
|
|
4144 |
using assms unfolding scale_ms_aux'_def Cons by (intro scale_shift_ms_aux) simp_all
|
|
4145 |
next
|
|
4146 |
from assms have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4147 |
by (simp add: basis_wf_Cons filterlim_at_top_dense)
|
|
4148 |
thus "eventually (\<lambda>x. eval C x * b x powr 0 * f x = eval C x * f x) at_top"
|
|
4149 |
by eventually_elim simp
|
|
4150 |
qed
|
|
4151 |
|
|
4152 |
lemma shift_ms_aux:
|
|
4153 |
fixes xs :: "('a :: multiseries \<times> real) msllist"
|
|
4154 |
assumes "is_expansion_aux xs f (b # basis)" "basis_wf (b # basis)"
|
|
4155 |
shows "is_expansion_aux (shift_ms_aux e xs) (\<lambda>x. b x powr e * f x) (b # basis)"
|
|
4156 |
unfolding shift_ms_aux_def
|
|
4157 |
using scale_shift_ms_aux[OF assms(2,1) is_expansion_const[of _ 1], of e] assms
|
|
4158 |
by (auto dest!: is_expansion_aux_expansion_level simp: basis_wf_Cons)
|
|
4159 |
|
|
4160 |
lemma expands_to_exp_0:
|
|
4161 |
assumes "(f expands_to MS (MSLCons (MS ys c, e) xs) g) (b # basis)"
|
|
4162 |
"basis_wf (b # basis)" "e - 0 = 0" "((\<lambda>x. exp (c x)) expands_to E) basis"
|
|
4163 |
shows "((\<lambda>x. exp (f x)) expands_to
|
|
4164 |
MS (scale_ms_aux' E (powser_ms_aux' exp_series_stream xs)) (\<lambda>x. exp (g x))) (b # basis)"
|
|
4165 |
(is "(_ expands_to MS ?H _) _")
|
|
4166 |
proof -
|
|
4167 |
from assms have "is_expansion_aux ?H (\<lambda>x. eval E x * exp (g x - c x * b x powr 0)) (b # basis)"
|
|
4168 |
by (intro scale_ms_aux' is_expansion_aux_exp_neg)
|
|
4169 |
(auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps)
|
|
4170 |
moreover {
|
|
4171 |
from \<open>basis_wf (b # basis)\<close> have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4172 |
by (simp add: filterlim_at_top_dense basis_wf_Cons)
|
|
4173 |
moreover from assms(4) have "eventually (\<lambda>x. eval E x = exp (c x)) at_top"
|
|
4174 |
by (simp add: expands_to.simps)
|
|
4175 |
ultimately have "eventually (\<lambda>x. eval E x * exp (g x - c x * b x powr 0) = exp (g x)) at_top"
|
|
4176 |
by eventually_elim (simp add: exp_diff)
|
|
4177 |
}
|
|
4178 |
ultimately have "is_expansion_aux ?H (\<lambda>x. exp (g x)) (b # basis)"
|
|
4179 |
by (rule is_expansion_aux_cong)
|
|
4180 |
with assms(1) show ?thesis by (simp add: expands_to.simps)
|
|
4181 |
qed
|
|
4182 |
|
|
4183 |
lemma expands_to_exp_0_real:
|
|
4184 |
assumes "(f expands_to MS (MSLCons (c::real, e) xs) g) [b]" "basis_wf [b]" "e - 0 = 0"
|
|
4185 |
shows "((\<lambda>x. exp (f x)) expands_to
|
|
4186 |
MS (scale_ms_aux (exp c) (powser_ms_aux' exp_series_stream xs)) (\<lambda>x. exp (g x))) [b]"
|
|
4187 |
(is "(_ expands_to MS ?H _) _")
|
|
4188 |
proof -
|
|
4189 |
from assms have "is_expansion_aux ?H (\<lambda>x. exp c * exp (g x - c * b x powr 0)) [b]"
|
|
4190 |
by (intro scale_ms_aux is_expansion_aux_exp_neg)
|
|
4191 |
(auto elim!: is_expansion_aux_MSLCons simp: expands_to.simps)
|
|
4192 |
moreover from \<open>basis_wf [b]\<close> have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4193 |
by (simp add: filterlim_at_top_dense basis_wf_Cons)
|
|
4194 |
hence "eventually (\<lambda>x. exp c * exp (g x - c * b x powr 0) = exp (g x)) at_top"
|
|
4195 |
by eventually_elim (simp add: exp_diff)
|
|
4196 |
ultimately have "is_expansion_aux ?H (\<lambda>x. exp (g x)) [b]"
|
|
4197 |
by (rule is_expansion_aux_cong)
|
|
4198 |
with assms(1) show ?thesis by (simp add: expands_to.simps)
|
|
4199 |
qed
|
|
4200 |
|
|
4201 |
lemma expands_to_exp_0_pull_out1:
|
|
4202 |
assumes "(f expands_to MS (MSLCons (C, e) xs) g) (b # basis)"
|
|
4203 |
assumes "((\<lambda>x. ln (b x)) expands_to L) basis" "basis_wf (b # basis)" "e - 0 = 0" "c \<equiv> c"
|
|
4204 |
shows "((\<lambda>x. f x - c * ln (b x)) expands_to
|
|
4205 |
MS (MSLCons (C - scale_ms c L, e) xs) (\<lambda>x. g x - c * ln (b x))) (b # basis)"
|
|
4206 |
proof -
|
|
4207 |
from \<open>basis_wf (b # basis)\<close> have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
4208 |
have "(\<lambda>x. c * ln (b x)) \<in> o(\<lambda>x. b x powr e')" if "e' > 0" for e'
|
|
4209 |
using that by (intro landau_o.small.compose[OF _ b_limit]) (simp add: ln_smallo_powr)
|
|
4210 |
hence *: "(\<lambda>x. g x - c * ln (b x)) \<in> o(\<lambda>x. b x powr e')" if "e' > 0" for e' using assms(1,4) that
|
|
4211 |
by (intro sum_in_smallo) (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons)
|
|
4212 |
|
|
4213 |
have "is_expansion_aux xs (\<lambda>x. (g x - c * b x powr 0 * eval L x) -
|
|
4214 |
eval (C - scale_ms c L) x * b x powr e) (b # basis)"
|
|
4215 |
(is "is_expansion_aux _ ?h _") using assms
|
|
4216 |
by (auto simp: expands_to.simps algebra_simps scale_ms_def elim!: is_expansion_aux_MSLCons)
|
|
4217 |
moreover from b_limit have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4218 |
by (simp add: filterlim_at_top_dense)
|
|
4219 |
hence "eventually (\<lambda>x. ?h x = g x - c * ln (b x) -
|
|
4220 |
eval (C - const_expansion c * L) x * b x powr e) at_top"
|
|
4221 |
(is "eventually (\<lambda>x. _ = ?h x) _") using assms(2)
|
|
4222 |
by (auto simp: expands_to.simps scale_ms_def elim: eventually_elim2)
|
|
4223 |
ultimately have **: "is_expansion_aux xs ?h (b # basis)" by (rule is_expansion_aux_cong)
|
|
4224 |
from assms(1,4) have "ms_exp_gt 0 (ms_aux_hd_exp xs)" "is_expansion C basis"
|
|
4225 |
by (auto simp: expands_to.simps elim!: is_expansion_aux_MSLCons)
|
|
4226 |
moreover from assms(1) have "length basis = expansion_level TYPE('a)"
|
|
4227 |
by (auto simp: expands_to.simps dest: is_expansion_aux_expansion_level)
|
|
4228 |
ultimately have "is_expansion_aux (MSLCons (C - scale_ms c L, e) xs)
|
|
4229 |
(\<lambda>x. g x - c * ln (b x)) (b # basis)" using assms unfolding scale_ms_def
|
|
4230 |
by (intro is_expansion_aux.intros is_expansion_minus is_expansion_mult is_expansion_const * **)
|
|
4231 |
(simp_all add: basis_wf_Cons expands_to.simps)
|
|
4232 |
with assms(1) show ?thesis by (auto simp: expands_to.simps)
|
|
4233 |
qed
|
|
4234 |
|
|
4235 |
lemma expands_to_exp_0_pull_out2:
|
|
4236 |
assumes "((\<lambda>x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)"
|
|
4237 |
assumes "basis_wf (b # basis)"
|
|
4238 |
shows "((\<lambda>x. exp (f x)) expands_to MS (shift_ms_aux c xs)
|
|
4239 |
(\<lambda>x. b x powr c * g x)) (b # basis)"
|
|
4240 |
proof (rule expands_to.intros)
|
|
4241 |
show "is_expansion (MS (shift_ms_aux c xs) (\<lambda>x. b x powr c * g x)) (b # basis)"
|
|
4242 |
using assms unfolding expands_to.simps by (auto intro!: shift_ms_aux)
|
|
4243 |
from assms have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4244 |
by (simp add: basis_wf_Cons filterlim_at_top_dense)
|
|
4245 |
with assms(1)
|
|
4246 |
show "\<forall>\<^sub>F x in at_top. eval (MS (shift_ms_aux c xs) (\<lambda>x. b x powr c * g x)) x = exp (f x)"
|
|
4247 |
by (auto simp: expands_to.simps exp_diff powr_def elim: eventually_elim2)
|
|
4248 |
qed
|
|
4249 |
|
|
4250 |
lemma expands_to_exp_0_pull_out2_nat:
|
|
4251 |
assumes "((\<lambda>x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)"
|
|
4252 |
assumes "basis_wf (b # basis)" "c = real n"
|
|
4253 |
shows "((\<lambda>x. exp (f x)) expands_to MS (shift_ms_aux c xs)
|
|
4254 |
(\<lambda>x. b x ^ n * g x)) (b # basis)"
|
|
4255 |
using expands_to_exp_0_pull_out2[OF assms(1-2)]
|
|
4256 |
proof (rule expands_to_cong')
|
|
4257 |
from assms(2) have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4258 |
by (simp add: filterlim_at_top_dense basis_wf_Cons)
|
|
4259 |
with assms(3) show "eventually (\<lambda>x. b x powr c * g x = b x ^ n * g x) at_top"
|
|
4260 |
by (auto elim!: eventually_mono simp: powr_realpow)
|
|
4261 |
qed
|
|
4262 |
|
|
4263 |
lemma expands_to_exp_0_pull_out2_neg_nat:
|
|
4264 |
assumes "((\<lambda>x. exp (f x - c * ln (b x))) expands_to MS xs g) (b # basis)"
|
|
4265 |
assumes "basis_wf (b # basis)" "c = -real n"
|
|
4266 |
shows "((\<lambda>x. exp (f x)) expands_to MS (shift_ms_aux c xs)
|
|
4267 |
(\<lambda>x. g x / b x ^ n)) (b # basis)"
|
|
4268 |
using expands_to_exp_0_pull_out2[OF assms(1-2)]
|
|
4269 |
proof (rule expands_to_cong')
|
|
4270 |
from assms(2) have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4271 |
by (simp add: filterlim_at_top_dense basis_wf_Cons)
|
|
4272 |
with assms(3) show "eventually (\<lambda>x. b x powr c * g x = g x / b x ^ n) at_top"
|
|
4273 |
by (auto elim!: eventually_mono simp: powr_realpow powr_minus divide_simps)
|
|
4274 |
qed
|
|
4275 |
|
|
4276 |
lemma eval_monom_collapse: "c * eval_monom (c', es) basis x = eval_monom (c * c', es) basis x"
|
|
4277 |
by (simp add: eval_monom_def)
|
|
4278 |
|
|
4279 |
lemma eval_monom_1_conv: "eval_monom a basis = (\<lambda>x. fst a * eval_monom (1, snd a) basis x)"
|
|
4280 |
by (simp add: eval_monom_def case_prod_unfold)
|
|
4281 |
|
|
4282 |
|
|
4283 |
subsubsection \<open>Comparing exponent lists\<close>
|
|
4284 |
|
|
4285 |
primrec flip_cmp_result where
|
|
4286 |
"flip_cmp_result LT = GT"
|
|
4287 |
| "flip_cmp_result GT = LT"
|
|
4288 |
| "flip_cmp_result EQ = EQ"
|
|
4289 |
|
|
4290 |
fun compare_lists :: "real list \<Rightarrow> real list \<Rightarrow> cmp_result" where
|
|
4291 |
"compare_lists [] [] = EQ"
|
|
4292 |
| "compare_lists (a # as) (b # bs) =
|
|
4293 |
(if a < b then LT else if b < a then GT else compare_lists as bs)"
|
|
4294 |
| "compare_lists _ _ = undefined"
|
|
4295 |
|
|
4296 |
declare compare_lists.simps [simp del]
|
|
4297 |
|
|
4298 |
lemma compare_lists_Nil [simp]: "compare_lists [] [] = EQ" by (fact compare_lists.simps)
|
|
4299 |
|
|
4300 |
lemma compare_lists_Cons [simp]:
|
|
4301 |
"a < b \<Longrightarrow> compare_lists (a # as) (b # bs) = LT"
|
|
4302 |
"a > b \<Longrightarrow> compare_lists (a # as) (b # bs) = GT"
|
|
4303 |
"a = b \<Longrightarrow> compare_lists (a # as) (b # bs) = compare_lists as bs"
|
|
4304 |
by (simp_all add: compare_lists.simps(2))
|
|
4305 |
|
|
4306 |
lemma flip_compare_lists:
|
|
4307 |
"length as = length bs \<Longrightarrow> flip_cmp_result (compare_lists as bs) = compare_lists bs as"
|
|
4308 |
by (induction as bs rule: compare_lists.induct) (auto simp: compare_lists.simps(2))
|
|
4309 |
|
|
4310 |
lemma compare_lists_induct [consumes 1, case_names Nil Eq Neq]:
|
|
4311 |
fixes as bs :: "real list"
|
|
4312 |
assumes "length as = length bs"
|
|
4313 |
assumes "P [] []"
|
|
4314 |
assumes "\<And>a as bs. P as bs \<Longrightarrow> P (a # as) (a # bs)"
|
|
4315 |
assumes "\<And>a as b bs. a \<noteq> b \<Longrightarrow> P (a # as) (b # bs)"
|
|
4316 |
shows "P as bs"
|
|
4317 |
using assms(1)
|
|
4318 |
proof (induction as bs rule: list_induct2)
|
|
4319 |
case (Cons a as b bs)
|
|
4320 |
thus ?case by (cases "a < b") (auto intro: assms)
|
|
4321 |
qed (simp_all add: assms)
|
|
4322 |
|
|
4323 |
|
|
4324 |
lemma eval_monom_smallo_eval_monom:
|
|
4325 |
assumes "length es1 = length es2" "length es2 = length basis" "basis_wf basis"
|
|
4326 |
assumes "compare_lists es1 es2 = LT"
|
|
4327 |
shows "eval_monom (1, es1) basis \<in> o(eval_monom (1, es2) basis)"
|
|
4328 |
using assms
|
|
4329 |
proof (induction es1 es2 basis rule: list_induct3)
|
|
4330 |
case (Cons e1 es1 e2 es2 b basis)
|
|
4331 |
show ?case
|
|
4332 |
proof (cases "e1 = e2")
|
|
4333 |
case True
|
|
4334 |
from Cons.prems have "eventually (\<lambda>x. b x > 0) at_top"
|
|
4335 |
by (simp add: basis_wf_Cons filterlim_at_top_dense)
|
|
4336 |
with Cons True show ?thesis
|
|
4337 |
by (auto intro: landau_o.small_big_mult simp: eval_monom_Cons basis_wf_Cons)
|
|
4338 |
next
|
|
4339 |
case False
|
|
4340 |
with Cons.prems have "e1 < e2" by (cases e1 e2 rule: linorder_cases) simp_all
|
|
4341 |
with Cons.prems have
|
|
4342 |
"(\<lambda>x. eval_monom (1, es1) basis x * eval_monom (inverse_monom (1, es2)) basis x *
|
|
4343 |
b x powr e1) \<in> o(\<lambda>x. b x powr ((e2 - e1)/2) * b x powr ((e2 - e1)/2) * b x powr e1)"
|
|
4344 |
(is "?f \<in> _") by (intro landau_o.small_big_mult[OF _ landau_o.big_refl] landau_o.small.mult
|
|
4345 |
eval_monom_smallo') simp_all
|
|
4346 |
also have "\<dots> = o(\<lambda>x. b x powr e2)" by (simp add: powr_add [symmetric])
|
|
4347 |
also have "?f = (\<lambda>x. eval_monom (1, e1 # es1) (b # basis) x / eval_monom (1, es2) basis x)"
|
|
4348 |
by (simp add: eval_inverse_monom field_simps eval_monom_Cons)
|
|
4349 |
also have "\<dots> \<in> o(\<lambda>x. b x powr e2) \<longleftrightarrow>
|
|
4350 |
eval_monom (1, e1 # es1) (b # basis) \<in> o(eval_monom (1, e2 # es2) (b # basis))"
|
|
4351 |
using Cons.prems
|
|
4352 |
by (subst landau_o.small.divide_eq2)
|
|
4353 |
(simp_all add: eval_monom_Cons mult_ac eval_monom_nonzero basis_wf_Cons)
|
|
4354 |
finally show ?thesis .
|
|
4355 |
qed
|
|
4356 |
qed simp_all
|
|
4357 |
|
|
4358 |
lemma eval_monom_eq:
|
|
4359 |
assumes "length es1 = length es2" "length es2 = length basis" "basis_wf basis"
|
|
4360 |
assumes "compare_lists es1 es2 = EQ"
|
|
4361 |
shows "eval_monom (1, es1) basis = eval_monom (1, es2) basis"
|
|
4362 |
using assms
|
|
4363 |
by (induction es1 es2 basis rule: list_induct3)
|
|
4364 |
(auto simp: basis_wf_Cons eval_monom_Cons compare_lists.simps(2) split: if_splits)
|
|
4365 |
|
|
4366 |
definition compare_expansions :: "'a :: multiseries \<Rightarrow> 'a \<Rightarrow> cmp_result \<times> real \<times> real" where
|
|
4367 |
"compare_expansions F G =
|
|
4368 |
(case compare_lists (snd (dominant_term F)) (snd (dominant_term G)) of
|
|
4369 |
LT \<Rightarrow> (LT, 0, 0)
|
|
4370 |
| GT \<Rightarrow> (GT, 0, 0)
|
|
4371 |
| EQ \<Rightarrow> (EQ, fst (dominant_term F), fst (dominant_term G)))"
|
|
4372 |
|
|
4373 |
lemma compare_expansions_real:
|
|
4374 |
"compare_expansions (c1 :: real) c2 = (EQ, c1, c2)"
|
|
4375 |
by (simp add: compare_expansions_def)
|
|
4376 |
|
|
4377 |
lemma compare_expansions_MSLCons_eval:
|
|
4378 |
"compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) =
|
|
4379 |
CMP_BRANCH (COMPARE e1 e2) (LT, 0, 0) (compare_expansions C1 C2) (GT, 0, 0)"
|
|
4380 |
by (simp add: compare_expansions_def dominant_term_ms_aux_def case_prod_unfold
|
|
4381 |
COMPARE_def CMP_BRANCH_def split: cmp_result.splits)
|
|
4382 |
|
|
4383 |
lemma compare_expansions_LT_I:
|
|
4384 |
assumes "e1 - e2 < 0"
|
|
4385 |
shows "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = (LT, 0, 0)"
|
|
4386 |
using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def)
|
|
4387 |
|
|
4388 |
lemma compare_expansions_GT_I:
|
|
4389 |
assumes "e1 - e2 > 0"
|
|
4390 |
shows "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = (GT, 0, 0)"
|
|
4391 |
using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def)
|
|
4392 |
|
|
4393 |
lemma compare_expansions_same_exp:
|
|
4394 |
assumes "e1 - e2 = 0" "compare_expansions C1 C2 = res"
|
|
4395 |
shows "compare_expansions (MS (MSLCons (C1, e1) xs) f) (MS (MSLCons (C2, e2) ys) g) = res"
|
|
4396 |
using assms by (simp add: compare_expansions_MSLCons_eval CMP_BRANCH_def COMPARE_def)
|
|
4397 |
|
|
4398 |
|
|
4399 |
lemma compare_expansions_GT_flip:
|
|
4400 |
"length (snd (dominant_term F)) = length (snd (dominant_term G)) \<Longrightarrow>
|
|
4401 |
compare_expansions F G = (GT, c) \<longleftrightarrow> compare_expansions G F = (LT, c)"
|
|
4402 |
using flip_compare_lists[of "snd (dominant_term F)" "snd (dominant_term G)"]
|
|
4403 |
by (auto simp: compare_expansions_def split: cmp_result.splits)
|
|
4404 |
|
|
4405 |
lemma compare_expansions_LT:
|
|
4406 |
assumes "compare_expansions F G = (LT, c, c')" "trimmed G"
|
|
4407 |
"(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
|
|
4408 |
shows "f \<in> o(g)"
|
|
4409 |
proof -
|
|
4410 |
from assms have "f \<in> \<Theta>(eval F)"
|
|
4411 |
by (auto simp: expands_to.simps eq_commute intro!: bigthetaI_cong)
|
|
4412 |
also have "eval F \<in> O(eval_monom (1, snd (dominant_term F)) basis)"
|
|
4413 |
using assms by (intro dominant_term_bigo) (auto simp: expands_to.simps)
|
|
4414 |
also have "eval_monom (1, snd (dominant_term F)) basis \<in>
|
|
4415 |
o(eval_monom (1, snd (dominant_term G)) basis)" using assms
|
|
4416 |
by (intro eval_monom_smallo_eval_monom)
|
|
4417 |
(auto simp: length_dominant_term expands_to.simps is_expansion_length compare_expansions_def
|
|
4418 |
split: cmp_result.splits)
|
|
4419 |
also have "eval_monom (1, snd (dominant_term G)) basis \<in> \<Theta>(eval_monom (dominant_term G) basis)"
|
|
4420 |
by (subst (2) eval_monom_1_conv, subst landau_theta.cmult)
|
|
4421 |
(insert assms, simp_all add: trimmed_imp_dominant_term_nz)
|
|
4422 |
also have "eval_monom (dominant_term G) basis \<in> \<Theta>(g)"
|
|
4423 |
by (intro asymp_equiv_imp_bigtheta[OF asymp_equiv_symI] dominant_term_expands_to
|
|
4424 |
asymp_equiv_imp_bigtheta assms)
|
|
4425 |
finally show ?thesis .
|
|
4426 |
qed
|
|
4427 |
|
|
4428 |
lemma compare_expansions_GT:
|
|
4429 |
assumes "compare_expansions F G = (GT, c, c')" "trimmed F"
|
|
4430 |
"(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
|
|
4431 |
shows "g \<in> o(f)"
|
|
4432 |
by (rule compare_expansions_LT[OF _ assms(2,4,3)])
|
|
4433 |
(insert assms, simp_all add: compare_expansions_GT_flip length_dominant_term)
|
|
4434 |
|
|
4435 |
lemma compare_expansions_EQ:
|
|
4436 |
assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G"
|
|
4437 |
"(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
|
|
4438 |
shows "(\<lambda>x. c' * f x) \<sim>[at_top] (\<lambda>x. c * g x)"
|
|
4439 |
proof -
|
|
4440 |
from assms(1) have c: "c = fst (dominant_term F)" "c' = fst (dominant_term G)"
|
|
4441 |
by (auto simp: compare_expansions_def split: cmp_result.splits)
|
|
4442 |
have "(\<lambda>x. c' * f x) \<sim>[at_top] (\<lambda>x. c' * (c * eval_monom (1, snd (dominant_term F)) basis x))"
|
|
4443 |
unfolding c by (rule asymp_equiv_mult, rule asymp_equiv_refl, subst eval_monom_collapse)
|
|
4444 |
(auto intro: dominant_term_expands_to assms)
|
|
4445 |
also have "eval_monom (1, snd (dominant_term F)) basis =
|
|
4446 |
eval_monom (1, snd (dominant_term G)) basis" using assms
|
|
4447 |
by (intro eval_monom_eq)
|
|
4448 |
(simp_all add: compare_expansions_def length_dominant_term is_expansion_length
|
|
4449 |
expands_to.simps split: cmp_result.splits)
|
|
4450 |
also have "(\<lambda>x. c' * (c * \<dots> x)) = (\<lambda>x. c * (c' * \<dots> x))" by (simp add: mult_ac)
|
|
4451 |
also have "\<dots> \<sim>[at_top] (\<lambda>x. c * g x)"
|
|
4452 |
unfolding c by (rule asymp_equiv_mult, rule asymp_equiv_refl, subst eval_monom_collapse,
|
|
4453 |
rule asymp_equiv_symI) (auto intro: dominant_term_expands_to assms)
|
|
4454 |
finally show ?thesis by (simp add: asymp_equiv_sym)
|
|
4455 |
qed
|
|
4456 |
|
|
4457 |
lemma compare_expansions_EQ_imp_bigo:
|
|
4458 |
assumes "compare_expansions F G = (EQ, c, c')" "trimmed G"
|
|
4459 |
"(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
|
|
4460 |
shows "f \<in> O(g)"
|
|
4461 |
proof -
|
|
4462 |
from assms(1) have c: "c = fst (dominant_term F)" "c' = fst (dominant_term G)"
|
|
4463 |
by (auto simp: compare_expansions_def split: cmp_result.splits)
|
|
4464 |
from assms(3) have [simp]: "expansion_level TYPE('a) = length basis"
|
|
4465 |
by (auto simp: expands_to.simps is_expansion_length)
|
|
4466 |
|
|
4467 |
have "f \<in> \<Theta>(eval F)"
|
|
4468 |
using assms by (intro bigthetaI_cong) (auto simp: expands_to.simps eq_commute)
|
|
4469 |
also have "eval F \<in> O(eval_monom (1, snd (dominant_term F)) basis)"
|
|
4470 |
using assms by (intro dominant_term_bigo assms) (auto simp: expands_to.simps)
|
|
4471 |
also have "eval_monom (1, snd (dominant_term F)) basis =
|
|
4472 |
eval_monom (1, snd (dominant_term G)) basis" using assms
|
|
4473 |
by (intro eval_monom_eq) (auto simp: compare_expansions_def case_prod_unfold
|
|
4474 |
length_dominant_term split: cmp_result.splits)
|
|
4475 |
also have "\<dots> \<in> O(eval_monom (dominant_term G) basis)" using assms(2)
|
|
4476 |
by (auto simp add: eval_monom_def case_prod_unfold dest: trimmed_imp_dominant_term_nz)
|
|
4477 |
also have "eval_monom (dominant_term G) basis \<in> \<Theta>(eval G)"
|
|
4478 |
by (rule asymp_equiv_imp_bigtheta, rule asymp_equiv_symI, rule dominant_term)
|
|
4479 |
(insert assms, auto simp: expands_to.simps)
|
|
4480 |
also have "eval G \<in> \<Theta>(g)"
|
|
4481 |
using assms by (intro bigthetaI_cong) (auto simp: expands_to.simps)
|
|
4482 |
finally show ?thesis .
|
|
4483 |
qed
|
|
4484 |
|
|
4485 |
lemma compare_expansions_EQ_same:
|
|
4486 |
assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G"
|
|
4487 |
"(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
|
|
4488 |
"c = c'"
|
|
4489 |
shows "f \<sim>[at_top] g"
|
|
4490 |
proof -
|
|
4491 |
from assms have [simp]: "c' \<noteq> 0"
|
|
4492 |
by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits)
|
|
4493 |
have "(\<lambda>x. inverse c * (c' * f x)) \<sim>[at_top] (\<lambda>x. inverse c * (c * g x))"
|
|
4494 |
by (rule asymp_equiv_mult[OF asymp_equiv_refl]) (rule compare_expansions_EQ[OF assms(1-6)])
|
|
4495 |
with assms(7) show ?thesis by (simp add: divide_simps)
|
|
4496 |
qed
|
|
4497 |
|
|
4498 |
lemma compare_expansions_EQ_imp_bigtheta:
|
|
4499 |
assumes "compare_expansions F G = (EQ, c, c')" "trimmed F" "trimmed G"
|
|
4500 |
"(f expands_to F) basis" "(g expands_to G) basis" "basis_wf basis"
|
|
4501 |
shows "f \<in> \<Theta>(g)"
|
|
4502 |
proof -
|
|
4503 |
from assms have "(\<lambda>x. c' * f x) \<in> \<Theta>(\<lambda>x. c * g x)"
|
|
4504 |
by (intro asymp_equiv_imp_bigtheta compare_expansions_EQ)
|
|
4505 |
moreover from assms have "c \<noteq> 0" "c' \<noteq> 0"
|
|
4506 |
by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits)
|
|
4507 |
ultimately show ?thesis by simp
|
|
4508 |
qed
|
|
4509 |
|
|
4510 |
lemma expands_to_MSLCons_0_asymp_equiv_hd:
|
|
4511 |
assumes "(f expands_to (MS (MSLCons (C, 0) xs) g)) basis" "trimmed (MS (MSLCons (C, 0) xs) f)"
|
|
4512 |
"basis_wf basis"
|
|
4513 |
shows "f \<sim>[at_top] eval C"
|
|
4514 |
proof -
|
|
4515 |
from assms(1) is_expansion_aux_basis_nonempty obtain b basis' where [simp]: "basis = b # basis'"
|
|
4516 |
by (cases basis) (auto simp: expands_to.simps)
|
|
4517 |
from assms have b_pos: "eventually (\<lambda>x. b x > 0) at_top"
|
|
4518 |
by (simp add: filterlim_at_top_dense basis_wf_Cons)
|
|
4519 |
from assms have "f \<sim>[at_top] eval_monom (dominant_term (MS (MSLCons (C, 0) xs) g)) basis"
|
|
4520 |
by (intro dominant_term_expands_to) simp_all
|
|
4521 |
also have "\<dots> = (\<lambda>x. eval_monom (dominant_term C) basis' x * b x powr 0)"
|
|
4522 |
by (simp_all add: dominant_term_ms_aux_def case_prod_unfold eval_monom_Cons)
|
|
4523 |
also have "eventually (\<lambda>x. \<dots> x = eval_monom (dominant_term C) basis' x) at_top"
|
|
4524 |
using b_pos by eventually_elim simp
|
|
4525 |
also from assms have "eval_monom (dominant_term C) basis' \<sim>[at_top] eval C"
|
|
4526 |
by (intro asymp_equiv_symI [OF dominant_term_expands_to]
|
|
4527 |
expands_to_hd''[of f C 0 xs g b basis']) (auto simp: trimmed_ms_aux_MSLCons basis_wf_Cons)
|
|
4528 |
finally show ?thesis .
|
|
4529 |
qed
|
|
4530 |
|
|
4531 |
|
|
4532 |
lemma compare_expansions_LT':
|
|
4533 |
assumes "compare_expansions (MS ys h) G \<equiv> (LT, c, c')" "trimmed G"
|
|
4534 |
"(f expands_to (MS (MSLCons (MS ys h, e) xs) f')) (b # basis)" "(g expands_to G) basis"
|
|
4535 |
"e = 0" "basis_wf (b # basis)"
|
|
4536 |
shows "h \<in> o(g)"
|
|
4537 |
proof -
|
|
4538 |
from assms show ?thesis
|
|
4539 |
by (intro compare_expansions_LT[OF _ assms(2) _ assms(4)])
|
|
4540 |
(auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons expands_to.simps
|
|
4541 |
elim!: is_expansion_aux_MSLCons)
|
|
4542 |
qed
|
|
4543 |
|
|
4544 |
lemma compare_expansions_GT':
|
|
4545 |
assumes "compare_expansions C G \<equiv> (GT, c, c')"
|
|
4546 |
"trimmed (MS (MSLCons (C, e) xs) f')"
|
|
4547 |
"(f expands_to (MS (MSLCons (C, e) xs) f')) (b # basis)" "(g expands_to G) basis"
|
|
4548 |
"e = 0" "basis_wf (b # basis)"
|
|
4549 |
shows "g \<in> o(f)"
|
|
4550 |
proof -
|
|
4551 |
from assms have "g \<in> o(eval C)"
|
|
4552 |
by (intro compare_expansions_GT[of C G c c' _ basis])
|
|
4553 |
(auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons)
|
|
4554 |
also from assms have "f \<in> \<Theta>(eval C)"
|
|
4555 |
by (intro asymp_equiv_imp_bigtheta expands_to_MSLCons_0_asymp_equiv_hd[of f C xs f' "b#basis"])
|
|
4556 |
auto
|
|
4557 |
finally show ?thesis .
|
|
4558 |
qed
|
|
4559 |
|
|
4560 |
lemma compare_expansions_EQ':
|
|
4561 |
assumes "compare_expansions C G = (EQ, c, c')"
|
|
4562 |
"trimmed (MS (MSLCons (C, e) xs) f')" "trimmed G"
|
|
4563 |
"(f expands_to (MS (MSLCons (C, e) xs) f')) (b # basis)" "(g expands_to G) basis"
|
|
4564 |
"e = 0" "basis_wf (b # basis)"
|
|
4565 |
shows "f \<sim>[at_top] (\<lambda>x. c / c' * g x)"
|
|
4566 |
proof -
|
|
4567 |
from assms have [simp]: "c' \<noteq> 0"
|
|
4568 |
by (auto simp: compare_expansions_def trimmed_imp_dominant_term_nz split: cmp_result.splits)
|
|
4569 |
from assms have "f \<sim>[at_top] eval C"
|
|
4570 |
by (intro expands_to_MSLCons_0_asymp_equiv_hd[of f C xs f' "b#basis"]) auto
|
|
4571 |
also from assms(2,4,7) have *: "(\<lambda>x. c' * eval C x) \<sim>[at_top] (\<lambda>x. c * g x)"
|
|
4572 |
by (intro compare_expansions_EQ[OF assms(1) _ assms(3) _ assms(5)])
|
|
4573 |
(auto intro: expands_to_hd'' simp: trimmed_ms_aux_MSLCons basis_wf_Cons)
|
|
4574 |
have "(\<lambda>x. inverse c' * (c' * eval C x)) \<sim>[at_top] (\<lambda>x. inverse c' * (c * g x))"
|
|
4575 |
by (rule asymp_equiv_mult) (insert *, simp_all)
|
|
4576 |
hence "eval C \<sim>[at_top] (\<lambda>x. c / c' * g x)" by (simp add: divide_simps)
|
|
4577 |
finally show ?thesis .
|
|
4578 |
qed
|
|
4579 |
|
|
4580 |
lemma expands_to_insert_ln:
|
|
4581 |
assumes "basis_wf [b]"
|
|
4582 |
shows "((\<lambda>x. ln (b x)) expands_to MS (MSLCons (1::real,1) MSLNil) (\<lambda>x. ln (b x))) [\<lambda>x. ln (b x)]"
|
|
4583 |
proof -
|
|
4584 |
from assms have b_limit: "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
4585 |
hence b: "eventually (\<lambda>x. b x > 1) at_top" by (simp add: filterlim_at_top_dense)
|
|
4586 |
have "(\<lambda>x::real. x) \<in> \<Theta>(\<lambda>x. x powr 1)"
|
|
4587 |
by (intro bigthetaI_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto
|
|
4588 |
also have "(\<lambda>x::real. x powr 1) \<in> o(\<lambda>a. a powr e)" if "e > 1" for e
|
|
4589 |
by (subst powr_smallo_iff) (auto simp: that filterlim_ident)
|
|
4590 |
finally show ?thesis using b assms
|
|
4591 |
by (auto intro!: is_expansion_aux.intros landau_o.small.compose[of _ at_top _ "\<lambda>x. ln (b x)"]
|
|
4592 |
filterlim_compose[OF ln_at_top b_limit] is_expansion_real.intros
|
|
4593 |
elim!: eventually_mono simp: expands_to.simps)
|
|
4594 |
qed
|
|
4595 |
|
|
4596 |
lemma trimmed_pos_insert_ln:
|
|
4597 |
assumes "basis_wf [b]"
|
|
4598 |
shows "trimmed_pos (MS (MSLCons (1::real, 1) MSLNil) (\<lambda>x. ln (b x)))"
|
|
4599 |
by (simp_all add: trimmed_ms_aux_def)
|
|
4600 |
|
|
4601 |
|
|
4602 |
primrec compare_list_0 where
|
|
4603 |
"compare_list_0 [] = EQ"
|
|
4604 |
| "compare_list_0 (c # cs) = CMP_BRANCH (COMPARE c 0) LT (compare_list_0 cs) GT"
|
|
4605 |
|
|
4606 |
|
|
4607 |
lemma compare_reals_diff_sgnD:
|
|
4608 |
"a - (b :: real) < 0 \<Longrightarrow> a < b" "a - b = 0 \<Longrightarrow> a = b" "a - b > 0 \<Longrightarrow> a > b"
|
|
4609 |
by simp_all
|
|
4610 |
|
|
4611 |
lemma expands_to_real_imp_filterlim:
|
|
4612 |
assumes "(f expands_to (c :: real)) basis"
|
|
4613 |
shows "(f \<longlongrightarrow> c) at_top"
|
|
4614 |
using assms by (auto elim!: expands_to.cases simp: eq_commute[of c] Lim_eventually)
|
|
4615 |
|
|
4616 |
lemma expands_to_MSLNil_imp_filterlim:
|
|
4617 |
assumes "(f expands_to MS MSLNil f') basis"
|
|
4618 |
shows "(f \<longlongrightarrow> 0) at_top"
|
|
4619 |
proof -
|
|
4620 |
from assms have "eventually (\<lambda>x. f' x = 0) at_top" "eventually (\<lambda>x. f' x = f x) at_top"
|
|
4621 |
by (auto elim!: expands_to.cases is_expansion_aux.cases[of MSLNil])
|
|
4622 |
hence "eventually (\<lambda>x. f x = 0) at_top" by eventually_elim auto
|
|
4623 |
thus ?thesis by (simp add: Lim_eventually)
|
|
4624 |
qed
|
|
4625 |
|
|
4626 |
lemma expands_to_neg_exponent_imp_filterlim:
|
|
4627 |
assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis" "basis_wf basis" "e < 0"
|
|
4628 |
shows "(f \<longlongrightarrow> 0) at_top"
|
|
4629 |
proof -
|
|
4630 |
let ?F = "MS (MSLCons (C, e) xs) f'"
|
|
4631 |
let ?es = "snd (dominant_term ?F)"
|
|
4632 |
from assms have exp: "is_expansion ?F basis"
|
|
4633 |
by (simp add: expands_to.simps)
|
|
4634 |
from assms have "f \<in> \<Theta>(f')" by (intro bigthetaI_cong) (simp add: expands_to.simps eq_commute)
|
|
4635 |
also from dominant_term_bigo[OF assms(2) exp]
|
|
4636 |
have "f' \<in> O(eval_monom (1, ?es) basis)" by simp
|
|
4637 |
also have "(eval_monom (1, ?es) basis) \<in> o(\<lambda>x. hd basis x powr 0)" using exp assms
|
|
4638 |
by (intro eval_monom_smallo)
|
|
4639 |
(auto simp: is_expansion_aux_basis_nonempty dominant_term_ms_aux_def
|
|
4640 |
case_prod_unfold length_dominant_term is_expansion_aux_expansion_level)
|
|
4641 |
also have "(\<lambda>x. hd basis x powr 0) \<in> O(\<lambda>_. 1)"
|
|
4642 |
by (intro landau_o.bigI[of 1]) auto
|
|
4643 |
finally show ?thesis by (auto dest!: smalloD_tendsto)
|
|
4644 |
qed
|
|
4645 |
|
|
4646 |
lemma expands_to_pos_exponent_imp_filterlim:
|
|
4647 |
assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis" "trimmed (MS (MSLCons (C, e) xs) f')"
|
|
4648 |
"basis_wf basis" "e > 0"
|
|
4649 |
shows "filterlim f at_infinity at_top"
|
|
4650 |
proof -
|
|
4651 |
let ?F = "MS (MSLCons (C, e) xs) f'"
|
|
4652 |
let ?es = "snd (dominant_term ?F)"
|
|
4653 |
from assms have exp: "is_expansion ?F basis"
|
|
4654 |
by (simp add: expands_to.simps)
|
|
4655 |
with assms have "filterlim (hd basis) at_top at_top"
|
|
4656 |
using is_expansion_aux_basis_nonempty[of "MSLCons (C, e) xs" f' basis]
|
|
4657 |
by (cases basis) (simp_all add: basis_wf_Cons)
|
|
4658 |
hence b_pos: "eventually (\<lambda>x. hd basis x > 0) at_top" using filterlim_at_top_dense by blast
|
|
4659 |
|
|
4660 |
from assms have "f \<in> \<Theta>(f')" by (intro bigthetaI_cong) (simp add: expands_to.simps eq_commute)
|
|
4661 |
also from dominant_term[OF assms(3) exp assms(2)]
|
|
4662 |
have "f' \<in> \<Theta>(eval_monom (dominant_term ?F) basis)" by (simp add: asymp_equiv_imp_bigtheta)
|
|
4663 |
also have "(eval_monom (dominant_term ?F) basis) \<in> \<Theta>(eval_monom (1, ?es) basis)"
|
|
4664 |
using assms by (simp add: eval_monom_def case_prod_unfold dominant_term_ms_aux_def
|
|
4665 |
trimmed_imp_dominant_term_nz trimmed_ms_aux_def)
|
|
4666 |
also have "eval_monom (1, ?es) basis \<in> \<omega>(\<lambda>x. hd basis x powr 0)" using exp assms
|
|
4667 |
by (intro eval_monom_smallomega)
|
|
4668 |
(auto simp: is_expansion_aux_basis_nonempty dominant_term_ms_aux_def
|
|
4669 |
case_prod_unfold length_dominant_term is_expansion_aux_expansion_level)
|
|
4670 |
also have "(\<lambda>x. hd basis x powr 0) \<in> \<Theta>(\<lambda>_. 1)"
|
|
4671 |
by (intro bigthetaI_cong eventually_mono[OF b_pos]) auto
|
|
4672 |
finally show ?thesis by (auto dest!: smallomegaD_filterlim_at_infinity)
|
|
4673 |
qed
|
|
4674 |
|
|
4675 |
lemma expands_to_zero_exponent_imp_filterlim:
|
|
4676 |
assumes "(f expands_to MS (MSLCons (C, e) xs) f') basis"
|
|
4677 |
"basis_wf basis" "e = 0"
|
|
4678 |
shows "filterlim (eval C) at_infinity at_top \<Longrightarrow> filterlim f at_infinity at_top"
|
|
4679 |
and "filterlim (eval C) (nhds L) at_top \<Longrightarrow> filterlim f (nhds L) at_top"
|
|
4680 |
proof -
|
|
4681 |
from assms obtain b basis' where *:
|
|
4682 |
"basis = b # basis'" "is_expansion C basis'" "ms_exp_gt 0 (ms_aux_hd_exp xs)"
|
|
4683 |
"is_expansion_aux xs (\<lambda>x. f' x - eval C x * b x powr 0) basis"
|
|
4684 |
by (auto elim!: expands_to.cases is_expansion_aux_MSLCons)
|
|
4685 |
|
|
4686 |
from *(1) assms have "filterlim b at_top at_top" by (simp add: basis_wf_Cons)
|
|
4687 |
hence b_pos: "eventually (\<lambda>x. b x > 0) at_top" using filterlim_at_top_dense by blast
|
|
4688 |
|
|
4689 |
from assms(1) have "eventually (\<lambda>x. f' x = f x) at_top" by (simp add: expands_to.simps)
|
|
4690 |
hence "eventually (\<lambda>x. f' x - eval C x * b x powr 0 = f x - eval C x) at_top"
|
|
4691 |
using b_pos by eventually_elim auto
|
|
4692 |
from *(4) and this have "is_expansion_aux xs (\<lambda>x. f x - eval C x) basis"
|
|
4693 |
by (rule is_expansion_aux_cong)
|
|
4694 |
hence "(\<lambda>x. f x - eval C x) \<in> o(\<lambda>x. hd basis x powr 0)"
|
|
4695 |
by (rule is_expansion_aux_imp_smallo) fact
|
|
4696 |
also have "(\<lambda>x. hd basis x powr 0) \<in> \<Theta>(\<lambda>_. 1)"
|
|
4697 |
by (intro bigthetaI_cong eventually_mono[OF b_pos]) (auto simp: *(1))
|
|
4698 |
finally have lim: "filterlim (\<lambda>x. f x - eval C x) (nhds 0) at_top"
|
|
4699 |
by (auto dest: smalloD_tendsto)
|
|
4700 |
|
|
4701 |
show "filterlim f at_infinity at_top" if "filterlim (eval C) at_infinity at_top"
|
|
4702 |
using tendsto_add_filterlim_at_infinity[OF lim that] by simp
|
|
4703 |
show "filterlim f (nhds L) at_top" if "filterlim (eval C) (nhds L) at_top"
|
|
4704 |
using tendsto_add[OF lim that] by simp
|
|
4705 |
qed
|
|
4706 |
|
|
4707 |
lemma expands_to_lift_function:
|
|
4708 |
assumes "eventually (\<lambda>x. f x - eval c x = 0) at_top"
|
|
4709 |
"((\<lambda>x. g (eval (c :: 'a :: multiseries) x)) expands_to G) bs'"
|
|
4710 |
shows "((\<lambda>x. g (f x)) expands_to G) bs'"
|
|
4711 |
by (rule expands_to_cong[OF assms(2)]) (insert assms, auto elim: eventually_mono)
|
|
4712 |
|
|
4713 |
lemma drop_zero_ms':
|
|
4714 |
fixes c f
|
|
4715 |
assumes "c = (0::real)" "(f expands_to MS (MSLCons (c, e) xs) g) basis"
|
|
4716 |
shows "(f expands_to MS xs g) basis"
|
|
4717 |
using assms drop_zero_ms[of c e xs g basis] by (simp add: expands_to.simps)
|
|
4718 |
|
|
4719 |
lemma trimmed_realI: "c \<noteq> (0::real) \<Longrightarrow> trimmed c"
|
|
4720 |
by simp
|
|
4721 |
|
|
4722 |
lemma trimmed_pos_realI: "c > (0::real) \<Longrightarrow> trimmed_pos c"
|
|
4723 |
by simp
|
|
4724 |
|
|
4725 |
lemma trimmed_neg_realI: "c < (0::real) \<Longrightarrow> trimmed_neg c"
|
|
4726 |
by (simp add: trimmed_neg_def)
|
|
4727 |
|
|
4728 |
lemma trimmed_pos_hd_coeff: "trimmed_pos (MS (MSLCons (C, e) xs) f) \<Longrightarrow> trimmed_pos C"
|
|
4729 |
by simp
|
|
4730 |
|
|
4731 |
lemma lift_trimmed: "trimmed C \<Longrightarrow> trimmed (MS (MSLCons (C, e) xs) f)"
|
|
4732 |
by (auto simp: trimmed_ms_aux_def)
|
|
4733 |
|
|
4734 |
lemma lift_trimmed_pos: "trimmed_pos C \<Longrightarrow> trimmed_pos (MS (MSLCons (C, e) xs) f)"
|
|
4735 |
by simp
|
|
4736 |
|
|
4737 |
lemma lift_trimmed_neg: "trimmed_neg C \<Longrightarrow> trimmed_neg (MS (MSLCons (C, e) xs) f)"
|
|
4738 |
by (simp add: trimmed_neg_def fst_dominant_term_ms_aux_MSLCons trimmed_ms_aux_MSLCons)
|
|
4739 |
|
|
4740 |
lemma lift_trimmed_pos':
|
|
4741 |
"trimmed_pos C \<Longrightarrow> MS (MSLCons (C, e) xs) f \<equiv> MS (MSLCons (C, e) xs) f \<Longrightarrow>
|
|
4742 |
trimmed_pos (MS (MSLCons (C, e) xs) f)"
|
|
4743 |
by simp
|
|
4744 |
|
|
4745 |
lemma lift_trimmed_neg':
|
|
4746 |
"trimmed_neg C \<Longrightarrow> MS (MSLCons (C, e) xs) f \<equiv> MS (MSLCons (C, e) xs) f \<Longrightarrow>
|
|
4747 |
trimmed_neg (MS (MSLCons (C, e) xs) f)"
|
|
4748 |
by (simp add: lift_trimmed_neg)
|
|
4749 |
|
|
4750 |
lemma trimmed_eq_cong: "trimmed C \<Longrightarrow> C \<equiv> C' \<Longrightarrow> trimmed C'"
|
|
4751 |
and trimmed_pos_eq_cong: "trimmed_pos C \<Longrightarrow> C \<equiv> C' \<Longrightarrow> trimmed_pos C'"
|
|
4752 |
and trimmed_neg_eq_cong: "trimmed_neg C \<Longrightarrow> C \<equiv> C' \<Longrightarrow> trimmed_neg C'"
|
|
4753 |
by simp_all
|
|
4754 |
|
|
4755 |
lemma trimmed_hd: "trimmed (MS (MSLCons (C, e) xs) f) \<Longrightarrow> trimmed C"
|
|
4756 |
by (simp add: trimmed_ms_aux_MSLCons)
|
|
4757 |
|
|
4758 |
lemma trim_lift_eq:
|
|
4759 |
assumes "A \<equiv> MS (MSLCons (C, e) xs) f" "C \<equiv> D"
|
|
4760 |
shows "A \<equiv> MS (MSLCons (D, e) xs) f"
|
|
4761 |
using assms by simp
|
|
4762 |
|
|
4763 |
lemma basis_wf_manyI:
|
|
4764 |
"filterlim b' at_top at_top \<Longrightarrow> (\<lambda>x. ln (b x)) \<in> o(\<lambda>x. ln (b' x)) \<Longrightarrow>
|
|
4765 |
basis_wf (b # basis) \<Longrightarrow> basis_wf (b' # b # basis)"
|
|
4766 |
by (simp_all add: basis_wf_many)
|
|
4767 |
|
|
4768 |
lemma ln_smallo_ln_exp: "(\<lambda>x. ln (b x)) \<in> o(f) \<Longrightarrow> (\<lambda>x. ln (b x)) \<in> o(\<lambda>x. ln (exp (f x :: real)))"
|
|
4769 |
by simp
|
|
4770 |
|
|
4771 |
|
|
4772 |
subsection \<open>Reification and other technical details\<close>
|
|
4773 |
|
|
4774 |
text \<open>
|
|
4775 |
The following is used by the automation in order to avoid writing terms like $x^2$ or $x^{-2}$
|
|
4776 |
as @{term "\<lambda>x::real. x powr 2"} etc.\ but as the more agreeable @{term "\<lambda>x::real. x ^ 2"} or
|
|
4777 |
@{term "\<lambda>x::real. inverse (x ^ 2)"}.
|
|
4778 |
\<close>
|
|
4779 |
|
|
4780 |
lemma intyness_0: "0 \<equiv> real 0"
|
|
4781 |
and intyness_1: "1 \<equiv> real 1"
|
|
4782 |
and intyness_numeral: "num \<equiv> num \<Longrightarrow> numeral num \<equiv> real (numeral num)"
|
|
4783 |
and intyness_uminus: "x \<equiv> real n \<Longrightarrow> -x \<equiv> -real n"
|
|
4784 |
and intyness_of_nat: "n \<equiv> n \<Longrightarrow> real n \<equiv> real n"
|
|
4785 |
by simp_all
|
|
4786 |
|
|
4787 |
lemma intyness_simps:
|
|
4788 |
"real a + real b = real (a + b)"
|
|
4789 |
"real a * real b = real (a * b)"
|
|
4790 |
"real a ^ n = real (a ^ n)"
|
|
4791 |
"1 = real 1" "0 = real 0" "numeral num = real (numeral num)"
|
|
4792 |
by simp_all
|
|
4793 |
|
|
4794 |
lemma odd_Numeral1: "odd (Numeral1)"
|
|
4795 |
by simp
|
|
4796 |
|
|
4797 |
lemma even_addI:
|
|
4798 |
"even a \<Longrightarrow> even b \<Longrightarrow> even (a + b)"
|
|
4799 |
"odd a \<Longrightarrow> odd b \<Longrightarrow> even (a + b)"
|
|
4800 |
by auto
|
|
4801 |
|
|
4802 |
lemma odd_addI:
|
|
4803 |
"even a \<Longrightarrow> odd b \<Longrightarrow> odd (a + b)"
|
|
4804 |
"odd a \<Longrightarrow> even b \<Longrightarrow> odd (a + b)"
|
|
4805 |
by auto
|
|
4806 |
|
|
4807 |
lemma even_diffI:
|
|
4808 |
"even a \<Longrightarrow> even b \<Longrightarrow> even (a - b :: 'a :: ring_parity)"
|
|
4809 |
"odd a \<Longrightarrow> odd b \<Longrightarrow> even (a - b)"
|
|
4810 |
by auto
|
|
4811 |
|
|
4812 |
lemma odd_diffI:
|
|
4813 |
"even a \<Longrightarrow> odd b \<Longrightarrow> odd (a - b :: 'a :: ring_parity)"
|
|
4814 |
"odd a \<Longrightarrow> even b \<Longrightarrow> odd (a - b)"
|
|
4815 |
by auto
|
|
4816 |
|
|
4817 |
lemma even_multI: "even a \<Longrightarrow> even (a * b)" "even b \<Longrightarrow> even (a * b)"
|
|
4818 |
by auto
|
|
4819 |
|
|
4820 |
lemma odd_multI: "odd a \<Longrightarrow> odd b \<Longrightarrow> odd (a * b)"
|
|
4821 |
by auto
|
|
4822 |
|
|
4823 |
lemma even_uminusI: "even a \<Longrightarrow> even (-a :: 'a :: ring_parity)"
|
|
4824 |
and odd_uminusI: "odd a \<Longrightarrow> odd (-a :: 'a :: ring_parity)"
|
|
4825 |
by auto
|
|
4826 |
|
|
4827 |
lemma odd_powerI: "odd a \<Longrightarrow> odd (a ^ n)"
|
|
4828 |
by auto
|
|
4829 |
|
|
4830 |
|
|
4831 |
text \<open>
|
|
4832 |
The following theorem collection is used to pre-process functions for multiseries expansion.
|
|
4833 |
\<close>
|
|
4834 |
named_theorems real_asymp_reify_simps
|
|
4835 |
|
|
4836 |
lemmas [real_asymp_reify_simps] =
|
|
4837 |
sinh_field_def cosh_field_def tanh_real_altdef arsinh_def arcosh_def artanh_def
|
|
4838 |
|
|
4839 |
|
|
4840 |
text \<open>
|
|
4841 |
This is needed in order to handle things like @{term "\<lambda>n. f n ^ n"}.
|
|
4842 |
\<close>
|
|
4843 |
definition powr_nat :: "real \<Rightarrow> real \<Rightarrow> real" where
|
|
4844 |
"powr_nat x y =
|
|
4845 |
(if y = 0 then 1
|
|
4846 |
else if x < 0 then cos (pi * y) * (-x) powr y else x powr y)"
|
|
4847 |
|
|
4848 |
lemma powr_nat_of_nat: "powr_nat x (of_nat n) = x ^ n"
|
|
4849 |
by (cases "x > 0") (auto simp: powr_nat_def powr_realpow not_less power_mult_distrib [symmetric])
|
|
4850 |
|
|
4851 |
lemma powr_nat_conv_powr: "x > 0 \<Longrightarrow> powr_nat x y = x powr y"
|
|
4852 |
by (simp_all add: powr_nat_def)
|
|
4853 |
|
|
4854 |
lemma reify_power: "x ^ n \<equiv> powr_nat x (of_nat n)"
|
|
4855 |
by (simp add: powr_nat_of_nat)
|
|
4856 |
|
|
4857 |
|
|
4858 |
lemma sqrt_conv_root [real_asymp_reify_simps]: "sqrt x = root 2 x"
|
|
4859 |
by (simp add: sqrt_def)
|
|
4860 |
|
|
4861 |
lemma tan_conv_sin_cos [real_asymp_reify_simps]: "tan x = sin x / cos x"
|
|
4862 |
by (simp add: tan_def)
|
|
4863 |
|
|
4864 |
definition rfloor :: "real \<Rightarrow> real" where "rfloor x = real_of_int (floor x)"
|
|
4865 |
definition rceil :: "real \<Rightarrow> real" where "rceil x = real_of_int (ceiling x)"
|
|
4866 |
definition rnatmod :: "real \<Rightarrow> real \<Rightarrow> real"
|
|
4867 |
where "rnatmod x y = real (nat \<lfloor>x\<rfloor> mod nat \<lfloor>y\<rfloor>)"
|
|
4868 |
definition rintmod :: "real \<Rightarrow> real \<Rightarrow> real"
|
|
4869 |
where "rintmod x y = real_of_int (\<lfloor>x\<rfloor> mod \<lfloor>y\<rfloor>)"
|
|
4870 |
|
|
4871 |
lemmas [real_asymp_reify_simps] =
|
|
4872 |
ln_exp log_def rfloor_def [symmetric] rceil_def [symmetric]
|
|
4873 |
|
|
4874 |
lemma expands_to_powr_nat_0_0:
|
|
4875 |
assumes "eventually (\<lambda>x. f x = 0) at_top" "eventually (\<lambda>x. g x = 0) at_top"
|
|
4876 |
"basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)"
|
|
4877 |
shows "((\<lambda>x. powr_nat (f x) (g x)) expands_to (const_expansion 1 :: 'a)) basis"
|
|
4878 |
proof (rule expands_to_cong [OF expands_to_const])
|
|
4879 |
from assms(1,2) show "eventually (\<lambda>x. 1 = powr_nat (f x) (g x)) at_top"
|
|
4880 |
by eventually_elim (simp add: powr_nat_def)
|
|
4881 |
qed fact+
|
|
4882 |
|
|
4883 |
lemma expands_to_powr_nat_0:
|
|
4884 |
assumes "eventually (\<lambda>x. f x = 0) at_top" "(g expands_to G) basis" "trimmed G"
|
|
4885 |
"basis_wf basis" "length basis = expansion_level TYPE('a :: multiseries)"
|
|
4886 |
shows "((\<lambda>x. powr_nat (f x) (g x)) expands_to (zero_expansion :: 'a)) basis"
|
|
4887 |
proof (rule expands_to_cong [OF expands_to_zero])
|
|
4888 |
from assms have "eventually (\<lambda>x. g x \<noteq> 0) at_top"
|
|
4889 |
using expands_to_imp_eventually_nz[of basis g G] by auto
|
|
4890 |
with assms(1) show "eventually (\<lambda>x. 0 = powr_nat (f x) (g x)) at_top"
|
|
4891 |
by eventually_elim (simp add: powr_nat_def)
|
|
4892 |
qed (insert assms, auto elim!: eventually_mono simp: powr_nat_def)
|
|
4893 |
|
|
4894 |
lemma expands_to_powr_nat:
|
|
4895 |
assumes "trimmed_pos F" "basis_wf basis'" "(f expands_to F) basis'"
|
|
4896 |
assumes "((\<lambda>x. exp (ln (f x) * g x)) expands_to E) basis"
|
|
4897 |
shows "((\<lambda>x. powr_nat (f x) (g x)) expands_to E) basis"
|
|
4898 |
using assms(4)
|
|
4899 |
proof (rule expands_to_cong)
|
|
4900 |
from eval_pos_if_dominant_term_pos[of basis' F] assms(1-4)
|
|
4901 |
show "eventually (\<lambda>x. exp (ln (f x) * g x) = powr_nat (f x) (g x)) at_top"
|
|
4902 |
by (auto simp: expands_to.simps trimmed_pos_def powr_def powr_nat_def elim: eventually_elim2)
|
|
4903 |
qed
|
|
4904 |
|
|
4905 |
|
|
4906 |
subsection \<open>Some technical lemmas\<close>
|
|
4907 |
|
|
4908 |
lemma landau_meta_eq_cong: "f \<in> L(g) \<Longrightarrow> f' \<equiv> f \<Longrightarrow> g' \<equiv> g \<Longrightarrow> f' \<in> L(g')"
|
|
4909 |
and asymp_equiv_meta_eq_cong: "f \<sim>[at_top] g \<Longrightarrow> f' \<equiv> f \<Longrightarrow> g' \<equiv> g \<Longrightarrow> f' \<sim>[at_top] g'"
|
|
4910 |
by simp_all
|
|
4911 |
|
|
4912 |
lemma filterlim_mono': "filterlim f F G \<Longrightarrow> F \<le> F' \<Longrightarrow> filterlim f F' G"
|
|
4913 |
by (erule (1) filterlim_mono) simp_all
|
|
4914 |
|
|
4915 |
lemma at_within_le_nhds: "at x within A \<le> nhds x"
|
|
4916 |
by (simp add: at_within_def)
|
|
4917 |
|
|
4918 |
lemma at_within_le_at: "at x within A \<le> at x"
|
|
4919 |
by (rule at_le) simp_all
|
|
4920 |
|
|
4921 |
lemma at_right_to_top': "at_right c = filtermap (\<lambda>x::real. c + inverse x) at_top"
|
|
4922 |
by (subst at_right_to_0, subst at_right_to_top) (simp add: filtermap_filtermap add_ac)
|
|
4923 |
|
|
4924 |
lemma at_left_to_top': "at_left c = filtermap (\<lambda>x::real. c - inverse x) at_top"
|
|
4925 |
by (subst at_left_minus, subst at_right_to_0, subst at_right_to_top)
|
|
4926 |
(simp add: filtermap_filtermap add_ac)
|
|
4927 |
|
|
4928 |
lemma at_left_to_top: "at_left 0 = filtermap (\<lambda>x::real. - inverse x) at_top"
|
|
4929 |
by (simp add: at_left_to_top')
|
|
4930 |
|
|
4931 |
lemma filterlim_conv_filtermap:
|
|
4932 |
"G = filtermap g G' \<Longrightarrow>
|
|
4933 |
PROP (Trueprop (filterlim f F G)) \<equiv> PROP (Trueprop (filterlim (\<lambda>x. f (g x)) F G'))"
|
|
4934 |
by (simp add: filterlim_filtermap)
|
|
4935 |
|
|
4936 |
lemma eventually_conv_filtermap:
|
|
4937 |
"G = filtermap g G' \<Longrightarrow>
|
|
4938 |
PROP (Trueprop (eventually P G)) \<equiv> PROP (Trueprop (eventually (\<lambda>x. P (g x)) G'))"
|
|
4939 |
by (simp add: eventually_filtermap)
|
|
4940 |
|
|
4941 |
lemma eventually_lt_imp_eventually_le:
|
|
4942 |
"eventually (\<lambda>x. f x < (g x :: real)) F \<Longrightarrow> eventually (\<lambda>x. f x \<le> g x) F"
|
|
4943 |
by (erule eventually_mono) simp
|
|
4944 |
|
|
4945 |
lemma smallo_imp_smallomega: "f \<in> o[F](g) \<Longrightarrow> g \<in> \<omega>[F](f)"
|
|
4946 |
by (simp add: smallomega_iff_smallo)
|
|
4947 |
|
|
4948 |
lemma bigo_imp_bigomega: "f \<in> O[F](g) \<Longrightarrow> g \<in> \<Omega>[F](f)"
|
|
4949 |
by (simp add: bigomega_iff_bigo)
|
|
4950 |
|
|
4951 |
context
|
|
4952 |
fixes L L' :: "real filter \<Rightarrow> (real \<Rightarrow> _) \<Rightarrow> _" and Lr :: "real filter \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> _"
|
|
4953 |
assumes LS: "landau_symbol L L' Lr"
|
|
4954 |
begin
|
|
4955 |
|
|
4956 |
interpretation landau_symbol L L' Lr by (fact LS)
|
|
4957 |
|
|
4958 |
lemma landau_symbol_at_top_imp_at_bot:
|
|
4959 |
"(\<lambda>x. f (-x)) \<in> L' at_top (\<lambda>x. g (-x)) \<Longrightarrow> f \<in> L at_bot g"
|
|
4960 |
by (simp add: in_filtermap_iff at_bot_mirror)
|
|
4961 |
|
|
4962 |
lemma landau_symbol_at_top_imp_at_right_0:
|
|
4963 |
"(\<lambda>x. f (inverse x)) \<in> L' at_top (\<lambda>x. g (inverse x)) \<Longrightarrow> f \<in> L (at_right 0) g"
|
|
4964 |
by (simp add: in_filtermap_iff at_right_to_top')
|
|
4965 |
|
|
4966 |
lemma landau_symbol_at_top_imp_at_left_0:
|
|
4967 |
"(\<lambda>x. f ( -inverse x)) \<in> L' at_top (\<lambda>x. g (-inverse x)) \<Longrightarrow> f \<in> L (at_left 0) g"
|
|
4968 |
by (simp add: in_filtermap_iff at_left_to_top')
|
|
4969 |
|
|
4970 |
lemma landau_symbol_at_top_imp_at_right:
|
|
4971 |
"(\<lambda>x. f (a + inverse x)) \<in> L' at_top (\<lambda>x. g (a + inverse x)) \<Longrightarrow> f \<in> L (at_right a) g"
|
|
4972 |
by (simp add: in_filtermap_iff at_right_to_top')
|
|
4973 |
|
|
4974 |
lemma landau_symbol_at_top_imp_at_left:
|
|
4975 |
"(\<lambda>x. f (a - inverse x)) \<in> L' at_top (\<lambda>x. g (a - inverse x)) \<Longrightarrow> f \<in> L (at_left a) g"
|
|
4976 |
by (simp add: in_filtermap_iff at_left_to_top')
|
|
4977 |
|
|
4978 |
lemma landau_symbol_at_top_imp_at:
|
|
4979 |
"(\<lambda>x. f (a - inverse x)) \<in> L' at_top (\<lambda>x. g (a - inverse x)) \<Longrightarrow>
|
|
4980 |
(\<lambda>x. f (a + inverse x)) \<in> L' at_top (\<lambda>x. g (a + inverse x)) \<Longrightarrow>
|
|
4981 |
f \<in> L (at a) g"
|
|
4982 |
by (simp add: sup at_eq_sup_left_right
|
|
4983 |
landau_symbol_at_top_imp_at_left landau_symbol_at_top_imp_at_right)
|
|
4984 |
|
|
4985 |
lemma landau_symbol_at_top_imp_at_0:
|
|
4986 |
"(\<lambda>x. f (-inverse x)) \<in> L' at_top (\<lambda>x. g (-inverse x)) \<Longrightarrow>
|
|
4987 |
(\<lambda>x. f (inverse x)) \<in> L' at_top (\<lambda>x. g (inverse x)) \<Longrightarrow>
|
|
4988 |
f \<in> L (at 0) g"
|
|
4989 |
by (rule landau_symbol_at_top_imp_at) simp_all
|
|
4990 |
|
|
4991 |
end
|
|
4992 |
|
|
4993 |
|
|
4994 |
context
|
|
4995 |
fixes f g :: "real \<Rightarrow> real"
|
|
4996 |
begin
|
|
4997 |
|
|
4998 |
lemma asymp_equiv_at_top_imp_at_bot:
|
|
4999 |
"(\<lambda>x. f (- x)) \<sim>[at_top] (\<lambda>x. g (-x)) \<Longrightarrow> f \<sim>[at_bot] g"
|
|
5000 |
by (simp add: asymp_equiv_def filterlim_at_bot_mirror)
|
|
5001 |
|
|
5002 |
lemma asymp_equiv_at_top_imp_at_right_0:
|
|
5003 |
"(\<lambda>x. f (inverse x)) \<sim>[at_top] (\<lambda>x. g (inverse x)) \<Longrightarrow> f \<sim>[at_right 0] g"
|
|
5004 |
by (simp add: at_right_to_top asymp_equiv_filtermap_iff)
|
|
5005 |
|
|
5006 |
lemma asymp_equiv_at_top_imp_at_left_0:
|
|
5007 |
"(\<lambda>x. f (-inverse x)) \<sim>[at_top] (\<lambda>x. g (-inverse x)) \<Longrightarrow> f \<sim>[at_left 0] g"
|
|
5008 |
by (simp add: at_left_to_top asymp_equiv_filtermap_iff)
|
|
5009 |
|
|
5010 |
lemma asymp_equiv_at_top_imp_at_right:
|
|
5011 |
"(\<lambda>x. f (a + inverse x)) \<sim>[at_top] (\<lambda>x. g (a + inverse x)) \<Longrightarrow> f \<sim>[at_right a] g"
|
|
5012 |
by (simp add: at_right_to_top' asymp_equiv_filtermap_iff)
|
|
5013 |
|
|
5014 |
lemma asymp_equiv_at_top_imp_at_left:
|
|
5015 |
"(\<lambda>x. f (a - inverse x)) \<sim>[at_top] (\<lambda>x. g (a - inverse x)) \<Longrightarrow> f \<sim>[at_left a] g"
|
|
5016 |
by (simp add: at_left_to_top' asymp_equiv_filtermap_iff)
|
|
5017 |
|
|
5018 |
lemma asymp_equiv_at_top_imp_at:
|
|
5019 |
"(\<lambda>x. f (a - inverse x)) \<sim>[at_top] (\<lambda>x. g (a - inverse x)) \<Longrightarrow>
|
|
5020 |
(\<lambda>x. f (a + inverse x)) \<sim>[at_top] (\<lambda>x. g (a + inverse x)) \<Longrightarrow> f \<sim>[at a] g"
|
|
5021 |
unfolding asymp_equiv_def
|
|
5022 |
using asymp_equiv_at_top_imp_at_left[of a] asymp_equiv_at_top_imp_at_right[of a]
|
|
5023 |
by (intro filterlim_split_at) (auto simp: asymp_equiv_def)
|
|
5024 |
|
|
5025 |
lemma asymp_equiv_at_top_imp_at_0:
|
|
5026 |
"(\<lambda>x. f (-inverse x)) \<sim>[at_top] (\<lambda>x. g (-inverse x)) \<Longrightarrow>
|
|
5027 |
(\<lambda>x. f (inverse x)) \<sim>[at_top] (\<lambda>x. g (inverse x)) \<Longrightarrow> f \<sim>[at 0] g"
|
|
5028 |
by (rule asymp_equiv_at_top_imp_at) auto
|
|
5029 |
|
|
5030 |
end
|
|
5031 |
|
|
5032 |
|
|
5033 |
lemmas eval_simps =
|
|
5034 |
eval_const eval_plus eval_minus eval_uminus eval_times eval_inverse eval_divide
|
|
5035 |
eval_map_ms eval_lift_ms eval_real_def eval_ms.simps
|
|
5036 |
|
|
5037 |
lemma real_eqI: "a - b = 0 \<Longrightarrow> a = (b::real)"
|
|
5038 |
by simp
|
|
5039 |
|
|
5040 |
lemma lift_ms_simps:
|
|
5041 |
"lift_ms (MS xs f) = MS (MSLCons (MS xs f, 0) MSLNil) f"
|
|
5042 |
"lift_ms (c :: real) = MS (MSLCons (c, 0) MSLNil) (\<lambda>_. c)"
|
|
5043 |
by (simp_all add: lift_ms_def)
|
|
5044 |
|
|
5045 |
lemmas landau_reduce_to_top =
|
|
5046 |
landau_symbols [THEN landau_symbol_at_top_imp_at_bot]
|
|
5047 |
landau_symbols [THEN landau_symbol_at_top_imp_at_left_0]
|
|
5048 |
landau_symbols [THEN landau_symbol_at_top_imp_at_right_0]
|
|
5049 |
landau_symbols [THEN landau_symbol_at_top_imp_at_left]
|
|
5050 |
landau_symbols [THEN landau_symbol_at_top_imp_at_right]
|
|
5051 |
asymp_equiv_at_top_imp_at_bot
|
|
5052 |
asymp_equiv_at_top_imp_at_left_0
|
|
5053 |
asymp_equiv_at_top_imp_at_right_0
|
|
5054 |
asymp_equiv_at_top_imp_at_left
|
|
5055 |
asymp_equiv_at_top_imp_at_right
|
|
5056 |
|
|
5057 |
lemmas landau_at_top_imp_at =
|
|
5058 |
landau_symbols [THEN landau_symbol_at_top_imp_at_0]
|
|
5059 |
landau_symbols [THEN landau_symbol_at_top_imp_at]
|
|
5060 |
asymp_equiv_at_top_imp_at_0
|
|
5061 |
asymp_equiv_at_top_imp_at
|
|
5062 |
|
|
5063 |
|
|
5064 |
lemma nhds_leI: "x = y \<Longrightarrow> nhds x \<le> nhds y"
|
|
5065 |
by simp
|
|
5066 |
|
|
5067 |
lemma of_nat_diff_real: "of_nat (a - b) = max (0::real) (of_nat a - of_nat b)"
|
|
5068 |
and of_nat_max_real: "of_nat (max a b) = max (of_nat a) (of_nat b :: real)"
|
|
5069 |
and of_nat_min_real: "of_nat (min a b) = min (of_nat a) (of_nat b :: real)"
|
|
5070 |
and of_int_max_real: "of_int (max c d) = max (of_int c) (of_int d :: real)"
|
|
5071 |
and of_int_min_real: "of_int (min c d) = min (of_int c) (of_int d :: real)"
|
|
5072 |
by simp_all
|
|
5073 |
|
|
5074 |
named_theorems real_asymp_nat_reify
|
|
5075 |
|
|
5076 |
lemmas [real_asymp_nat_reify] =
|
|
5077 |
of_nat_add of_nat_mult of_nat_diff_real of_nat_max_real of_nat_min_real of_nat_power
|
|
5078 |
of_nat_Suc of_nat_numeral
|
|
5079 |
|
|
5080 |
lemma of_nat_div_real [real_asymp_nat_reify]:
|
|
5081 |
"of_nat (a div b) = max 0 (rfloor (of_nat a / of_nat b))"
|
|
5082 |
by (simp add: rfloor_def floor_divide_of_nat_eq)
|
|
5083 |
|
|
5084 |
lemma of_nat_mod_real [real_asymp_nat_reify]: "of_nat (a mod b) = rnatmod (of_nat a) (of_nat b)"
|
|
5085 |
by (simp add: rnatmod_def)
|
|
5086 |
|
|
5087 |
lemma real_nat [real_asymp_nat_reify]: "real (nat a) = max 0 (of_int a)"
|
|
5088 |
by simp
|
|
5089 |
|
|
5090 |
|
|
5091 |
named_theorems real_asymp_int_reify
|
|
5092 |
|
|
5093 |
lemmas [real_asymp_int_reify] =
|
|
5094 |
of_int_add of_int_mult of_int_diff of_int_minus of_int_max_real of_int_min_real
|
|
5095 |
of_int_power of_int_of_nat_eq of_int_numeral
|
|
5096 |
|
|
5097 |
lemma of_int_div_real [real_asymp_int_reify]:
|
|
5098 |
"of_int (a div b) = rfloor (of_int a / of_int b)"
|
|
5099 |
by (simp add: rfloor_def floor_divide_of_int_eq)
|
|
5100 |
|
|
5101 |
named_theorems real_asymp_preproc
|
|
5102 |
|
|
5103 |
lemmas [real_asymp_preproc] =
|
|
5104 |
of_nat_add of_nat_mult of_nat_diff_real of_nat_max_real of_nat_min_real of_nat_power
|
|
5105 |
of_nat_Suc of_nat_numeral
|
|
5106 |
of_int_add of_int_mult of_int_diff of_int_minus of_int_max_real of_int_min_real
|
|
5107 |
of_int_power of_int_of_nat_eq of_int_numeral real_nat
|
|
5108 |
|
|
5109 |
lemma of_int_mod_real [real_asymp_int_reify]: "of_int (a mod b) = rintmod (of_int a) (of_int b)"
|
|
5110 |
by (simp add: rintmod_def)
|
|
5111 |
|
|
5112 |
|
|
5113 |
lemma filterlim_of_nat_at_top:
|
|
5114 |
"filterlim f F at_top \<Longrightarrow> filterlim (\<lambda>x. f (of_nat x :: real)) F at_top"
|
|
5115 |
by (erule filterlim_compose[OF _filterlim_real_sequentially])
|
|
5116 |
|
|
5117 |
lemma asymp_equiv_real_nat_transfer:
|
|
5118 |
"f \<sim>[at_top] g \<Longrightarrow> (\<lambda>x. f (of_nat x :: real)) \<sim>[at_top] (\<lambda>x. g (of_nat x))"
|
|
5119 |
unfolding asymp_equiv_def by (rule filterlim_of_nat_at_top)
|
|
5120 |
|
|
5121 |
lemma eventually_nat_real:
|
|
5122 |
assumes "eventually P (at_top :: real filter)"
|
|
5123 |
shows "eventually (\<lambda>x. P (real x)) (at_top :: nat filter)"
|
|
5124 |
using assms filterlim_real_sequentially
|
|
5125 |
unfolding filterlim_def le_filter_def eventually_filtermap by auto
|
|
5126 |
|
|
5127 |
lemmas real_asymp_nat_intros =
|
|
5128 |
filterlim_of_nat_at_top eventually_nat_real smallo_real_nat_transfer bigo_real_nat_transfer
|
|
5129 |
smallomega_real_nat_transfer bigomega_real_nat_transfer bigtheta_real_nat_transfer
|
|
5130 |
asymp_equiv_real_nat_transfer
|
|
5131 |
|
|
5132 |
|
|
5133 |
lemma filterlim_of_int_at_top:
|
|
5134 |
"filterlim f F at_top \<Longrightarrow> filterlim (\<lambda>x. f (of_int x :: real)) F at_top"
|
|
5135 |
by (erule filterlim_compose[OF _ filterlim_real_of_int_at_top])
|
|
5136 |
|
|
5137 |
(* TODO Move *)
|
|
5138 |
lemma filterlim_real_of_int_at_bot [tendsto_intros]:
|
|
5139 |
"filterlim real_of_int at_bot at_bot"
|
|
5140 |
unfolding filterlim_at_bot
|
|
5141 |
proof
|
|
5142 |
fix C :: real
|
|
5143 |
show "eventually (\<lambda>n. real_of_int n \<le> C) at_bot"
|
|
5144 |
using eventually_le_at_bot[of "\<lfloor>C\<rfloor>"] by eventually_elim linarith
|
|
5145 |
qed
|
|
5146 |
|
|
5147 |
lemma filterlim_of_int_at_bot:
|
|
5148 |
"filterlim f F at_bot \<Longrightarrow> filterlim (\<lambda>x. f (of_int x :: real)) F at_bot"
|
|
5149 |
by (erule filterlim_compose[OF _ filterlim_real_of_int_at_bot])
|
|
5150 |
find_theorems of_int at_bot
|
|
5151 |
|
|
5152 |
lemma asymp_equiv_real_int_transfer_at_top:
|
|
5153 |
"f \<sim>[at_top] g \<Longrightarrow> (\<lambda>x. f (of_int x :: real)) \<sim>[at_top] (\<lambda>x. g (of_int x))"
|
|
5154 |
unfolding asymp_equiv_def by (rule filterlim_of_int_at_top)
|
|
5155 |
|
|
5156 |
lemma asymp_equiv_real_int_transfer_at_bot:
|
|
5157 |
"f \<sim>[at_bot] g \<Longrightarrow> (\<lambda>x. f (of_int x :: real)) \<sim>[at_bot] (\<lambda>x. g (of_int x))"
|
|
5158 |
unfolding asymp_equiv_def by (rule filterlim_of_int_at_bot)
|
|
5159 |
|
|
5160 |
lemma eventually_int_real_at_top:
|
|
5161 |
assumes "eventually P (at_top :: real filter)"
|
|
5162 |
shows "eventually (\<lambda>x. P (of_int x)) (at_top :: int filter)"
|
|
5163 |
using assms filterlim_of_int_at_top filterlim_iff filterlim_real_of_int_at_top by blast
|
|
5164 |
|
|
5165 |
lemma eventually_int_real_at_bot:
|
|
5166 |
assumes "eventually P (at_bot :: real filter)"
|
|
5167 |
shows "eventually (\<lambda>x. P (of_int x)) (at_bot :: int filter)"
|
|
5168 |
using assms filterlim_of_int_at_bot filterlim_iff filterlim_real_of_int_at_bot by blast
|
|
5169 |
|
|
5170 |
|
|
5171 |
lemmas real_asymp_int_intros =
|
|
5172 |
filterlim_of_int_at_bot filterlim_of_int_at_top
|
|
5173 |
landau_symbols[THEN landau_symbol.compose[OF _ _ filterlim_real_of_int_at_top]]
|
|
5174 |
landau_symbols[THEN landau_symbol.compose[OF _ _ filterlim_real_of_int_at_bot]]
|
|
5175 |
asymp_equiv_real_int_transfer_at_top asymp_equiv_real_int_transfer_at_bot
|
|
5176 |
|
|
5177 |
(* TODO: Move? *)
|
|
5178 |
lemma tendsto_discrete_iff:
|
|
5179 |
"filterlim f (nhds (c :: 'a :: discrete_topology)) F \<longleftrightarrow> eventually (\<lambda>x. f x = c) F"
|
|
5180 |
by (auto simp: nhds_discrete filterlim_principal)
|
|
5181 |
|
|
5182 |
lemma tendsto_of_nat_iff:
|
|
5183 |
"filterlim (\<lambda>n. of_nat (f n)) (nhds (of_nat c :: 'a :: real_normed_div_algebra)) F \<longleftrightarrow>
|
|
5184 |
eventually (\<lambda>n. f n = c) F"
|
|
5185 |
proof
|
|
5186 |
assume "((\<lambda>n. of_nat (f n)) \<longlongrightarrow> (of_nat c :: 'a)) F"
|
|
5187 |
hence "eventually (\<lambda>n. \<bar>real (f n) - real c\<bar> < 1/2) F"
|
|
5188 |
by (force simp: tendsto_iff dist_of_nat dest: spec[of _ "1/2"])
|
|
5189 |
thus "eventually (\<lambda>n. f n = c) F"
|
|
5190 |
by eventually_elim (simp add: abs_if split: if_splits)
|
|
5191 |
next
|
|
5192 |
assume "eventually (\<lambda>n. f n = c) F"
|
|
5193 |
hence "eventually (\<lambda>n. of_nat (f n) = (of_nat c :: 'a)) F"
|
|
5194 |
by eventually_elim simp
|
|
5195 |
thus "filterlim (\<lambda>n. of_nat (f n)) (nhds (of_nat c :: 'a)) F"
|
|
5196 |
by (rule Lim_eventually)
|
|
5197 |
qed
|
|
5198 |
|
|
5199 |
lemma tendsto_of_int_iff:
|
|
5200 |
"filterlim (\<lambda>n. of_int (f n)) (nhds (of_int c :: 'a :: real_normed_div_algebra)) F \<longleftrightarrow>
|
|
5201 |
eventually (\<lambda>n. f n = c) F"
|
|
5202 |
proof
|
|
5203 |
assume "((\<lambda>n. of_int (f n)) \<longlongrightarrow> (of_int c :: 'a)) F"
|
|
5204 |
hence "eventually (\<lambda>n. \<bar>real_of_int (f n) - of_int c\<bar> < 1/2) F"
|
|
5205 |
by (force simp: tendsto_iff dist_of_int dest: spec[of _ "1/2"])
|
|
5206 |
thus "eventually (\<lambda>n. f n = c) F"
|
|
5207 |
by eventually_elim (simp add: abs_if split: if_splits)
|
|
5208 |
next
|
|
5209 |
assume "eventually (\<lambda>n. f n = c) F"
|
|
5210 |
hence "eventually (\<lambda>n. of_int (f n) = (of_int c :: 'a)) F"
|
|
5211 |
by eventually_elim simp
|
|
5212 |
thus "filterlim (\<lambda>n. of_int (f n)) (nhds (of_int c :: 'a)) F"
|
|
5213 |
by (rule Lim_eventually)
|
|
5214 |
qed
|
|
5215 |
|
|
5216 |
lemma filterlim_at_top_int_iff_filterlim_real:
|
|
5217 |
"filterlim f at_top F \<longleftrightarrow> filterlim (\<lambda>x. real_of_int (f x)) at_top F" (is "?lhs = ?rhs")
|
|
5218 |
proof
|
|
5219 |
assume ?rhs
|
|
5220 |
hence "filterlim (\<lambda>x. floor (real_of_int (f x))) at_top F"
|
|
5221 |
by (intro filterlim_compose[OF filterlim_floor_sequentially])
|
|
5222 |
also have "?this \<longleftrightarrow> ?lhs" by (intro filterlim_cong refl) auto
|
|
5223 |
finally show ?lhs .
|
|
5224 |
qed (auto intro: filterlim_compose[OF filterlim_real_of_int_at_top])
|
|
5225 |
|
|
5226 |
lemma filterlim_floor_at_bot: "filterlim (floor :: real \<Rightarrow> int) at_bot at_bot"
|
|
5227 |
proof (subst filterlim_at_bot, rule allI)
|
|
5228 |
fix C :: int show "eventually (\<lambda>x::real. \<lfloor>x\<rfloor> \<le> C) at_bot"
|
|
5229 |
using eventually_le_at_bot[of "real_of_int C"] by eventually_elim linarith
|
|
5230 |
qed
|
|
5231 |
|
|
5232 |
lemma filterlim_at_bot_int_iff_filterlim_real:
|
|
5233 |
"filterlim f at_bot F \<longleftrightarrow> filterlim (\<lambda>x. real_of_int (f x)) at_bot F" (is "?lhs = ?rhs")
|
|
5234 |
proof
|
|
5235 |
assume ?rhs
|
|
5236 |
hence "filterlim (\<lambda>x. floor (real_of_int (f x))) at_bot F"
|
|
5237 |
by (intro filterlim_compose[OF filterlim_floor_at_bot])
|
|
5238 |
also have "?this \<longleftrightarrow> ?lhs" by (intro filterlim_cong refl) auto
|
|
5239 |
finally show ?lhs .
|
|
5240 |
qed (auto intro: filterlim_compose[OF filterlim_real_of_int_at_bot])
|
|
5241 |
(* END TODO *)
|
|
5242 |
|
|
5243 |
|
|
5244 |
lemma real_asymp_real_nat_transfer:
|
|
5245 |
"filterlim (\<lambda>n. real (f n)) at_top F \<Longrightarrow> filterlim f at_top F"
|
|
5246 |
"filterlim (\<lambda>n. real (f n)) (nhds (real c)) F \<Longrightarrow> eventually (\<lambda>n. f n = c) F"
|
|
5247 |
"eventually (\<lambda>n. real (f n) \<le> real (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n \<le> g n) F"
|
|
5248 |
"eventually (\<lambda>n. real (f n) < real (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n < g n) F"
|
|
5249 |
by (simp_all add: filterlim_sequentially_iff_filterlim_real tendsto_of_nat_iff)
|
|
5250 |
|
|
5251 |
lemma real_asymp_real_int_transfer:
|
|
5252 |
"filterlim (\<lambda>n. real_of_int (f n)) at_top F \<Longrightarrow> filterlim f at_top F"
|
|
5253 |
"filterlim (\<lambda>n. real_of_int (f n)) at_bot F \<Longrightarrow> filterlim f at_bot F"
|
|
5254 |
"filterlim (\<lambda>n. real_of_int (f n)) (nhds (of_int c)) F \<Longrightarrow> eventually (\<lambda>n. f n = c) F"
|
|
5255 |
"eventually (\<lambda>n. real_of_int (f n) \<le> real_of_int (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n \<le> g n) F"
|
|
5256 |
"eventually (\<lambda>n. real_of_int (f n) < real_of_int (g n)) F \<Longrightarrow> eventually (\<lambda>n. f n < g n) F"
|
|
5257 |
by (simp_all add: tendsto_of_int_iff filterlim_at_top_int_iff_filterlim_real
|
|
5258 |
filterlim_at_bot_int_iff_filterlim_real)
|
|
5259 |
|
|
5260 |
|
|
5261 |
lemma eventually_at_left_at_right_imp_at:
|
|
5262 |
"eventually P (at_left a) \<Longrightarrow> eventually P (at_right a) \<Longrightarrow> eventually P (at (a::real))"
|
|
5263 |
by (simp add: eventually_at_split)
|
|
5264 |
|
|
5265 |
lemma filtermap_at_left_shift: "filtermap (\<lambda>x. x - d) (at_left a) = at_left (a - d)"
|
|
5266 |
for a d :: "real"
|
|
5267 |
by (simp add: filter_eq_iff eventually_filtermap eventually_at_filter filtermap_nhds_shift[symmetric])
|
|
5268 |
|
|
5269 |
lemma at_left_to_0: "at_left a = filtermap (\<lambda>x. x + a) (at_left 0)"
|
|
5270 |
for a :: real using filtermap_at_left_shift[of "-a" 0] by simp
|
|
5271 |
|
|
5272 |
lemma filterlim_at_leftI:
|
|
5273 |
assumes "filterlim (\<lambda>x. f x - c) (at_left 0) F"
|
|
5274 |
shows "filterlim f (at_left (c::real)) F"
|
|
5275 |
proof -
|
|
5276 |
from assms have "filtermap (\<lambda>x. f x - c) F \<le> at_left 0" by (simp add: filterlim_def)
|
|
5277 |
hence "filtermap (\<lambda>x. x + c) (filtermap (\<lambda>x. f x - c) F) \<le> filtermap (\<lambda>x. x + c) (at_left 0)"
|
|
5278 |
by (rule filtermap_mono)
|
|
5279 |
thus ?thesis by (subst at_left_to_0) (simp_all add: filterlim_def filtermap_filtermap)
|
|
5280 |
qed
|
|
5281 |
|
|
5282 |
lemma filterlim_at_rightI:
|
|
5283 |
assumes "filterlim (\<lambda>x. f x - c) (at_right 0) F"
|
|
5284 |
shows "filterlim f (at_right (c::real)) F"
|
|
5285 |
proof -
|
|
5286 |
from assms have "filtermap (\<lambda>x. f x - c) F \<le> at_right 0" by (simp add: filterlim_def)
|
|
5287 |
hence "filtermap (\<lambda>x. x + c) (filtermap (\<lambda>x. f x - c) F) \<le> filtermap (\<lambda>x. x + c) (at_right 0)"
|
|
5288 |
by (rule filtermap_mono)
|
|
5289 |
thus ?thesis by (subst at_right_to_0) (simp_all add: filterlim_def filtermap_filtermap)
|
|
5290 |
qed
|
|
5291 |
|
|
5292 |
lemma filterlim_atI':
|
|
5293 |
assumes "filterlim (\<lambda>x. f x - c) (at 0) F"
|
|
5294 |
shows "filterlim f (at (c::real)) F"
|
|
5295 |
proof -
|
|
5296 |
from assms have "filtermap (\<lambda>x. f x - c) F \<le> at 0" by (simp add: filterlim_def)
|
|
5297 |
hence "filtermap (\<lambda>x. x + c) (filtermap (\<lambda>x. f x - c) F) \<le> filtermap (\<lambda>x. x + c) (at 0)"
|
|
5298 |
by (rule filtermap_mono)
|
|
5299 |
thus ?thesis by (subst at_to_0) (simp_all add: filterlim_def filtermap_filtermap)
|
|
5300 |
qed
|
|
5301 |
|
|
5302 |
lemma gbinomial_series_aux_altdef:
|
|
5303 |
"gbinomial_series_aux False x n acc =
|
|
5304 |
MSLCons acc (gbinomial_series_aux False x (n + 1) ((x - n) / (n + 1) * acc))"
|
|
5305 |
"gbinomial_series_aux True x n acc =
|
|
5306 |
(if acc = 0 then MSLNil else
|
|
5307 |
MSLCons acc (gbinomial_series_aux True x (n + 1) ((x - n) / (n + 1) * acc)))"
|
|
5308 |
by (subst gbinomial_series_aux.code, simp)+
|
|
5309 |
|
|
5310 |
|
|
5311 |
|
|
5312 |
subsection \<open>Proof method setup\<close>
|
|
5313 |
|
|
5314 |
text \<open>
|
|
5315 |
The following theorem collection is the list of rewrite equations to be used by the
|
|
5316 |
lazy evaluation package. If something is missing here, normalisation of terms into
|
|
5317 |
head normal form will fail.
|
|
5318 |
\<close>
|
|
5319 |
|
|
5320 |
named_theorems real_asymp_eval_eqs
|
|
5321 |
|
|
5322 |
lemmas [real_asymp_eval_eqs] =
|
|
5323 |
msllist.sel fst_conv snd_conv CMP_BRANCH.simps plus_ms.simps minus_ms_altdef
|
|
5324 |
minus_ms_aux_MSLNil_left minus_ms_aux_MSLNil_right minus_ms_aux_MSLCons_eval times_ms.simps
|
|
5325 |
uminus_ms.simps divide_ms.simps inverse_ms.simps inverse_ms_aux_MSLCons const_expansion_ms_eval
|
|
5326 |
const_expansion_real_def times_ms_aux_MSLNil times_ms_aux_MSLCons_eval scale_shift_ms_aux_MSLCons_eval
|
|
5327 |
scale_shift_ms_aux_MSLNil uminus_ms_aux_MSLCons_eval uminus_ms_aux_MSLNil
|
|
5328 |
scale_ms_aux_MSLNil scale_ms_aux_MSLCons scale_ms_def shift_ms_aux_MSLNil shift_ms_aux_MSLCons
|
|
5329 |
scale_ms_aux'_MSLNil scale_ms_aux'_MSLCons exp_series_stream_def exp_series_stream_aux.code
|
|
5330 |
plus_ms_aux_MSLNil plus_ms_aux_MSLCons_eval map_ms_altdef map_ms_aux_MSLCons
|
|
5331 |
map_ms_aux_MSLNil lift_ms_simps powser_ms_aux_MSLNil powser_ms_aux_MSLCons
|
|
5332 |
ln_series_stream_aux_code gbinomial_series_def gbinomial_series_aux_altdef
|
|
5333 |
mssalternate.code powr_expansion_ms.simps powr_expansion_real_def powr_ms_aux_MSLCons
|
|
5334 |
power_expansion_ms.simps power_expansion_real_def power_ms_aux_MSLCons
|
|
5335 |
powser_ms_aux'_MSSCons sin_ms_aux'_altdef cos_ms_aux'_altdef const_ms_aux_def
|
|
5336 |
compare_expansions_MSLCons_eval compare_expansions_real zero_expansion_ms_def
|
|
5337 |
zero_expansion_real_def sin_series_stream_def cos_series_stream_def arctan_series_stream_def
|
|
5338 |
arctan_ms_aux_def sin_series_stream_aux_code arctan_series_stream_aux_code if_True if_False
|
|
5339 |
|
|
5340 |
text \<open>
|
|
5341 |
The following constant and theorem collection exist in order to register constructors for
|
|
5342 |
lazy evaluation. All constants marked in such a way will be passed to the lazy evaluation
|
|
5343 |
package as free constructors upon which pattern-matching can be done.
|
|
5344 |
\<close>
|
|
5345 |
definition REAL_ASYMP_EVAL_CONSTRUCTOR :: "'a \<Rightarrow> bool"
|
|
5346 |
where [simp]: "REAL_ASYMP_EVAL_CONSTRUCTOR x = True"
|
|
5347 |
|
|
5348 |
named_theorems exp_log_eval_constructor
|
|
5349 |
|
|
5350 |
lemma exp_log_eval_constructors [exp_log_eval_constructor]:
|
|
5351 |
"REAL_ASYMP_EVAL_CONSTRUCTOR MSLNil"
|
|
5352 |
"REAL_ASYMP_EVAL_CONSTRUCTOR MSLCons"
|
|
5353 |
"REAL_ASYMP_EVAL_CONSTRUCTOR MSSCons"
|
|
5354 |
"REAL_ASYMP_EVAL_CONSTRUCTOR LT"
|
|
5355 |
"REAL_ASYMP_EVAL_CONSTRUCTOR EQ"
|
|
5356 |
"REAL_ASYMP_EVAL_CONSTRUCTOR GT"
|
|
5357 |
"REAL_ASYMP_EVAL_CONSTRUCTOR Pair"
|
|
5358 |
"REAL_ASYMP_EVAL_CONSTRUCTOR True"
|
|
5359 |
"REAL_ASYMP_EVAL_CONSTRUCTOR False"
|
|
5360 |
"REAL_ASYMP_EVAL_CONSTRUCTOR MS"
|
|
5361 |
by simp_all
|
|
5362 |
|
|
5363 |
text \<open>
|
|
5364 |
The following constant exists in order to mark functions as having plug-in support
|
|
5365 |
for multiseries expansions (i.\,e.\ this can be used to add support for arbitrary functions
|
|
5366 |
later on)
|
|
5367 |
\<close>
|
|
5368 |
definition REAL_ASYMP_CUSTOM :: "'a \<Rightarrow> bool"
|
|
5369 |
where [simp]: "REAL_ASYMP_CUSTOM x = True"
|
|
5370 |
|
|
5371 |
lemmas [simp del] = ms.map inverse_ms_aux.simps divide_ms.simps
|
|
5372 |
|
|
5373 |
definition expansion_with_remainder_term :: "(real \<Rightarrow> real) \<Rightarrow> (real \<Rightarrow> real) set \<Rightarrow> bool" where
|
|
5374 |
"expansion_with_remainder_term _ _ = True"
|
|
5375 |
|
|
5376 |
notation (output) expansion_with_remainder_term (infixl "+" 10)
|
|
5377 |
|
|
5378 |
ML_file \<open>asymptotic_basis.ML\<close>
|
|
5379 |
ML_file \<open>exp_log_expression.ML\<close>
|
|
5380 |
ML_file \<open>expansion_interface.ML\<close>
|
|
5381 |
ML_file \<open>multiseries_expansion.ML\<close>
|
|
5382 |
ML_file \<open>real_asymp.ML\<close>
|
|
5383 |
|
|
5384 |
method_setup real_asymp =
|
|
5385 |
\<open>(Scan.lift (Scan.optional (Args.parens (Args.$$$ "verbose") >> K true) false) --|
|
|
5386 |
Method.sections Simplifier.simp_modifiers) >>
|
|
5387 |
(fn verbose => fn ctxt => SIMPLE_METHOD' (Real_Asymp_Basic.tac verbose ctxt))\<close>
|
|
5388 |
"Semi-automatic decision procedure for asymptotics of exp-log functions"
|
|
5389 |
|
|
5390 |
end
|