29687
|
1 |
(* Title: Formal_Power_Series.thy
|
|
2 |
ID:
|
|
3 |
Author: Amine Chaieb, University of Cambridge
|
|
4 |
*)
|
|
5 |
|
|
6 |
header{* A formalization of formal power series *}
|
|
7 |
|
|
8 |
theory Formal_Power_Series
|
|
9 |
imports Main Fact Parity
|
|
10 |
begin
|
|
11 |
|
|
12 |
section {* The type of formal power series*}
|
|
13 |
|
|
14 |
typedef 'a fps = "{f :: nat \<Rightarrow> 'a. True}"
|
|
15 |
by simp
|
|
16 |
|
|
17 |
text{* Definition of the basic elements 0 and 1 and the basic operations of addition, negation and multiplication *}
|
|
18 |
|
|
19 |
instantiation fps :: (zero) zero
|
|
20 |
begin
|
|
21 |
|
|
22 |
definition fps_zero_def: "(0 :: 'a fps) \<equiv> Abs_fps (\<lambda>(n::nat). 0)"
|
|
23 |
instance ..
|
|
24 |
end
|
|
25 |
|
|
26 |
instantiation fps :: ("{one,zero}") one
|
|
27 |
begin
|
|
28 |
|
|
29 |
definition fps_one_def: "(1 :: 'a fps) \<equiv> Abs_fps (\<lambda>(n::nat). if n = 0 then 1 else 0)"
|
|
30 |
instance ..
|
|
31 |
end
|
|
32 |
|
|
33 |
instantiation fps :: (plus) plus
|
|
34 |
begin
|
|
35 |
|
|
36 |
definition fps_plus_def: "op + \<equiv> (\<lambda>(f::'a fps) (g:: 'a fps). Abs_fps (\<lambda>(n::nat). Rep_fps (f) n + Rep_fps (g) n))"
|
|
37 |
instance ..
|
|
38 |
end
|
|
39 |
|
|
40 |
instantiation fps :: (minus) minus
|
|
41 |
begin
|
|
42 |
|
|
43 |
definition fps_minus_def: "op - \<equiv> (\<lambda>(f::'a fps) (g:: 'a fps). Abs_fps (\<lambda>(n::nat). Rep_fps (f) n - Rep_fps (g) n))"
|
|
44 |
instance ..
|
|
45 |
end
|
|
46 |
|
|
47 |
instantiation fps :: (uminus) uminus
|
|
48 |
begin
|
|
49 |
|
|
50 |
definition fps_uminus_def: "uminus \<equiv> (\<lambda>(f::'a fps). Abs_fps (\<lambda>(n::nat). - Rep_fps (f) n))"
|
|
51 |
instance ..
|
|
52 |
end
|
|
53 |
|
|
54 |
instantiation fps :: ("{comm_monoid_add, times}") times
|
|
55 |
begin
|
|
56 |
|
|
57 |
definition fps_times_def:
|
|
58 |
"op * \<equiv> (\<lambda>(f::'a fps) (g:: 'a fps). Abs_fps (\<lambda>(n::nat). setsum (\<lambda>i. Rep_fps (f) i * Rep_fps (g) (n - i)) {0.. n}))"
|
|
59 |
instance ..
|
|
60 |
end
|
|
61 |
|
|
62 |
text{* Some useful theorems to get rid of Abs and Rep *}
|
|
63 |
|
|
64 |
lemma mem_fps_set_trivial[intro, simp]: "f \<in> fps" unfolding fps_def by blast
|
|
65 |
lemma Rep_fps_Abs_fps[simp]: "Rep_fps (Abs_fps f) = f"
|
|
66 |
by (blast intro: Abs_fps_inverse)
|
|
67 |
lemma Abs_fps_Rep_fps[simp]: "Abs_fps (Rep_fps f) = f"
|
|
68 |
by (blast intro: Rep_fps_inverse)
|
|
69 |
lemma Abs_fps_eq[simp]: "Abs_fps f = Abs_fps g \<longleftrightarrow> f = g"
|
|
70 |
proof-
|
|
71 |
{assume "f = g" hence "Abs_fps f = Abs_fps g" by simp}
|
|
72 |
moreover
|
|
73 |
{assume a: "Abs_fps f = Abs_fps g"
|
|
74 |
from a have "Rep_fps (Abs_fps f) = Rep_fps (Abs_fps g)" by simp
|
|
75 |
hence "f = g" by simp}
|
|
76 |
ultimately show ?thesis by blast
|
|
77 |
qed
|
|
78 |
|
|
79 |
lemma Rep_fps_eq[simp]: "Rep_fps f = Rep_fps g \<longleftrightarrow> f = g"
|
|
80 |
proof-
|
|
81 |
{assume "Rep_fps f = Rep_fps g"
|
|
82 |
hence "Abs_fps (Rep_fps f) = Abs_fps (Rep_fps g)" by simp hence "f=g" by simp}
|
|
83 |
moreover
|
|
84 |
{assume "f = g" hence "Rep_fps f = Rep_fps g" by simp}
|
|
85 |
ultimately show ?thesis by blast
|
|
86 |
qed
|
|
87 |
|
|
88 |
declare atLeastAtMost_iff[presburger]
|
|
89 |
declare Bex_def[presburger]
|
|
90 |
declare Ball_def[presburger]
|
|
91 |
|
|
92 |
lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"
|
|
93 |
by auto
|
|
94 |
lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
|
|
95 |
by auto
|
|
96 |
|
|
97 |
section{* Formal power series form a commutative ring with unity, if the range of sequences
|
|
98 |
they represent is a commutative ring with unity*}
|
|
99 |
|
|
100 |
instantiation fps :: (semigroup_add) semigroup_add
|
|
101 |
begin
|
|
102 |
|
|
103 |
instance
|
|
104 |
proof
|
|
105 |
fix a b c :: "'a fps" show "a + b + c = a + (b + c)"
|
|
106 |
by (auto simp add: fps_plus_def expand_fun_eq add_assoc)
|
|
107 |
qed
|
|
108 |
|
|
109 |
end
|
|
110 |
|
|
111 |
instantiation fps :: (ab_semigroup_add) ab_semigroup_add
|
|
112 |
begin
|
|
113 |
|
|
114 |
instance by (intro_classes, simp add: fps_plus_def expand_fun_eq add_commute)
|
|
115 |
end
|
|
116 |
|
|
117 |
instantiation fps :: (semiring_1) semigroup_mult
|
|
118 |
begin
|
|
119 |
|
|
120 |
instance
|
|
121 |
proof
|
|
122 |
fix a b c :: "'a fps"
|
|
123 |
let ?a = "Rep_fps a"
|
|
124 |
let ?b = "Rep_fps b"
|
|
125 |
let ?c = "Rep_fps c"
|
|
126 |
let ?x = "\<lambda> i k. if k \<le> i then (1::'a) else 0"
|
|
127 |
show "a*b*c = a* (b * c)"
|
|
128 |
proof(auto simp add: fps_times_def setsum_right_distrib setsum_left_distrib, rule ext)
|
|
129 |
fix n::nat
|
|
130 |
let ?r = "\<lambda>i. n - i"
|
|
131 |
have i: "inj_on ?r {0..n}" by (auto simp add: inj_on_def)
|
|
132 |
have ri: "{0 .. n} = ?r ` {0..n}" apply (auto simp add: image_iff)
|
|
133 |
by presburger
|
|
134 |
let ?f = "\<lambda>i j. ?a j * ?b (i - j) * ?c (n -i)"
|
|
135 |
let ?g = "\<lambda>i j. ?a i * (?b j * ?c (n - (i + j)))"
|
|
136 |
have "setsum (\<lambda>i. setsum (?f i) {0..i}) {0..n}
|
|
137 |
= setsum (\<lambda>i. setsum (\<lambda>j. ?f i j * ?x i j) {0..i}) {0..n}"
|
|
138 |
by (rule setsum_cong2)+ auto
|
|
139 |
also have "\<dots> = setsum (\<lambda>i. setsum (\<lambda>j. ?f i j * ?x i j) {0..n}) {0..n}"
|
|
140 |
proof(rule setsum_cong2)
|
|
141 |
fix i assume i: "i \<in> {0..n}"
|
|
142 |
have eq: "{0 .. n} = {0 ..i} \<union> {i+1 .. n}" using i by auto
|
|
143 |
have d: "{0 ..i} \<inter> {i+1 .. n} = {}" using i by auto
|
|
144 |
have f: "finite {0..i}" "finite {i+1 ..n}" by auto
|
|
145 |
have s0: "setsum (\<lambda>j. ?f i j * ?x i j) {i+1 ..n} = 0" by simp
|
|
146 |
show "setsum (\<lambda>j. ?f i j * ?x i j) {0..i} = setsum (\<lambda>j. ?f i j * ?x i j) {0..n}"
|
|
147 |
unfolding eq setsum_Un_disjoint[OF f d] s0
|
|
148 |
by simp
|
|
149 |
qed
|
|
150 |
also have "\<dots> = setsum (\<lambda>i. setsum (\<lambda>j. ?f j i * ?x j i) {0 .. n}) {0 .. n}"
|
|
151 |
by (rule setsum_commute)
|
|
152 |
also have "\<dots> = setsum (\<lambda>i. setsum (\<lambda>j. ?f j i * ?x j i) {i .. n}) {0 .. n}"
|
|
153 |
apply(rule setsum_cong2)
|
|
154 |
apply (rule setsum_mono_zero_right)
|
|
155 |
apply auto
|
|
156 |
done
|
|
157 |
also have "\<dots> = setsum (\<lambda>i. setsum (?g i) {0..n - i}) {0..n}"
|
|
158 |
apply (rule setsum_cong2)
|
|
159 |
apply (rule_tac f="\<lambda>i. i + x" in setsum_reindex_cong)
|
|
160 |
apply (simp add: inj_on_def)
|
|
161 |
apply (rule set_ext)
|
|
162 |
apply (presburger add: image_iff)
|
|
163 |
by (simp add: add_ac mult_assoc)
|
|
164 |
finally show "setsum (\<lambda>i. setsum (\<lambda>j. ?a j * ?b (i - j) * ?c (n -i)) {0..i}) {0..n}
|
|
165 |
= setsum (\<lambda>i. setsum (\<lambda>j. ?a i * (?b j * ?c (n - (i + j)))) {0..n - i}) {0..n}".
|
|
166 |
qed
|
|
167 |
qed
|
|
168 |
|
|
169 |
end
|
|
170 |
|
|
171 |
instantiation fps :: (comm_semiring_1) ab_semigroup_mult
|
|
172 |
begin
|
|
173 |
|
|
174 |
instance
|
|
175 |
proof
|
|
176 |
fix a b :: "'a fps"
|
|
177 |
show "a*b = b*a"
|
|
178 |
apply(auto simp add: fps_times_def setsum_right_distrib setsum_left_distrib, rule ext)
|
|
179 |
apply (rule_tac f = "\<lambda>i. n - i" in setsum_reindex_cong)
|
|
180 |
apply (simp add: inj_on_def)
|
|
181 |
apply presburger
|
|
182 |
apply (rule set_ext)
|
|
183 |
apply (presburger add: image_iff)
|
|
184 |
by (simp add: mult_commute)
|
|
185 |
qed
|
|
186 |
end
|
|
187 |
|
|
188 |
instantiation fps :: (monoid_add) monoid_add
|
|
189 |
begin
|
|
190 |
|
|
191 |
instance
|
|
192 |
proof
|
|
193 |
fix a :: "'a fps" show "0 + a = a "
|
|
194 |
by (auto simp add: fps_plus_def fps_zero_def intro: ext)
|
|
195 |
next
|
|
196 |
fix a :: "'a fps" show "a + 0 = a "
|
|
197 |
by (auto simp add: fps_plus_def fps_zero_def intro: ext)
|
|
198 |
qed
|
|
199 |
|
|
200 |
end
|
|
201 |
instantiation fps :: (comm_monoid_add) comm_monoid_add
|
|
202 |
begin
|
|
203 |
|
|
204 |
instance
|
|
205 |
proof
|
|
206 |
fix a :: "'a fps" show "0 + a = a "
|
|
207 |
by (auto simp add: fps_plus_def fps_zero_def intro: ext)
|
|
208 |
qed
|
|
209 |
|
|
210 |
end
|
|
211 |
|
|
212 |
instantiation fps :: (semiring_1) monoid_mult
|
|
213 |
begin
|
|
214 |
|
|
215 |
instance
|
|
216 |
proof
|
|
217 |
fix a :: "'a fps" show "1 * a = a"
|
|
218 |
apply (auto simp add: fps_one_def fps_times_def)
|
|
219 |
apply (subst (2) Abs_fps_Rep_fps[of a, symmetric])
|
|
220 |
unfolding Abs_fps_eq
|
|
221 |
apply (rule ext)
|
|
222 |
by (simp add: cond_value_iff cond_application_beta setsum_delta cong del: if_weak_cong)
|
|
223 |
next
|
|
224 |
fix a :: "'a fps" show "a*1 = a"
|
|
225 |
apply (auto simp add: fps_one_def fps_times_def)
|
|
226 |
apply (subst (2) Abs_fps_Rep_fps[of a, symmetric])
|
|
227 |
unfolding Abs_fps_eq
|
|
228 |
apply (rule ext)
|
|
229 |
by (simp add: cond_value_iff cond_application_beta setsum_delta' cong del: if_weak_cong)
|
|
230 |
qed
|
|
231 |
end
|
|
232 |
|
|
233 |
instantiation fps :: (cancel_semigroup_add) cancel_semigroup_add
|
|
234 |
begin
|
|
235 |
|
|
236 |
instance by (intro_classes) (auto simp add: fps_plus_def expand_fun_eq Rep_fps_eq[symmetric])
|
|
237 |
end
|
|
238 |
|
|
239 |
instantiation fps :: (group_add) group_add
|
|
240 |
begin
|
|
241 |
|
|
242 |
instance
|
|
243 |
proof
|
|
244 |
fix a :: "'a fps" show "- a + a = 0"
|
|
245 |
by (auto simp add: fps_plus_def fps_uminus_def fps_zero_def intro: ext)
|
|
246 |
next
|
|
247 |
fix a b :: "'a fps" show "a - b = a + - b"
|
|
248 |
by (auto simp add: fps_plus_def fps_uminus_def fps_zero_def
|
|
249 |
fps_minus_def expand_fun_eq diff_minus)
|
|
250 |
qed
|
|
251 |
end
|
|
252 |
|
|
253 |
context comm_ring_1
|
|
254 |
begin
|
|
255 |
subclass group_add proof qed
|
|
256 |
end
|
|
257 |
|
|
258 |
instantiation fps :: (zero_neq_one) zero_neq_one
|
|
259 |
begin
|
|
260 |
instance by (intro_classes, auto simp add: zero_neq_one
|
|
261 |
fps_one_def fps_zero_def expand_fun_eq)
|
|
262 |
end
|
|
263 |
|
|
264 |
instantiation fps :: (semiring_1) semiring
|
|
265 |
begin
|
|
266 |
|
|
267 |
instance
|
|
268 |
proof
|
|
269 |
fix a b c :: "'a fps"
|
|
270 |
show "(a + b) * c = a * c + b*c"
|
|
271 |
apply (auto simp add: fps_plus_def fps_times_def, rule ext)
|
|
272 |
unfolding setsum_addf[symmetric]
|
|
273 |
apply (simp add: ring_simps)
|
|
274 |
done
|
|
275 |
next
|
|
276 |
fix a b c :: "'a fps"
|
|
277 |
show "a * (b + c) = a * b + a*c"
|
|
278 |
apply (auto simp add: fps_plus_def fps_times_def, rule ext)
|
|
279 |
unfolding setsum_addf[symmetric]
|
|
280 |
apply (simp add: ring_simps)
|
|
281 |
done
|
|
282 |
qed
|
|
283 |
end
|
|
284 |
|
|
285 |
instantiation fps :: (semiring_1) semiring_0
|
|
286 |
begin
|
|
287 |
|
|
288 |
instance
|
|
289 |
proof
|
|
290 |
fix a:: "'a fps" show "0 * a = 0" by (simp add: fps_zero_def fps_times_def)
|
|
291 |
next
|
|
292 |
fix a:: "'a fps" show "a*0 = 0" by (simp add: fps_zero_def fps_times_def)
|
|
293 |
qed
|
|
294 |
end
|
|
295 |
|
|
296 |
section {* Selection of the nth power of the implicit variable in the infinite sum*}
|
|
297 |
|
|
298 |
definition fps_nth:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a" (infixl "$" 75)
|
|
299 |
where "f $ n = Rep_fps f n"
|
|
300 |
|
|
301 |
lemma fps_nth_Abs_fps[simp]: "Abs_fps a $ n = a n" by (simp add: fps_nth_def)
|
|
302 |
lemma fps_zero_nth[simp]: "0 $ n = 0" by (simp add: fps_zero_def)
|
|
303 |
lemma fps_one_nth[simp]: "1 $ n = (if n = 0 then 1 else 0)"
|
|
304 |
by (simp add: fps_one_def)
|
|
305 |
lemma fps_add_nth[simp]: "(f + g) $ n = f$n + g$n" by (simp add: fps_plus_def fps_nth_def)
|
|
306 |
lemma fps_mult_nth: "(f * g) $ n = setsum (\<lambda>i. f$i * g$(n - i)) {0..n}"
|
|
307 |
by (simp add: fps_times_def fps_nth_def)
|
|
308 |
lemma fps_neg_nth[simp]: "(- f) $n = - (f $n)" by (simp add: fps_nth_def fps_uminus_def)
|
|
309 |
lemma fps_sub_nth[simp]: "(f - g)$n = f$n - g$n" by (simp add: fps_nth_def fps_minus_def)
|
|
310 |
|
|
311 |
lemma fps_nonzero_nth: "f \<noteq> 0 \<longleftrightarrow> (\<exists> n. f $n \<noteq> 0)"
|
|
312 |
proof-
|
|
313 |
{assume "f \<noteq> 0"
|
|
314 |
hence "Rep_fps f \<noteq> Rep_fps 0" by simp
|
|
315 |
hence "\<exists>n. f $n \<noteq> 0" by (simp add: expand_fun_eq fps_zero_def fps_nth_def )}
|
|
316 |
moreover
|
|
317 |
{assume "\<exists>n. f$n \<noteq> 0" and "f = 0"
|
|
318 |
then have False by (simp add: expand_fun_eq fps_zero_def fps_nth_def )}
|
|
319 |
ultimately show ?thesis by blast
|
|
320 |
qed
|
|
321 |
|
|
322 |
lemma fps_nonzero_nth_minimal: "f \<noteq> 0 \<longleftrightarrow> (\<exists> n. f $n \<noteq> 0 \<and> (\<forall>m <n. f $m = 0))"
|
|
323 |
proof-
|
|
324 |
let ?S = "{n. f$n \<noteq> 0}"
|
|
325 |
{assume "\<exists>n. f$n \<noteq> 0 \<and> (\<forall>m <n. f $m = 0)" and "f = 0"
|
|
326 |
then have False by (simp add: expand_fun_eq fps_zero_def fps_nth_def )}
|
|
327 |
moreover
|
|
328 |
{assume f0: "f \<noteq> 0"
|
|
329 |
from f0 fps_nonzero_nth have ex: "\<exists>n. f$n \<noteq> 0" by blast
|
|
330 |
hence Se: "?S\<noteq> {}" by blast
|
|
331 |
from ex obtain n where n: "f$n \<noteq> 0" by blast
|
|
332 |
from n have nS: "n \<in> ?S" by blast
|
|
333 |
let ?U = "?S \<inter> {0..n}"
|
|
334 |
have fU: "finite ?U" by auto
|
|
335 |
from n have Ue: "?U \<noteq> {}" by auto
|
|
336 |
let ?m = "Min ?U"
|
|
337 |
have mU: "?m \<in> ?U" using Min_in[OF fU Ue] .
|
|
338 |
hence mn: "?m \<le> n" by simp
|
|
339 |
from mU have mf: "f $ ?m \<noteq> 0" by blast
|
|
340 |
{fix m assume m: "m < ?m" and f: "f $m \<noteq> 0"
|
|
341 |
from m mn have mn': "m < n" by arith
|
|
342 |
with f have mU': "m \<in> ?U" by simp
|
|
343 |
from Min_le[OF fU mU'] m have False by arith}
|
|
344 |
hence "\<forall>m <?m. f$m = 0" by blast
|
|
345 |
with mf have "\<exists> n. f $n \<noteq> 0 \<and> (\<forall>m <n. f $m = 0)" by blast}
|
|
346 |
ultimately show ?thesis by blast
|
|
347 |
qed
|
|
348 |
|
|
349 |
lemma fps_eq_iff: "f = g \<longleftrightarrow> (\<forall>n. f $ n = g $n)"
|
|
350 |
by (auto simp add: fps_nth_def Rep_fps_eq[unfolded expand_fun_eq])
|
|
351 |
|
|
352 |
lemma fps_setsum_nth: "(setsum f S) $ n = setsum (\<lambda>k. (f k) $ n) S"
|
|
353 |
proof-
|
|
354 |
{assume "\<not> finite S" hence ?thesis by simp}
|
|
355 |
moreover
|
|
356 |
{assume fS: "finite S"
|
|
357 |
have ?thesis by(induct rule: finite_induct[OF fS]) auto}
|
|
358 |
ultimately show ?thesis by blast
|
|
359 |
qed
|
|
360 |
|
|
361 |
section{* Injection of the basic ring elements and multiplication by scalars *}
|
|
362 |
|
|
363 |
definition "fps_const c = Abs_fps (\<lambda>n. if n = 0 then c else 0)"
|
|
364 |
lemma fps_const_0_eq_0[simp]: "fps_const 0 = 0" by (simp add: fps_const_def fps_eq_iff)
|
|
365 |
lemma fps_const_1_eq_1[simp]: "fps_const 1 = 1" by (simp add: fps_const_def fps_eq_iff)
|
|
366 |
lemma fps_const_neg[simp]: "- (fps_const (c::'a::ring)) = fps_const (- c)"
|
|
367 |
by (simp add: fps_uminus_def fps_const_def fps_eq_iff)
|
|
368 |
lemma fps_const_add[simp]: "fps_const (c::'a\<Colon>monoid_add) + fps_const d = fps_const (c + d)"
|
|
369 |
by (simp add: fps_plus_def fps_const_def fps_eq_iff)
|
|
370 |
lemma fps_const_mult[simp]: "fps_const (c::'a\<Colon>ring) * fps_const d = fps_const (c * d)"
|
|
371 |
by (auto simp add: fps_times_def fps_const_def fps_eq_iff intro: setsum_0')
|
|
372 |
|
|
373 |
lemma fps_const_add_left: "fps_const (c::'a\<Colon>monoid_add) + f = Abs_fps (\<lambda>n. if n = 0 then c + f$0 else f$n)"
|
|
374 |
unfolding fps_eq_iff fps_add_nth by (simp add: fps_const_def)
|
|
375 |
lemma fps_const_add_right: "f + fps_const (c::'a\<Colon>monoid_add) = Abs_fps (\<lambda>n. if n = 0 then f$0 + c else f$n)"
|
|
376 |
unfolding fps_eq_iff fps_add_nth by (simp add: fps_const_def)
|
|
377 |
|
|
378 |
lemma fps_const_mult_left: "fps_const (c::'a\<Colon>semiring_0) * f = Abs_fps (\<lambda>n. c * f$n)"
|
|
379 |
unfolding fps_eq_iff fps_mult_nth
|
|
380 |
by (simp add: fps_const_def cond_application_beta cond_value_iff setsum_delta cong del: if_weak_cong)
|
|
381 |
lemma fps_const_mult_right: "f * fps_const (c::'a\<Colon>semiring_0) = Abs_fps (\<lambda>n. f$n * c)"
|
|
382 |
unfolding fps_eq_iff fps_mult_nth
|
|
383 |
by (simp add: fps_const_def cond_application_beta cond_value_iff setsum_delta' cong del: if_weak_cong)
|
|
384 |
|
|
385 |
lemma fps_const_nth[simp]: "(fps_const c) $n = (if n = 0 then c else 0)"
|
|
386 |
by (simp add: fps_const_def)
|
|
387 |
|
|
388 |
lemma fps_mult_left_const_nth[simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
|
|
389 |
by (simp add: fps_mult_nth fps_const_nth cond_application_beta cond_value_iff setsum_delta cong del: if_weak_cong)
|
|
390 |
|
|
391 |
lemma fps_mult_right_const_nth[simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
|
|
392 |
by (simp add: fps_mult_nth fps_const_nth cond_application_beta cond_value_iff setsum_delta' cong del: if_weak_cong)
|
|
393 |
|
|
394 |
section {* Formal power series form an integral domain*}
|
|
395 |
|
|
396 |
instantiation fps :: (ring_1) ring_1
|
|
397 |
begin
|
|
398 |
|
|
399 |
instance by (intro_classes, auto simp add: diff_minus left_distrib)
|
|
400 |
end
|
|
401 |
|
|
402 |
instantiation fps :: (comm_ring_1) comm_ring_1
|
|
403 |
begin
|
|
404 |
|
|
405 |
instance by (intro_classes, auto simp add: diff_minus left_distrib)
|
|
406 |
end
|
|
407 |
instantiation fps :: ("{ring_no_zero_divisors, comm_ring_1}") ring_no_zero_divisors
|
|
408 |
begin
|
|
409 |
|
|
410 |
instance
|
|
411 |
proof
|
|
412 |
fix a b :: "'a fps"
|
|
413 |
assume a0: "a \<noteq> 0" and b0: "b \<noteq> 0"
|
|
414 |
then obtain i j where i: "a$i\<noteq>0" "\<forall>k<i. a$k=0"
|
|
415 |
and j: "b$j \<noteq>0" "\<forall>k<j. b$k =0" unfolding fps_nonzero_nth_minimal
|
|
416 |
by blast+
|
|
417 |
have eq: "({0..i+j} -{i}) \<union> {i} = {0..i+j}" by auto
|
|
418 |
have d: "({0..i+j} -{i}) \<inter> {i} = {}" by auto
|
|
419 |
have f: "finite ({0..i+j} -{i})" "finite {i}" by auto
|
|
420 |
have th0: "setsum (\<lambda>k. a$k * b$(i+j - k)) ({0..i+j} -{i}) = 0"
|
|
421 |
apply (rule setsum_0')
|
|
422 |
apply auto
|
|
423 |
apply (case_tac "aa < i")
|
|
424 |
using i
|
|
425 |
apply auto
|
|
426 |
apply (subgoal_tac "b $ (i+j - aa) = 0")
|
|
427 |
apply blast
|
|
428 |
apply (rule j(2)[rule_format])
|
|
429 |
by arith
|
|
430 |
have "(a*b) $ (i+j) = setsum (\<lambda>k. a$k * b$(i+j - k)) {0..i+j}"
|
|
431 |
by (rule fps_mult_nth)
|
|
432 |
hence "(a*b) $ (i+j) = a$i * b$j"
|
|
433 |
unfolding setsum_Un_disjoint[OF f d, unfolded eq] th0 by simp
|
|
434 |
with i j have "(a*b) $ (i+j) \<noteq> 0" by simp
|
|
435 |
then show "a*b \<noteq> 0" unfolding fps_nonzero_nth by blast
|
|
436 |
qed
|
|
437 |
end
|
|
438 |
|
|
439 |
instantiation fps :: (idom) idom
|
|
440 |
begin
|
|
441 |
|
|
442 |
instance ..
|
|
443 |
end
|
|
444 |
|
|
445 |
section{* Inverses of formal power series *}
|
|
446 |
|
|
447 |
declare setsum_cong[fundef_cong]
|
|
448 |
|
|
449 |
|
|
450 |
instantiation fps :: ("{comm_monoid_add,inverse, times, uminus}") inverse
|
|
451 |
begin
|
|
452 |
|
|
453 |
fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a" where
|
|
454 |
"natfun_inverse f 0 = inverse (f$0)"
|
|
455 |
| "natfun_inverse f n = - inverse (f$0) * setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
|
|
456 |
|
|
457 |
definition fps_inverse_def:
|
|
458 |
"inverse f = (if f$0 = 0 then 0 else Abs_fps (natfun_inverse f))"
|
|
459 |
definition fps_divide_def: "divide \<equiv> (\<lambda>(f::'a fps) g. f * inverse g)"
|
|
460 |
instance ..
|
|
461 |
end
|
|
462 |
|
|
463 |
lemma fps_inverse_zero[simp]:
|
|
464 |
"inverse (0 :: 'a::{comm_monoid_add,inverse, times, uminus} fps) = 0"
|
|
465 |
by (simp add: fps_zero_def fps_inverse_def)
|
|
466 |
|
|
467 |
lemma fps_inverse_one[simp]: "inverse (1 :: 'a::{division_ring,zero_neq_one} fps) = 1"
|
|
468 |
apply (auto simp add: fps_one_def fps_inverse_def expand_fun_eq)
|
|
469 |
by (case_tac x, auto)
|
|
470 |
|
|
471 |
instantiation fps :: ("{comm_monoid_add,inverse, times, uminus}") division_by_zero
|
|
472 |
begin
|
|
473 |
instance
|
|
474 |
apply (intro_classes)
|
|
475 |
by (rule fps_inverse_zero)
|
|
476 |
end
|
|
477 |
|
|
478 |
lemma inverse_mult_eq_1[intro]: assumes f0: "f$0 \<noteq> (0::'a::field)"
|
|
479 |
shows "inverse f * f = 1"
|
|
480 |
proof-
|
|
481 |
have c: "inverse f * f = f * inverse f" by (simp add: mult_commute)
|
|
482 |
from f0 have ifn: "\<And>n. inverse f $ n = natfun_inverse f n"
|
|
483 |
by (simp add: fps_inverse_def)
|
|
484 |
from f0 have th0: "(inverse f * f) $ 0 = 1"
|
|
485 |
by (simp add: fps_inverse_def fps_one_def fps_mult_nth)
|
|
486 |
{fix n::nat assume np: "n >0 "
|
|
487 |
from np have eq: "{0..n} = {0} \<union> {1 .. n}" by auto
|
|
488 |
have d: "{0} \<inter> {1 .. n} = {}" by auto
|
|
489 |
have f: "finite {0::nat}" "finite {1..n}" by auto
|
|
490 |
from f0 np have th0: "- (inverse f$n) =
|
|
491 |
(setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
|
|
492 |
by (cases n, simp_all add: divide_inverse fps_inverse_def fps_nth_def ring_simps)
|
|
493 |
from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
|
|
494 |
have th1: "setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} =
|
|
495 |
- (f$0) * (inverse f)$n"
|
|
496 |
by (simp add: ring_simps)
|
|
497 |
have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
|
|
498 |
unfolding fps_mult_nth ifn ..
|
|
499 |
also have "\<dots> = f$0 * natfun_inverse f n
|
|
500 |
+ (\<Sum>i = 1..n. f$i * natfun_inverse f (n-i))"
|
|
501 |
unfolding setsum_Un_disjoint[OF f d, unfolded eq[symmetric]]
|
|
502 |
by simp
|
|
503 |
also have "\<dots> = 0" unfolding th1 ifn by simp
|
|
504 |
finally have "(inverse f * f)$n = 0" unfolding c . }
|
|
505 |
with th0 show ?thesis by (simp add: fps_eq_iff)
|
|
506 |
qed
|
|
507 |
|
|
508 |
lemma fps_inverse_0_iff[simp]: "(inverse f)$0 = (0::'a::division_ring) \<longleftrightarrow> f$0 = 0"
|
|
509 |
apply (simp add: fps_inverse_def)
|
|
510 |
by (metis fps_nth_def fps_nth_def inverse_zero_imp_zero)
|
|
511 |
|
|
512 |
lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::field) fps) \<longleftrightarrow> f $0 = 0"
|
|
513 |
proof-
|
|
514 |
{assume "f$0 = 0" hence "inverse f = 0" by (simp add: fps_inverse_def)}
|
|
515 |
moreover
|
|
516 |
{assume h: "inverse f = 0" and c: "f $0 \<noteq> 0"
|
|
517 |
from inverse_mult_eq_1[OF c] h have False by simp}
|
|
518 |
ultimately show ?thesis by blast
|
|
519 |
qed
|
|
520 |
|
|
521 |
lemma fps_inverse_idempotent[intro]: assumes f0: "f$0 \<noteq> (0::'a::field)"
|
|
522 |
shows "inverse (inverse f) = f"
|
|
523 |
proof-
|
|
524 |
from f0 have if0: "inverse f $ 0 \<noteq> 0" by simp
|
|
525 |
from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0]
|
|
526 |
have th0: "inverse f * f = inverse f * inverse (inverse f)" by (simp add: mult_ac)
|
|
527 |
then show ?thesis using f0 unfolding mult_cancel_left by simp
|
|
528 |
qed
|
|
529 |
|
|
530 |
lemma fps_inverse_unique: assumes f0: "f$0 \<noteq> (0::'a::field)" and fg: "f*g = 1"
|
|
531 |
shows "inverse f = g"
|
|
532 |
proof-
|
|
533 |
from inverse_mult_eq_1[OF f0] fg
|
|
534 |
have th0: "inverse f * f = g * f" by (simp add: mult_ac)
|
|
535 |
then show ?thesis using f0 unfolding mult_cancel_right
|
|
536 |
unfolding Rep_fps_eq[of f 0, symmetric]
|
|
537 |
by (auto simp add: fps_zero_def expand_fun_eq fps_nth_def)
|
|
538 |
qed
|
|
539 |
|
|
540 |
lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
|
|
541 |
= Abs_fps (\<lambda>n. if n= 0 then 1 else if n=1 then - 1 else 0)"
|
|
542 |
apply (rule fps_inverse_unique)
|
|
543 |
apply simp
|
|
544 |
apply (simp add: fps_eq_iff fps_nth_def fps_times_def fps_one_def)
|
|
545 |
proof(clarsimp)
|
|
546 |
fix n::nat assume n: "n > 0"
|
|
547 |
let ?f = "\<lambda>i. if n = i then (1\<Colon>'a) else if n - i = 1 then - 1 else 0"
|
|
548 |
let ?g = "\<lambda>i. if i = n then 1 else if i=n - 1 then - 1 else 0"
|
|
549 |
let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
|
|
550 |
have th1: "setsum ?f {0..n} = setsum ?g {0..n}"
|
|
551 |
by (rule setsum_cong2) auto
|
|
552 |
have th2: "setsum ?g {0..n - 1} = setsum ?h {0..n - 1}"
|
|
553 |
using n apply - by (rule setsum_cong2) auto
|
|
554 |
have eq: "{0 .. n} = {0.. n - 1} \<union> {n}" by auto
|
|
555 |
from n have d: "{0.. n - 1} \<inter> {n} = {}" by auto
|
|
556 |
have f: "finite {0.. n - 1}" "finite {n}" by auto
|
|
557 |
show "setsum ?f {0..n} = 0"
|
|
558 |
unfolding th1
|
|
559 |
apply (simp add: setsum_Un_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
|
|
560 |
unfolding th2
|
|
561 |
by(simp add: setsum_delta)
|
|
562 |
qed
|
|
563 |
|
|
564 |
section{* Formal Derivatives, and the McLauren theorem around 0*}
|
|
565 |
|
|
566 |
definition "fps_deriv f = Abs_fps (\<lambda>n. of_nat (n + 1) * f $ (n + 1))"
|
|
567 |
|
|
568 |
lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n+1)" by (simp add: fps_deriv_def)
|
|
569 |
|
|
570 |
lemma fps_deriv_linear[simp]: "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) = fps_const a * fps_deriv f + fps_const b * fps_deriv g"
|
|
571 |
unfolding fps_eq_iff fps_add_nth fps_const_mult_left fps_deriv_nth by (simp add: ring_simps)
|
|
572 |
|
|
573 |
lemma fps_deriv_mult[simp]:
|
|
574 |
fixes f :: "('a :: comm_ring_1) fps"
|
|
575 |
shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g"
|
|
576 |
proof-
|
|
577 |
let ?D = "fps_deriv"
|
|
578 |
{fix n::nat
|
|
579 |
let ?Zn = "{0 ..n}"
|
|
580 |
let ?Zn1 = "{0 .. n + 1}"
|
|
581 |
let ?f = "\<lambda>i. i + 1"
|
|
582 |
have fi: "inj_on ?f {0..n}" by (simp add: inj_on_def)
|
|
583 |
have eq: "{1.. n+1} = ?f ` {0..n}" by auto
|
|
584 |
let ?g = "\<lambda>i. of_nat (i+1) * g $ (i+1) * f $ (n - i) +
|
|
585 |
of_nat (i+1)* f $ (i+1) * g $ (n - i)"
|
|
586 |
let ?h = "\<lambda>i. of_nat i * g $ i * f $ ((n+1) - i) +
|
|
587 |
of_nat i* f $ i * g $ ((n + 1) - i)"
|
|
588 |
{fix k assume k: "k \<in> {0..n}"
|
|
589 |
have "?h (k + 1) = ?g k" using k by auto}
|
|
590 |
note th0 = this
|
|
591 |
have eq': "{0..n +1}- {1 .. n+1} = {0}" by auto
|
|
592 |
have s0: "setsum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 = setsum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
|
|
593 |
apply (rule setsum_reindex_cong[where f="\<lambda>i. n + 1 - i"])
|
|
594 |
apply (simp add: inj_on_def Ball_def)
|
|
595 |
apply presburger
|
|
596 |
apply (rule set_ext)
|
|
597 |
apply (presburger add: image_iff)
|
|
598 |
by simp
|
|
599 |
have s1: "setsum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 = setsum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
|
|
600 |
apply (rule setsum_reindex_cong[where f="\<lambda>i. n + 1 - i"])
|
|
601 |
apply (simp add: inj_on_def Ball_def)
|
|
602 |
apply presburger
|
|
603 |
apply (rule set_ext)
|
|
604 |
apply (presburger add: image_iff)
|
|
605 |
by simp
|
|
606 |
have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n" by (simp only: mult_commute)
|
|
607 |
also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
|
|
608 |
by (simp add: fps_mult_nth setsum_addf[symmetric])
|
|
609 |
also have "\<dots> = setsum ?h {1..n+1}"
|
|
610 |
using th0 setsum_reindex_cong[OF fi eq, of "?g" "?h"] by auto
|
|
611 |
also have "\<dots> = setsum ?h {0..n+1}"
|
|
612 |
apply (rule setsum_mono_zero_left)
|
|
613 |
apply simp
|
|
614 |
apply (simp add: subset_eq)
|
|
615 |
unfolding eq'
|
|
616 |
by simp
|
|
617 |
also have "\<dots> = (fps_deriv (f * g)) $ n"
|
|
618 |
apply (simp only: fps_deriv_nth fps_mult_nth setsum_addf)
|
|
619 |
unfolding s0 s1
|
|
620 |
unfolding setsum_addf[symmetric] setsum_right_distrib
|
|
621 |
apply (rule setsum_cong2)
|
|
622 |
by (auto simp add: of_nat_diff ring_simps)
|
|
623 |
finally have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" .}
|
|
624 |
then show ?thesis unfolding fps_eq_iff by auto
|
|
625 |
qed
|
|
626 |
|
|
627 |
lemma fps_deriv_neg[simp]: "fps_deriv (- (f:: ('a:: comm_ring_1) fps)) = - (fps_deriv f)"
|
|
628 |
by (simp add: fps_uminus_def fps_eq_iff fps_deriv_def fps_nth_def)
|
|
629 |
lemma fps_deriv_add[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) + g) = fps_deriv f + fps_deriv g"
|
|
630 |
using fps_deriv_linear[of 1 f 1 g] by simp
|
|
631 |
|
|
632 |
lemma fps_deriv_sub[simp]: "fps_deriv ((f:: ('a::comm_ring_1) fps) - g) = fps_deriv f - fps_deriv g"
|
|
633 |
unfolding diff_minus by simp
|
|
634 |
|
|
635 |
lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
|
|
636 |
by (simp add: fps_deriv_def fps_const_def fps_zero_def)
|
|
637 |
|
|
638 |
lemma fps_deriv_mult_const_left[simp]: "fps_deriv (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_deriv f"
|
|
639 |
by simp
|
|
640 |
|
|
641 |
lemma fps_deriv_0[simp]: "fps_deriv 0 = 0"
|
|
642 |
by (simp add: fps_deriv_def fps_eq_iff)
|
|
643 |
|
|
644 |
lemma fps_deriv_1[simp]: "fps_deriv 1 = 0"
|
|
645 |
by (simp add: fps_deriv_def fps_eq_iff )
|
|
646 |
|
|
647 |
lemma fps_deriv_mult_const_right[simp]: "fps_deriv (f * fps_const (c::'a::comm_ring_1)) = fps_deriv f * fps_const c"
|
|
648 |
by simp
|
|
649 |
|
|
650 |
lemma fps_deriv_setsum: "fps_deriv (setsum f S) = setsum (\<lambda>i. fps_deriv (f i :: ('a::comm_ring_1) fps)) S"
|
|
651 |
proof-
|
|
652 |
{assume "\<not> finite S" hence ?thesis by simp}
|
|
653 |
moreover
|
|
654 |
{assume fS: "finite S"
|
|
655 |
have ?thesis by (induct rule: finite_induct[OF fS], simp_all)}
|
|
656 |
ultimately show ?thesis by blast
|
|
657 |
qed
|
|
658 |
|
|
659 |
lemma fps_deriv_eq_0_iff[simp]: "fps_deriv f = 0 \<longleftrightarrow> (f = fps_const (f$0 :: 'a::{idom,semiring_char_0}))"
|
|
660 |
proof-
|
|
661 |
{assume "f= fps_const (f$0)" hence "fps_deriv f = fps_deriv (fps_const (f$0))" by simp
|
|
662 |
hence "fps_deriv f = 0" by simp }
|
|
663 |
moreover
|
|
664 |
{assume z: "fps_deriv f = 0"
|
|
665 |
hence "\<forall>n. (fps_deriv f)$n = 0" by simp
|
|
666 |
hence "\<forall>n. f$(n+1) = 0" by (simp del: of_nat_Suc of_nat_add One_nat_def)
|
|
667 |
hence "f = fps_const (f$0)"
|
|
668 |
apply (clarsimp simp add: fps_eq_iff fps_const_def)
|
|
669 |
apply (erule_tac x="n - 1" in allE)
|
|
670 |
by simp}
|
|
671 |
ultimately show ?thesis by blast
|
|
672 |
qed
|
|
673 |
|
|
674 |
lemma fps_deriv_eq_iff:
|
|
675 |
fixes f:: "('a::{idom,semiring_char_0}) fps"
|
|
676 |
shows "fps_deriv f = fps_deriv g \<longleftrightarrow> (f = fps_const(f$0 - g$0) + g)"
|
|
677 |
proof-
|
|
678 |
have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0" by simp
|
|
679 |
also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f-g)$0)" unfolding fps_deriv_eq_0_iff ..
|
|
680 |
finally show ?thesis by (simp add: ring_simps)
|
|
681 |
qed
|
|
682 |
|
|
683 |
lemma fps_deriv_eq_iff_ex: "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>(c::'a::{idom,semiring_char_0}). f = fps_const c + g)"
|
|
684 |
apply auto unfolding fps_deriv_eq_iff by blast
|
|
685 |
|
|
686 |
|
|
687 |
fun fps_nth_deriv :: "nat \<Rightarrow> ('a::semiring_1) fps \<Rightarrow> 'a fps" where
|
|
688 |
"fps_nth_deriv 0 f = f"
|
|
689 |
| "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)"
|
|
690 |
|
|
691 |
lemma fps_nth_deriv_commute: "fps_nth_deriv (Suc n) f = fps_deriv (fps_nth_deriv n f)"
|
|
692 |
by (induct n arbitrary: f, auto)
|
|
693 |
|
|
694 |
lemma fps_nth_deriv_linear[simp]: "fps_nth_deriv n (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) = fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g"
|
|
695 |
by (induct n arbitrary: f g, auto simp add: fps_nth_deriv_commute)
|
|
696 |
|
|
697 |
lemma fps_nth_deriv_neg[simp]: "fps_nth_deriv n (- (f:: ('a:: comm_ring_1) fps)) = - (fps_nth_deriv n f)"
|
|
698 |
by (induct n arbitrary: f, simp_all)
|
|
699 |
|
|
700 |
lemma fps_nth_deriv_add[simp]: "fps_nth_deriv n ((f:: ('a::comm_ring_1) fps) + g) = fps_nth_deriv n f + fps_nth_deriv n g"
|
|
701 |
using fps_nth_deriv_linear[of n 1 f 1 g] by simp
|
|
702 |
|
|
703 |
lemma fps_nth_deriv_sub[simp]: "fps_nth_deriv n ((f:: ('a::comm_ring_1) fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
|
|
704 |
unfolding diff_minus fps_nth_deriv_add by simp
|
|
705 |
|
|
706 |
lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
|
|
707 |
by (induct n, simp_all )
|
|
708 |
|
|
709 |
lemma fps_nth_deriv_1[simp]: "fps_nth_deriv n 1 = (if n = 0 then 1 else 0)"
|
|
710 |
by (induct n, simp_all )
|
|
711 |
|
|
712 |
lemma fps_nth_deriv_const[simp]: "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)"
|
|
713 |
by (cases n, simp_all)
|
|
714 |
|
|
715 |
lemma fps_nth_deriv_mult_const_left[simp]: "fps_nth_deriv n (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_nth_deriv n f"
|
|
716 |
using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp
|
|
717 |
|
|
718 |
lemma fps_nth_deriv_mult_const_right[simp]: "fps_nth_deriv n (f * fps_const (c::'a::comm_ring_1)) = fps_nth_deriv n f * fps_const c"
|
|
719 |
using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult_commute)
|
|
720 |
|
|
721 |
lemma fps_nth_deriv_setsum: "fps_nth_deriv n (setsum f S) = setsum (\<lambda>i. fps_nth_deriv n (f i :: ('a::comm_ring_1) fps)) S"
|
|
722 |
proof-
|
|
723 |
{assume "\<not> finite S" hence ?thesis by simp}
|
|
724 |
moreover
|
|
725 |
{assume fS: "finite S"
|
|
726 |
have ?thesis by (induct rule: finite_induct[OF fS], simp_all)}
|
|
727 |
ultimately show ?thesis by blast
|
|
728 |
qed
|
|
729 |
|
|
730 |
lemma fps_deriv_maclauren_0: "(fps_nth_deriv k (f:: ('a::comm_semiring_1) fps)) $ 0 = of_nat (fact k) * f$(k)"
|
|
731 |
by (induct k arbitrary: f) (auto simp add: ring_simps of_nat_mult)
|
|
732 |
|
|
733 |
section {* Powers*}
|
|
734 |
|
|
735 |
instantiation fps :: (semiring_1) power
|
|
736 |
begin
|
|
737 |
|
|
738 |
fun fps_pow :: "nat \<Rightarrow> 'a fps \<Rightarrow> 'a fps" where
|
|
739 |
"fps_pow 0 f = 1"
|
|
740 |
| "fps_pow (Suc n) f = f * fps_pow n f"
|
|
741 |
|
|
742 |
definition fps_power_def: "power (f::'a fps) n = fps_pow n f"
|
|
743 |
instance ..
|
|
744 |
end
|
|
745 |
|
|
746 |
instantiation fps :: (comm_ring_1) recpower
|
|
747 |
begin
|
|
748 |
instance
|
|
749 |
apply (intro_classes)
|
|
750 |
by (simp_all add: fps_power_def)
|
|
751 |
end
|
|
752 |
|
|
753 |
lemma eq_neg_iff_add_eq_0: "(a::'a::ring) = -b \<longleftrightarrow> a + b = 0"
|
|
754 |
proof-
|
|
755 |
{assume "a = -b" hence "b + a = b + -b" by simp
|
|
756 |
hence "a + b = 0" by (simp add: ring_simps)}
|
|
757 |
moreover
|
|
758 |
{assume "a + b = 0" hence "a + b - b = -b" by simp
|
|
759 |
hence "a = -b" by simp}
|
|
760 |
ultimately show ?thesis by blast
|
|
761 |
qed
|
|
762 |
|
|
763 |
lemma fps_sqrare_eq_iff: "(a:: 'a::idom fps)^ 2 = b^2 \<longleftrightarrow> (a = b \<or> a = -b)"
|
|
764 |
proof-
|
|
765 |
{assume "a = b \<or> a = -b" hence "a^2 = b^2" by auto}
|
|
766 |
moreover
|
|
767 |
{assume "a^2 = b^2 "
|
|
768 |
hence "a^2 - b^2 = 0" by simp
|
|
769 |
hence "(a-b) * (a+b) = 0" by (simp add: power2_eq_square ring_simps)
|
|
770 |
hence "a = b \<or> a = -b" by (simp add: eq_neg_iff_add_eq_0)}
|
|
771 |
ultimately show ?thesis by blast
|
|
772 |
qed
|
|
773 |
|
|
774 |
lemma fps_power_zeroth_eq_one: "a$0 =1 \<Longrightarrow> a^n $ 0 = (1::'a::semiring_1)"
|
|
775 |
by (induct n, auto simp add: fps_power_def fps_times_def fps_nth_def fps_one_def)
|
|
776 |
|
|
777 |
lemma fps_power_first_eq: "(a:: 'a::comm_ring_1 fps)$0 =1 \<Longrightarrow> a^n $ 1 = of_nat n * a$1"
|
|
778 |
proof(induct n)
|
|
779 |
case 0 thus ?case by (simp add: fps_power_def)
|
|
780 |
next
|
|
781 |
case (Suc n)
|
|
782 |
note h = Suc.hyps[OF `a$0 = 1`]
|
|
783 |
show ?case unfolding power_Suc fps_mult_nth
|
|
784 |
using h `a$0 = 1` fps_power_zeroth_eq_one[OF `a$0=1`] by (simp add: ring_simps)
|
|
785 |
qed
|
|
786 |
|
|
787 |
lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \<Longrightarrow> a^n $ 0 = 1"
|
|
788 |
by (induct n, auto simp add: fps_power_def fps_mult_nth)
|
|
789 |
|
|
790 |
lemma startsby_zero_power:"a $0 = (0::'a::comm_ring_1) \<Longrightarrow> n > 0 \<Longrightarrow> a^n $0 = 0"
|
|
791 |
by (induct n, auto simp add: fps_power_def fps_mult_nth)
|
|
792 |
|
|
793 |
lemma startsby_power:"a $0 = (v::'a::{comm_ring_1, recpower}) \<Longrightarrow> a^n $0 = v^n"
|
|
794 |
by (induct n, auto simp add: fps_power_def fps_mult_nth power_Suc)
|
|
795 |
|
|
796 |
lemma startsby_zero_power_iff[simp]:
|
|
797 |
"a^n $0 = (0::'a::{idom, recpower}) \<longleftrightarrow> (n \<noteq> 0 \<and> a$0 = 0)"
|
|
798 |
apply (rule iffI)
|
|
799 |
apply (induct n, auto simp add: power_Suc fps_mult_nth)
|
|
800 |
by (rule startsby_zero_power, simp_all)
|
|
801 |
|
|
802 |
lemma startsby_zero_power_prefix:
|
|
803 |
assumes a0: "a $0 = (0::'a::idom)"
|
|
804 |
shows "\<forall>n < k. a ^ k $ n = 0"
|
|
805 |
using a0
|
|
806 |
proof(induct k rule: nat_less_induct)
|
|
807 |
fix k assume H: "\<forall>m<k. a $0 = 0 \<longrightarrow> (\<forall>n<m. a ^ m $ n = 0)" and a0: "a $0 = (0\<Colon>'a)"
|
|
808 |
let ?ths = "\<forall>m<k. a ^ k $ m = 0"
|
|
809 |
{assume "k = 0" then have ?ths by simp}
|
|
810 |
moreover
|
|
811 |
{fix l assume k: "k = Suc l"
|
|
812 |
{fix m assume mk: "m < k"
|
|
813 |
{assume "m=0" hence "a^k $ m = 0" using startsby_zero_power[of a k] k a0
|
|
814 |
by simp}
|
|
815 |
moreover
|
|
816 |
{assume m0: "m \<noteq> 0"
|
|
817 |
have "a ^k $ m = (a^l * a) $m" by (simp add: k power_Suc mult_commute)
|
|
818 |
also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))" by (simp add: fps_mult_nth)
|
|
819 |
also have "\<dots> = 0" apply (rule setsum_0')
|
|
820 |
apply auto
|
|
821 |
apply (case_tac "aa = m")
|
|
822 |
using a0
|
|
823 |
apply simp
|
|
824 |
apply (rule H[rule_format])
|
|
825 |
using a0 k mk by auto
|
|
826 |
finally have "a^k $ m = 0" .}
|
|
827 |
ultimately have "a^k $ m = 0" by blast}
|
|
828 |
hence ?ths by blast}
|
|
829 |
ultimately show ?ths by (cases k, auto)
|
|
830 |
qed
|
|
831 |
|
|
832 |
lemma startsby_zero_setsum_depends:
|
|
833 |
assumes a0: "a $0 = (0::'a::idom)" and kn: "n \<ge> k"
|
|
834 |
shows "setsum (\<lambda>i. (a ^ i)$k) {0 .. n} = setsum (\<lambda>i. (a ^ i)$k) {0 .. k}"
|
|
835 |
apply (rule setsum_mono_zero_right)
|
|
836 |
using kn apply auto
|
|
837 |
apply (rule startsby_zero_power_prefix[rule_format, OF a0])
|
|
838 |
by arith
|
|
839 |
|
|
840 |
lemma startsby_zero_power_nth_same: assumes a0: "a$0 = (0::'a::{recpower, idom})"
|
|
841 |
shows "a^n $ n = (a$1) ^ n"
|
|
842 |
proof(induct n)
|
|
843 |
case 0 thus ?case by (simp add: power_0)
|
|
844 |
next
|
|
845 |
case (Suc n)
|
|
846 |
have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)" by (simp add: ring_simps power_Suc)
|
|
847 |
also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}" by (simp add: fps_mult_nth)
|
|
848 |
also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
|
|
849 |
apply (rule setsum_mono_zero_right)
|
|
850 |
apply simp
|
|
851 |
apply clarsimp
|
|
852 |
apply clarsimp
|
|
853 |
apply (rule startsby_zero_power_prefix[rule_format, OF a0])
|
|
854 |
apply arith
|
|
855 |
done
|
|
856 |
also have "\<dots> = a^n $ n * a$1" using a0 by simp
|
|
857 |
finally show ?case using Suc.hyps by (simp add: power_Suc)
|
|
858 |
qed
|
|
859 |
|
|
860 |
lemma fps_inverse_power:
|
|
861 |
fixes a :: "('a::{field, recpower}) fps"
|
|
862 |
shows "inverse (a^n) = inverse a ^ n"
|
|
863 |
proof-
|
|
864 |
{assume a0: "a$0 = 0"
|
|
865 |
hence eq: "inverse a = 0" by (simp add: fps_inverse_def)
|
|
866 |
{assume "n = 0" hence ?thesis by simp}
|
|
867 |
moreover
|
|
868 |
{assume n: "n > 0"
|
|
869 |
from startsby_zero_power[OF a0 n] eq a0 n have ?thesis
|
|
870 |
by (simp add: fps_inverse_def)}
|
|
871 |
ultimately have ?thesis by blast}
|
|
872 |
moreover
|
|
873 |
{assume a0: "a$0 \<noteq> 0"
|
|
874 |
have ?thesis
|
|
875 |
apply (rule fps_inverse_unique)
|
|
876 |
apply (simp add: a0)
|
|
877 |
unfolding power_mult_distrib[symmetric]
|
|
878 |
apply (rule ssubst[where t = "a * inverse a" and s= 1])
|
|
879 |
apply simp_all
|
|
880 |
apply (subst mult_commute)
|
|
881 |
by (rule inverse_mult_eq_1[OF a0])}
|
|
882 |
ultimately show ?thesis by blast
|
|
883 |
qed
|
|
884 |
|
|
885 |
lemma fps_deriv_power: "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a:: comm_ring_1) * fps_deriv a * a ^ (n - 1)"
|
|
886 |
apply (induct n, auto simp add: power_Suc ring_simps fps_const_add[symmetric] simp del: fps_const_add)
|
|
887 |
by (case_tac n, auto simp add: power_Suc ring_simps)
|
|
888 |
|
|
889 |
lemma fps_inverse_deriv:
|
|
890 |
fixes a:: "('a :: field) fps"
|
|
891 |
assumes a0: "a$0 \<noteq> 0"
|
|
892 |
shows "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2"
|
|
893 |
proof-
|
|
894 |
from inverse_mult_eq_1[OF a0]
|
|
895 |
have "fps_deriv (inverse a * a) = 0" by simp
|
|
896 |
hence "inverse a * fps_deriv a + fps_deriv (inverse a) * a = 0" by simp
|
|
897 |
hence "inverse a * (inverse a * fps_deriv a + fps_deriv (inverse a) * a) = 0" by simp
|
|
898 |
with inverse_mult_eq_1[OF a0]
|
|
899 |
have "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) = 0"
|
|
900 |
unfolding power2_eq_square
|
|
901 |
apply (simp add: ring_simps)
|
|
902 |
by (simp add: mult_assoc[symmetric])
|
|
903 |
hence "inverse a ^ 2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * inverse a ^ 2 = 0 - fps_deriv a * inverse a ^ 2"
|
|
904 |
by simp
|
|
905 |
then show "fps_deriv (inverse a) = - fps_deriv a * inverse a ^ 2" by (simp add: ring_simps)
|
|
906 |
qed
|
|
907 |
|
|
908 |
lemma fps_inverse_mult:
|
|
909 |
fixes a::"('a :: field) fps"
|
|
910 |
shows "inverse (a * b) = inverse a * inverse b"
|
|
911 |
proof-
|
|
912 |
{assume a0: "a$0 = 0" hence ab0: "(a*b)$0 = 0" by (simp add: fps_mult_nth)
|
|
913 |
from a0 ab0 have th: "inverse a = 0" "inverse (a*b) = 0" by simp_all
|
|
914 |
have ?thesis unfolding th by simp}
|
|
915 |
moreover
|
|
916 |
{assume b0: "b$0 = 0" hence ab0: "(a*b)$0 = 0" by (simp add: fps_mult_nth)
|
|
917 |
from b0 ab0 have th: "inverse b = 0" "inverse (a*b) = 0" by simp_all
|
|
918 |
have ?thesis unfolding th by simp}
|
|
919 |
moreover
|
|
920 |
{assume a0: "a$0 \<noteq> 0" and b0: "b$0 \<noteq> 0"
|
|
921 |
from a0 b0 have ab0:"(a*b) $ 0 \<noteq> 0" by (simp add: fps_mult_nth)
|
|
922 |
from inverse_mult_eq_1[OF ab0]
|
|
923 |
have "inverse (a*b) * (a*b) * inverse a * inverse b = 1 * inverse a * inverse b" by simp
|
|
924 |
then have "inverse (a*b) * (inverse a * a) * (inverse b * b) = inverse a * inverse b"
|
|
925 |
by (simp add: ring_simps)
|
|
926 |
then have ?thesis using inverse_mult_eq_1[OF a0] inverse_mult_eq_1[OF b0] by simp}
|
|
927 |
ultimately show ?thesis by blast
|
|
928 |
qed
|
|
929 |
|
|
930 |
lemma fps_inverse_deriv':
|
|
931 |
fixes a:: "('a :: field) fps"
|
|
932 |
assumes a0: "a$0 \<noteq> 0"
|
|
933 |
shows "fps_deriv (inverse a) = - fps_deriv a / a ^ 2"
|
|
934 |
using fps_inverse_deriv[OF a0]
|
|
935 |
unfolding power2_eq_square fps_divide_def
|
|
936 |
fps_inverse_mult by simp
|
|
937 |
|
|
938 |
lemma inverse_mult_eq_1': assumes f0: "f$0 \<noteq> (0::'a::field)"
|
|
939 |
shows "f * inverse f= 1"
|
|
940 |
by (metis mult_commute inverse_mult_eq_1 f0)
|
|
941 |
|
|
942 |
lemma fps_divide_deriv: fixes a:: "('a :: field) fps"
|
|
943 |
assumes a0: "b$0 \<noteq> 0"
|
|
944 |
shows "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b ^ 2"
|
|
945 |
using fps_inverse_deriv[OF a0]
|
|
946 |
by (simp add: fps_divide_def ring_simps power2_eq_square fps_inverse_mult inverse_mult_eq_1'[OF a0])
|
|
947 |
|
|
948 |
section{* The eXtractor series X*}
|
|
949 |
|
|
950 |
lemma minus_one_power_iff: "(- (1::'a :: {recpower, comm_ring_1})) ^ n = (if even n then 1 else - 1)"
|
|
951 |
by (induct n, auto)
|
|
952 |
|
|
953 |
definition "X = Abs_fps (\<lambda>n. if n = 1 then 1 else 0)"
|
|
954 |
|
|
955 |
lemma fps_inverse_gp': "inverse (Abs_fps(\<lambda>n. (1::'a::field)))
|
|
956 |
= 1 - X"
|
|
957 |
by (simp add: fps_inverse_gp fps_eq_iff X_def fps_minus_def fps_one_def)
|
|
958 |
|
|
959 |
lemma X_mult_nth[simp]: "(X * (f :: ('a::semiring_1) fps)) $n = (if n = 0 then 0 else f $ (n - 1))"
|
|
960 |
proof-
|
|
961 |
{assume n: "n \<noteq> 0"
|
|
962 |
have fN: "finite {0 .. n}" by simp
|
|
963 |
have "(X * f) $n = (\<Sum>i = 0..n. X $ i * f $ (n - i))" by (simp add: fps_mult_nth)
|
|
964 |
also have "\<dots> = f $ (n - 1)"
|
|
965 |
using n by (simp add: X_def cond_value_iff cond_application_beta setsum_delta[OF fN]
|
|
966 |
del: One_nat_def cong del: if_weak_cong)
|
|
967 |
finally have ?thesis using n by simp }
|
|
968 |
moreover
|
|
969 |
{assume n: "n=0" hence ?thesis by (simp add: fps_mult_nth X_def)}
|
|
970 |
ultimately show ?thesis by blast
|
|
971 |
qed
|
|
972 |
|
|
973 |
lemma X_mult_right_nth[simp]: "((f :: ('a::comm_semiring_1) fps) * X) $n = (if n = 0 then 0 else f $ (n - 1))"
|
|
974 |
by (metis X_mult_nth mult_commute)
|
|
975 |
|
|
976 |
lemma X_power_iff: "X^k = Abs_fps (\<lambda>n. if n = k then (1::'a::comm_ring_1) else 0)"
|
|
977 |
proof(induct k)
|
|
978 |
case 0 thus ?case by (simp add: X_def fps_power_def fps_eq_iff)
|
|
979 |
next
|
|
980 |
case (Suc k)
|
|
981 |
{fix m
|
|
982 |
have "(X^Suc k) $ m = (if m = 0 then (0::'a) else (X^k) $ (m - 1))"
|
|
983 |
by (simp add: power_Suc del: One_nat_def)
|
|
984 |
then have "(X^Suc k) $ m = (if m = Suc k then (1::'a) else 0)"
|
|
985 |
using Suc.hyps by (auto cong del: if_weak_cong)}
|
|
986 |
then show ?case by (simp add: fps_eq_iff)
|
|
987 |
qed
|
|
988 |
|
|
989 |
lemma X_power_mult_nth: "(X^k * (f :: ('a::comm_ring_1) fps)) $n = (if n < k then 0 else f $ (n - k))"
|
|
990 |
apply (induct k arbitrary: n)
|
|
991 |
apply (simp)
|
|
992 |
unfolding power_Suc mult_assoc
|
|
993 |
by (case_tac n, auto)
|
|
994 |
|
|
995 |
lemma X_power_mult_right_nth: "((f :: ('a::comm_ring_1) fps) * X^k) $n = (if n < k then 0 else f $ (n - k))"
|
|
996 |
by (metis X_power_mult_nth mult_commute)
|
|
997 |
lemma fps_deriv_X[simp]: "fps_deriv X = 1"
|
|
998 |
by (simp add: fps_deriv_def X_def fps_eq_iff)
|
|
999 |
|
|
1000 |
lemma fps_nth_deriv_X[simp]: "fps_nth_deriv n X = (if n = 0 then X else if n=1 then 1 else 0)"
|
|
1001 |
by (cases "n", simp_all)
|
|
1002 |
|
|
1003 |
lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)" by (simp add: X_def)
|
|
1004 |
lemma X_power_nth[simp]: "(X^k) $n = (if n = k then 1 else (0::'a::comm_ring_1))"
|
|
1005 |
by (simp add: X_power_iff)
|
|
1006 |
|
|
1007 |
lemma fps_inverse_X_plus1:
|
|
1008 |
"inverse (1 + X) = Abs_fps (\<lambda>n. (- (1::'a::{recpower, field})) ^ n)" (is "_ = ?r")
|
|
1009 |
proof-
|
|
1010 |
have eq: "(1 + X) * ?r = 1"
|
|
1011 |
unfolding minus_one_power_iff
|
|
1012 |
apply (auto simp add: ring_simps fps_eq_iff)
|
|
1013 |
by presburger+
|
|
1014 |
show ?thesis by (auto simp add: eq intro: fps_inverse_unique)
|
|
1015 |
qed
|
|
1016 |
|
|
1017 |
|
|
1018 |
section{* Integration *}
|
|
1019 |
definition "fps_integral a a0 = Abs_fps (\<lambda>n. if n = 0 then a0 else (a$(n - 1) / of_nat n))"
|
|
1020 |
|
|
1021 |
lemma fps_deriv_fps_integral: "fps_deriv (fps_integral a (a0 :: 'a :: {field, ring_char_0})) = a"
|
|
1022 |
by (simp add: fps_integral_def fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
|
|
1023 |
|
|
1024 |
lemma fps_integral_linear: "fps_integral (fps_const (a::'a::{field, ring_char_0}) * f + fps_const b * g) (a*a0 + b*b0) = fps_const a * fps_integral f a0 + fps_const b * fps_integral g b0" (is "?l = ?r")
|
|
1025 |
proof-
|
|
1026 |
have "fps_deriv ?l = fps_deriv ?r" by (simp add: fps_deriv_fps_integral)
|
|
1027 |
moreover have "?l$0 = ?r$0" by (simp add: fps_integral_def)
|
|
1028 |
ultimately show ?thesis
|
|
1029 |
unfolding fps_deriv_eq_iff by auto
|
|
1030 |
qed
|
|
1031 |
|
|
1032 |
section {* Composition of FPSs *}
|
|
1033 |
definition fps_compose :: "('a::semiring_1) fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps" (infixl "oo" 55) where
|
|
1034 |
fps_compose_def: "a oo b = Abs_fps (\<lambda>n. setsum (\<lambda>i. a$i * (b^i$n)) {0..n})"
|
|
1035 |
|
|
1036 |
lemma fps_compose_nth: "(a oo b)$n = setsum (\<lambda>i. a$i * (b^i$n)) {0..n}" by (simp add: fps_compose_def)
|
|
1037 |
|
|
1038 |
lemma fps_compose_X[simp]: "a oo X = (a :: ('a :: comm_ring_1) fps)"
|
|
1039 |
by (auto simp add: fps_compose_def X_power_iff fps_eq_iff cond_application_beta cond_value_iff setsum_delta' cong del: if_weak_cong)
|
|
1040 |
|
|
1041 |
lemma fps_const_compose[simp]:
|
|
1042 |
"fps_const (a::'a::{comm_ring_1}) oo b = fps_const (a)"
|
|
1043 |
apply (auto simp add: fps_eq_iff fps_compose_nth fps_mult_nth
|
|
1044 |
cond_application_beta cond_value_iff cong del: if_weak_cong)
|
|
1045 |
by (simp add: setsum_delta )
|
|
1046 |
|
|
1047 |
lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: ('a :: comm_ring_1) fps)"
|
|
1048 |
apply (auto simp add: fps_compose_def fps_eq_iff cond_application_beta cond_value_iff setsum_delta cong del: if_weak_cong)
|
|
1049 |
apply (simp add: power_Suc)
|
|
1050 |
apply (subgoal_tac "n = 0")
|
|
1051 |
by simp_all
|
|
1052 |
|
|
1053 |
|
|
1054 |
section {* Rules from Herbert Wilf's Generatingfunctionology*}
|
|
1055 |
|
|
1056 |
subsection {* Rule 1 *}
|
|
1057 |
(* {a_{n+k}}_0^infty Corresponds to (f - setsum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
|
|
1058 |
|
|
1059 |
lemma fps_power_mult_eq_shift:
|
|
1060 |
"X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) = Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a:: field) * X^i) {0 .. k}" (is "?lhs = ?rhs")
|
|
1061 |
proof-
|
|
1062 |
{fix n:: nat
|
|
1063 |
have "?lhs $ n = (if n < Suc k then 0 else a n)"
|
|
1064 |
unfolding X_power_mult_nth by auto
|
|
1065 |
also have "\<dots> = ?rhs $ n"
|
|
1066 |
proof(induct k)
|
|
1067 |
case 0 thus ?case by (simp add: fps_setsum_nth power_Suc)
|
|
1068 |
next
|
|
1069 |
case (Suc k)
|
|
1070 |
note th = Suc.hyps[symmetric]
|
|
1071 |
have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a:: field) * X^i) {0 .. Suc k})$n = (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a:: field) * X^i) {0 .. k} - fps_const (a (Suc k)) * X^ Suc k) $ n" by (simp add: ring_simps)
|
|
1072 |
also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
|
|
1073 |
using th
|
|
1074 |
unfolding fps_sub_nth by simp
|
|
1075 |
also have "\<dots> = (if n < Suc (Suc k) then 0 else a n)"
|
|
1076 |
unfolding X_power_mult_right_nth
|
|
1077 |
apply (auto simp add: not_less fps_const_def)
|
|
1078 |
apply (rule cong[of a a, OF refl])
|
|
1079 |
by arith
|
|
1080 |
finally show ?case by simp
|
|
1081 |
qed
|
|
1082 |
finally have "?lhs $ n = ?rhs $ n" .}
|
|
1083 |
then show ?thesis by (simp add: fps_eq_iff)
|
|
1084 |
qed
|
|
1085 |
|
|
1086 |
subsection{* Rule 2*}
|
|
1087 |
|
|
1088 |
(* We can not reach the form of Wilf, but still near to it using rewrite rules*)
|
|
1089 |
(* If f reprents {a_n} and P is a polynomial, then
|
|
1090 |
P(xD) f represents {P(n) a_n}*)
|
|
1091 |
|
|
1092 |
definition "XD = op * X o fps_deriv"
|
|
1093 |
|
|
1094 |
lemma XD_add[simp]:"XD (a + b) = XD a + XD (b :: ('a::comm_ring_1) fps)"
|
|
1095 |
by (simp add: XD_def ring_simps)
|
|
1096 |
|
|
1097 |
lemma XD_mult_const[simp]:"XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * XD a"
|
|
1098 |
by (simp add: XD_def ring_simps)
|
|
1099 |
|
|
1100 |
lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) = fps_const c * XD a + fps_const d * XD (b :: ('a::comm_ring_1) fps)"
|
|
1101 |
by simp
|
|
1102 |
|
|
1103 |
|
|
1104 |
fun funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
|
|
1105 |
"funpow 0 f = id"
|
|
1106 |
| "funpow (Suc n) f = f o funpow n f"
|
|
1107 |
|
|
1108 |
lemma XDN_linear: "(funpow n XD) (fps_const c * a + fps_const d * b) = fps_const c * (funpow n XD) a + fps_const d * (funpow n XD) (b :: ('a::comm_ring_1) fps)"
|
|
1109 |
by (induct n, simp_all)
|
|
1110 |
|
|
1111 |
lemma fps_mult_X_deriv_shift: "X* fps_deriv a = Abs_fps (\<lambda>n. of_nat n* a$n)" by (simp add: fps_eq_iff)
|
|
1112 |
|
|
1113 |
lemma fps_mult_XD_shift: "funpow k XD (a:: ('a::{comm_ring_1, recpower, ring_char_0}) fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
|
|
1114 |
by (induct k arbitrary: a) (simp_all add: power_Suc XD_def fps_eq_iff ring_simps del: One_nat_def)
|
|
1115 |
|
|
1116 |
subsection{* Rule 3 is trivial and is given by fps_times_def*}
|
|
1117 |
subsection{* Rule 5 --- summation and "division" by (1 - X)*}
|
|
1118 |
|
|
1119 |
lemma fps_divide_X_minus1_setsum_lemma:
|
|
1120 |
"a = ((1::('a::comm_ring_1) fps) - X) * Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
|
|
1121 |
proof-
|
|
1122 |
let ?X = "X::('a::comm_ring_1) fps"
|
|
1123 |
let ?sa = "Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
|
|
1124 |
have th0: "\<And>i. (1 - (X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)" by simp
|
|
1125 |
{fix n:: nat
|
|
1126 |
{assume "n=0" hence "a$n = ((1 - ?X) * ?sa) $ n"
|
|
1127 |
by (simp add: fps_mult_nth)}
|
|
1128 |
moreover
|
|
1129 |
{assume n0: "n \<noteq> 0"
|
|
1130 |
then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1}\<union>{2..n} = {1..n}"
|
|
1131 |
"{0..n - 1}\<union>{n} = {0..n}"
|
|
1132 |
apply (simp_all add: expand_set_eq) by presburger+
|
|
1133 |
have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}"
|
|
1134 |
"{0..n - 1}\<inter>{n} ={}" using n0
|
|
1135 |
by (simp_all add: expand_set_eq, presburger+)
|
|
1136 |
have f: "finite {0}" "finite {1}" "finite {2 .. n}"
|
|
1137 |
"finite {0 .. n - 1}" "finite {n}" by simp_all
|
|
1138 |
have "((1 - ?X) * ?sa) $ n = setsum (\<lambda>i. (1 - ?X)$ i * ?sa $ (n - i)) {0 .. n}"
|
|
1139 |
by (simp add: fps_mult_nth)
|
|
1140 |
also have "\<dots> = a$n" unfolding th0
|
|
1141 |
unfolding setsum_Un_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
|
|
1142 |
unfolding setsum_Un_disjoint[OF f(2) f(3) d(2)]
|
|
1143 |
apply (simp)
|
|
1144 |
unfolding setsum_Un_disjoint[OF f(4,5) d(3), unfolded u(3)]
|
|
1145 |
by simp
|
|
1146 |
finally have "a$n = ((1 - ?X) * ?sa) $ n" by simp}
|
|
1147 |
ultimately have "a$n = ((1 - ?X) * ?sa) $ n" by blast}
|
|
1148 |
then show ?thesis
|
|
1149 |
unfolding fps_eq_iff by blast
|
|
1150 |
qed
|
|
1151 |
|
|
1152 |
lemma fps_divide_X_minus1_setsum:
|
|
1153 |
"a /((1::('a::field) fps) - X) = Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
|
|
1154 |
proof-
|
|
1155 |
let ?X = "1 - (X::('a::field) fps)"
|
|
1156 |
have th0: "?X $ 0 \<noteq> 0" by simp
|
|
1157 |
have "a /?X = ?X * Abs_fps (\<lambda>n\<Colon>nat. setsum (op $ a) {0..n}) * inverse ?X"
|
|
1158 |
using fps_divide_X_minus1_setsum_lemma[of a, symmetric] th0
|
|
1159 |
by (simp add: fps_divide_def mult_assoc)
|
|
1160 |
also have "\<dots> = (inverse ?X * ?X) * Abs_fps (\<lambda>n\<Colon>nat. setsum (op $ a) {0..n}) "
|
|
1161 |
by (simp add: mult_ac)
|
|
1162 |
finally show ?thesis by (simp add: inverse_mult_eq_1[OF th0])
|
|
1163 |
qed
|
|
1164 |
|
|
1165 |
subsection{* Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
|
|
1166 |
finite product of FPS, also the relvant instance of powers of a FPS*}
|
|
1167 |
|
|
1168 |
definition "natpermute n k = {l:: nat list. length l = k \<and> foldl op + 0 l = n}"
|
|
1169 |
|
|
1170 |
lemma natlist_trivial_1: "natpermute n 1 = {[n]}"
|
|
1171 |
apply (auto simp add: natpermute_def)
|
|
1172 |
apply (case_tac x, auto)
|
|
1173 |
done
|
|
1174 |
|
|
1175 |
lemma foldl_add_start0:
|
|
1176 |
"foldl op + x xs = x + foldl op + (0::nat) xs"
|
|
1177 |
apply (induct xs arbitrary: x)
|
|
1178 |
apply simp
|
|
1179 |
unfolding foldl.simps
|
|
1180 |
apply atomize
|
|
1181 |
apply (subgoal_tac "\<forall>x\<Colon>nat. foldl op + x xs = x + foldl op + (0\<Colon>nat) xs")
|
|
1182 |
apply (erule_tac x="x + a" in allE)
|
|
1183 |
apply (erule_tac x="a" in allE)
|
|
1184 |
apply simp
|
|
1185 |
apply assumption
|
|
1186 |
done
|
|
1187 |
|
|
1188 |
lemma foldl_add_append: "foldl op + (x::nat) (xs@ys) = foldl op + x xs + foldl op + 0 ys"
|
|
1189 |
apply (induct ys arbitrary: x xs)
|
|
1190 |
apply auto
|
|
1191 |
apply (subst (2) foldl_add_start0)
|
|
1192 |
apply simp
|
|
1193 |
apply (subst (2) foldl_add_start0)
|
|
1194 |
by simp
|
|
1195 |
|
|
1196 |
lemma foldl_add_setsum: "foldl op + (x::nat) xs = x + setsum (nth xs) {0..<length xs}"
|
|
1197 |
proof(induct xs arbitrary: x)
|
|
1198 |
case Nil thus ?case by simp
|
|
1199 |
next
|
|
1200 |
case (Cons a as x)
|
|
1201 |
have eq: "setsum (op ! (a#as)) {1..<length (a#as)} = setsum (op ! as) {0..<length as}"
|
|
1202 |
apply (rule setsum_reindex_cong [where f=Suc])
|
|
1203 |
by (simp_all add: inj_on_def)
|
|
1204 |
have f: "finite {0}" "finite {1 ..< length (a#as)}" by simp_all
|
|
1205 |
have d: "{0} \<inter> {1..<length (a#as)} = {}" by simp
|
|
1206 |
have seq: "{0} \<union> {1..<length(a#as)} = {0 ..<length (a#as)}" by auto
|
|
1207 |
have "foldl op + x (a#as) = x + foldl op + a as "
|
|
1208 |
apply (subst foldl_add_start0) by simp
|
|
1209 |
also have "\<dots> = x + a + setsum (op ! as) {0..<length as}" unfolding Cons.hyps by simp
|
|
1210 |
also have "\<dots> = x + setsum (op ! (a#as)) {0..<length (a#as)}"
|
|
1211 |
unfolding eq[symmetric]
|
|
1212 |
unfolding setsum_Un_disjoint[OF f d, unfolded seq]
|
|
1213 |
by simp
|
|
1214 |
finally show ?case .
|
|
1215 |
qed
|
|
1216 |
|
|
1217 |
|
|
1218 |
lemma append_natpermute_less_eq:
|
|
1219 |
assumes h: "xs@ys \<in> natpermute n k" shows "foldl op + 0 xs \<le> n" and "foldl op + 0 ys \<le> n"
|
|
1220 |
proof-
|
|
1221 |
{from h have "foldl op + 0 (xs@ ys) = n" by (simp add: natpermute_def)
|
|
1222 |
hence "foldl op + 0 xs + foldl op + 0 ys = n" unfolding foldl_add_append .}
|
|
1223 |
note th = this
|
|
1224 |
{from th show "foldl op + 0 xs \<le> n" by simp}
|
|
1225 |
{from th show "foldl op + 0 ys \<le> n" by simp}
|
|
1226 |
qed
|
|
1227 |
|
|
1228 |
lemma natpermute_split:
|
|
1229 |
assumes mn: "h \<le> k"
|
|
1230 |
shows "natpermute n k = (\<Union>m \<in>{0..n}. {l1 @ l2 |l1 l2. l1 \<in> natpermute m h \<and> l2 \<in> natpermute (n - m) (k - h)})" (is "?L = ?R" is "?L = (\<Union>m \<in>{0..n}. ?S m)")
|
|
1231 |
proof-
|
|
1232 |
{fix l assume l: "l \<in> ?R"
|
|
1233 |
from l obtain m xs ys where h: "m \<in> {0..n}" and xs: "xs \<in> natpermute m h" and ys: "ys \<in> natpermute (n - m) (k - h)" and leq: "l = xs@ys" by blast
|
|
1234 |
from xs have xs': "foldl op + 0 xs = m" by (simp add: natpermute_def)
|
|
1235 |
from ys have ys': "foldl op + 0 ys = n - m" by (simp add: natpermute_def)
|
|
1236 |
have "l \<in> ?L" using leq xs ys h
|
|
1237 |
apply simp
|
|
1238 |
apply (clarsimp simp add: natpermute_def simp del: foldl_append)
|
|
1239 |
apply (simp add: foldl_add_append[unfolded foldl_append])
|
|
1240 |
unfolding xs' ys'
|
|
1241 |
using mn xs ys
|
|
1242 |
unfolding natpermute_def by simp}
|
|
1243 |
moreover
|
|
1244 |
{fix l assume l: "l \<in> natpermute n k"
|
|
1245 |
let ?xs = "take h l"
|
|
1246 |
let ?ys = "drop h l"
|
|
1247 |
let ?m = "foldl op + 0 ?xs"
|
|
1248 |
from l have ls: "foldl op + 0 (?xs @ ?ys) = n" by (simp add: natpermute_def)
|
|
1249 |
have xs: "?xs \<in> natpermute ?m h" using l mn by (simp add: natpermute_def)
|
|
1250 |
have ys: "?ys \<in> natpermute (n - ?m) (k - h)" using l mn ls[unfolded foldl_add_append]
|
|
1251 |
by (simp add: natpermute_def)
|
|
1252 |
from ls have m: "?m \<in> {0..n}" unfolding foldl_add_append by simp
|
|
1253 |
from xs ys ls have "l \<in> ?R"
|
|
1254 |
apply auto
|
|
1255 |
apply (rule bexI[where x = "?m"])
|
|
1256 |
apply (rule exI[where x = "?xs"])
|
|
1257 |
apply (rule exI[where x = "?ys"])
|
|
1258 |
using ls l unfolding foldl_add_append
|
|
1259 |
by (auto simp add: natpermute_def)}
|
|
1260 |
ultimately show ?thesis by blast
|
|
1261 |
qed
|
|
1262 |
|
|
1263 |
lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})"
|
|
1264 |
by (auto simp add: natpermute_def)
|
|
1265 |
lemma natpermute_0'[simp]: "natpermute 0 k = (if k = 0 then {[]} else {replicate k 0})"
|
|
1266 |
apply (auto simp add: set_replicate_conv_if natpermute_def)
|
|
1267 |
apply (rule nth_equalityI)
|
|
1268 |
by simp_all
|
|
1269 |
|
|
1270 |
lemma natpermute_finite: "finite (natpermute n k)"
|
|
1271 |
proof(induct k arbitrary: n)
|
|
1272 |
case 0 thus ?case
|
|
1273 |
apply (subst natpermute_split[of 0 0, simplified])
|
|
1274 |
by (simp add: natpermute_0)
|
|
1275 |
next
|
|
1276 |
case (Suc k)
|
|
1277 |
then show ?case unfolding natpermute_split[of k "Suc k", simplified]
|
|
1278 |
apply -
|
|
1279 |
apply (rule finite_UN_I)
|
|
1280 |
apply simp
|
|
1281 |
unfolding One_nat_def[symmetric] natlist_trivial_1
|
|
1282 |
apply simp
|
|
1283 |
unfolding image_Collect[symmetric]
|
|
1284 |
unfolding Collect_def mem_def
|
|
1285 |
apply (rule finite_imageI)
|
|
1286 |
apply blast
|
|
1287 |
done
|
|
1288 |
qed
|
|
1289 |
|
|
1290 |
lemma natpermute_contain_maximal:
|
|
1291 |
"{xs \<in> natpermute n (k+1). n \<in> set xs} = UNION {0 .. k} (\<lambda>i. {(replicate (k+1) 0) [i:=n]})"
|
|
1292 |
(is "?A = ?B")
|
|
1293 |
proof-
|
|
1294 |
{fix xs assume H: "xs \<in> natpermute n (k+1)" and n: "n \<in> set xs"
|
|
1295 |
from n obtain i where i: "i \<in> {0.. k}" "xs!i = n" using H
|
|
1296 |
unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def)
|
|
1297 |
have eqs: "({0..k} - {i}) \<union> {i} = {0..k}" using i by auto
|
|
1298 |
have f: "finite({0..k} - {i})" "finite {i}" by auto
|
|
1299 |
have d: "({0..k} - {i}) \<inter> {i} = {}" using i by auto
|
|
1300 |
from H have "n = setsum (nth xs) {0..k}" apply (simp add: natpermute_def)
|
|
1301 |
unfolding foldl_add_setsum by (auto simp add: atLeastLessThanSuc_atLeastAtMost)
|
|
1302 |
also have "\<dots> = n + setsum (nth xs) ({0..k} - {i})"
|
|
1303 |
unfolding setsum_Un_disjoint[OF f d, unfolded eqs] using i by simp
|
|
1304 |
finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0" by auto
|
|
1305 |
from H have xsl: "length xs = k+1" by (simp add: natpermute_def)
|
|
1306 |
from i have i': "i < length (replicate (k+1) 0)" "i < k+1"
|
|
1307 |
unfolding length_replicate by arith+
|
|
1308 |
have "xs = replicate (k+1) 0 [i := n]"
|
|
1309 |
apply (rule nth_equalityI)
|
|
1310 |
unfolding xsl length_list_update length_replicate
|
|
1311 |
apply simp
|
|
1312 |
apply clarify
|
|
1313 |
unfolding nth_list_update[OF i'(1)]
|
|
1314 |
using i zxs
|
|
1315 |
by (case_tac "ia=i", auto simp del: replicate.simps)
|
|
1316 |
then have "xs \<in> ?B" using i by blast}
|
|
1317 |
moreover
|
|
1318 |
{fix i assume i: "i \<in> {0..k}"
|
|
1319 |
let ?xs = "replicate (k+1) 0 [i:=n]"
|
|
1320 |
have nxs: "n \<in> set ?xs"
|
|
1321 |
apply (rule set_update_memI) using i by simp
|
|
1322 |
have xsl: "length ?xs = k+1" by (simp only: length_replicate length_list_update)
|
|
1323 |
have "foldl op + 0 ?xs = setsum (nth ?xs) {0..<k+1}"
|
|
1324 |
unfolding foldl_add_setsum add_0 length_replicate length_list_update ..
|
|
1325 |
also have "\<dots> = setsum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
|
|
1326 |
apply (rule setsum_cong2) by (simp del: replicate.simps)
|
|
1327 |
also have "\<dots> = n" using i by (simp add: setsum_delta)
|
|
1328 |
finally
|
|
1329 |
have "?xs \<in> natpermute n (k+1)" using xsl unfolding natpermute_def Collect_def mem_def
|
|
1330 |
by blast
|
|
1331 |
then have "?xs \<in> ?A" using nxs by blast}
|
|
1332 |
ultimately show ?thesis by auto
|
|
1333 |
qed
|
|
1334 |
|
|
1335 |
(* The general form *)
|
|
1336 |
lemma fps_setprod_nth:
|
|
1337 |
fixes m :: nat and a :: "nat \<Rightarrow> ('a::comm_ring_1) fps"
|
|
1338 |
shows "(setprod a {0 .. m})$n = setsum (\<lambda>v. setprod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
|
|
1339 |
(is "?P m n")
|
|
1340 |
proof(induct m arbitrary: n rule: nat_less_induct)
|
|
1341 |
fix m n assume H: "\<forall>m' < m. \<forall>n. ?P m' n"
|
|
1342 |
{assume m0: "m = 0"
|
|
1343 |
hence "?P m n" apply simp
|
|
1344 |
unfolding natlist_trivial_1[where n = n, unfolded One_nat_def] by simp}
|
|
1345 |
moreover
|
|
1346 |
{fix k assume k: "m = Suc k"
|
|
1347 |
have km: "k < m" using k by arith
|
|
1348 |
have u0: "{0 .. k} \<union> {m} = {0..m}" using k apply (simp add: expand_set_eq) by presburger
|
|
1349 |
have f0: "finite {0 .. k}" "finite {m}" by auto
|
|
1350 |
have d0: "{0 .. k} \<inter> {m} = {}" using k by auto
|
|
1351 |
have "(setprod a {0 .. m}) $ n = (setprod a {0 .. k} * a m) $ n"
|
|
1352 |
unfolding setprod_Un_disjoint[OF f0 d0, unfolded u0] by simp
|
|
1353 |
also have "\<dots> = (\<Sum>i = 0..n. (\<Sum>v\<in>natpermute i (k + 1). \<Prod>j\<in>{0..k}. a j $ v ! j) * a m $ (n - i))"
|
|
1354 |
unfolding fps_mult_nth H[rule_format, OF km] ..
|
|
1355 |
also have "\<dots> = (\<Sum>v\<in>natpermute n (m + 1). \<Prod>j\<in>{0..m}. a j $ v ! j)"
|
|
1356 |
apply (simp add: k)
|
|
1357 |
unfolding natpermute_split[of m "m + 1", simplified, of n, unfolded natlist_trivial_1[unfolded One_nat_def] k]
|
|
1358 |
apply (subst setsum_UN_disjoint)
|
|
1359 |
apply simp
|
|
1360 |
apply simp
|
|
1361 |
unfolding image_Collect[symmetric]
|
|
1362 |
apply clarsimp
|
|
1363 |
apply (rule finite_imageI)
|
|
1364 |
apply (rule natpermute_finite)
|
|
1365 |
apply (clarsimp simp add: expand_set_eq)
|
|
1366 |
apply auto
|
|
1367 |
apply (rule setsum_cong2)
|
|
1368 |
unfolding setsum_left_distrib
|
|
1369 |
apply (rule sym)
|
|
1370 |
apply (rule_tac f="\<lambda>xs. xs @[n - x]" in setsum_reindex_cong)
|
|
1371 |
apply (simp add: inj_on_def)
|
|
1372 |
apply auto
|
|
1373 |
unfolding setprod_Un_disjoint[OF f0 d0, unfolded u0, unfolded k]
|
|
1374 |
apply (clarsimp simp add: natpermute_def nth_append)
|
|
1375 |
apply (rule_tac f="\<lambda>x. x * a (Suc k) $ (n - foldl op + 0 aa)" in cong[OF refl])
|
|
1376 |
apply (rule setprod_cong)
|
|
1377 |
apply simp
|
|
1378 |
apply simp
|
|
1379 |
done
|
|
1380 |
finally have "?P m n" .}
|
|
1381 |
ultimately show "?P m n " by (cases m, auto)
|
|
1382 |
qed
|
|
1383 |
|
|
1384 |
text{* The special form for powers *}
|
|
1385 |
lemma fps_power_nth_Suc:
|
|
1386 |
fixes m :: nat and a :: "('a::comm_ring_1) fps"
|
|
1387 |
shows "(a ^ Suc m)$n = setsum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m}) (natpermute n (m+1))"
|
|
1388 |
proof-
|
|
1389 |
have f: "finite {0 ..m}" by simp
|
|
1390 |
have th0: "a^Suc m = setprod (\<lambda>i. a) {0..m}" unfolding setprod_constant[OF f, of a] by simp
|
|
1391 |
show ?thesis unfolding th0 fps_setprod_nth ..
|
|
1392 |
qed
|
|
1393 |
lemma fps_power_nth:
|
|
1394 |
fixes m :: nat and a :: "('a::comm_ring_1) fps"
|
|
1395 |
shows "(a ^m)$n = (if m=0 then 1$n else setsum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
|
|
1396 |
by (cases m, simp_all add: fps_power_nth_Suc)
|
|
1397 |
|
|
1398 |
lemma fps_nth_power_0:
|
|
1399 |
fixes m :: nat and a :: "('a::{comm_ring_1, recpower}) fps"
|
|
1400 |
shows "(a ^m)$0 = (a$0) ^ m"
|
|
1401 |
proof-
|
|
1402 |
{assume "m=0" hence ?thesis by simp}
|
|
1403 |
moreover
|
|
1404 |
{fix n assume m: "m = Suc n"
|
|
1405 |
have c: "m = card {0..n}" using m by simp
|
|
1406 |
have "(a ^m)$0 = setprod (\<lambda>i. a$0) {0..n}"
|
|
1407 |
apply (simp add: m fps_power_nth del: replicate.simps)
|
|
1408 |
apply (rule setprod_cong)
|
|
1409 |
by (simp_all del: replicate.simps)
|
|
1410 |
also have "\<dots> = (a$0) ^ m"
|
|
1411 |
unfolding c by (rule setprod_constant, simp)
|
|
1412 |
finally have ?thesis .}
|
|
1413 |
ultimately show ?thesis by (cases m, auto)
|
|
1414 |
qed
|
|
1415 |
|
|
1416 |
lemma fps_compose_inj_right:
|
|
1417 |
assumes a0: "a$0 = (0::'a::{recpower,idom})"
|
|
1418 |
and a1: "a$1 \<noteq> 0"
|
|
1419 |
shows "(b oo a = c oo a) \<longleftrightarrow> b = c" (is "?lhs \<longleftrightarrow>?rhs")
|
|
1420 |
proof-
|
|
1421 |
{assume ?rhs then have "?lhs" by simp}
|
|
1422 |
moreover
|
|
1423 |
{assume h: ?lhs
|
|
1424 |
{fix n have "b$n = c$n"
|
|
1425 |
proof(induct n rule: nat_less_induct)
|
|
1426 |
fix n assume H: "\<forall>m<n. b$m = c$m"
|
|
1427 |
{assume n0: "n=0"
|
|
1428 |
from h have "(b oo a)$n = (c oo a)$n" by simp
|
|
1429 |
hence "b$n = c$n" using n0 by (simp add: fps_compose_nth)}
|
|
1430 |
moreover
|
|
1431 |
{fix n1 assume n1: "n = Suc n1"
|
|
1432 |
have f: "finite {0 .. n1}" "finite {n}" by simp_all
|
|
1433 |
have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using n1 by auto
|
|