| author | paulson <lp15@cam.ac.uk> | 
| Wed, 10 Apr 2019 13:34:55 +0100 | |
| changeset 70097 | 4005298550a6 | 
| parent 69739 | 8b47c021666e | 
| child 70113 | c8deb8ba6d05 | 
| permissions | -rw-r--r-- | 
| 56215 | 1  | 
(* Author: John Harrison and Valentina Bruno  | 
2  | 
Ported from "hol_light/Multivariate/complexes.ml" by L C Paulson  | 
|
3  | 
*)  | 
|
4  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
5  | 
section \<open>Polynomial Functions: Extremal Behaviour and Root Counts\<close>  | 
| 61560 | 6  | 
|
| 
63594
 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 
hoelzl 
parents: 
62626 
diff
changeset
 | 
7  | 
theory Poly_Roots  | 
| 56215 | 8  | 
imports Complex_Main  | 
9  | 
begin  | 
|
10  | 
||
| 69683 | 11  | 
subsection\<open>Basics about polynomial functions: extremal behaviour and root counts\<close>  | 
| 56215 | 12  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
13  | 
lemma sub_polyfun:  | 
| 56215 | 14  | 
  fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
15  | 
shows "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =  | 
| 56215 | 16  | 
(x - y) * (\<Sum>j<n. \<Sum>k= Suc j..n. a k * y^(k - Suc j) * x^j)"  | 
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
17  | 
proof -  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
18  | 
have "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =  | 
| 56215 | 19  | 
(\<Sum>i\<le>n. a i * (x^i - y^i))"  | 
| 64267 | 20  | 
by (simp add: algebra_simps sum_subtractf [symmetric])  | 
| 56215 | 21  | 
also have "... = (\<Sum>i\<le>n. a i * (x - y) * (\<Sum>j<i. y^(i - Suc j) * x^j))"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
57512 
diff
changeset
 | 
22  | 
by (simp add: power_diff_sumr2 ac_simps)  | 
| 56215 | 23  | 
also have "... = (x - y) * (\<Sum>i\<le>n. (\<Sum>j<i. a i * y^(i - Suc j) * x^j))"  | 
| 64267 | 24  | 
by (simp add: sum_distrib_left ac_simps)  | 
| 56215 | 25  | 
also have "... = (x - y) * (\<Sum>j<n. (\<Sum>i=Suc j..n. a i * y^(i - Suc j) * x^j))"  | 
| 64267 | 26  | 
by (simp add: nested_sum_swap')  | 
| 56215 | 27  | 
finally show ?thesis .  | 
28  | 
qed  | 
|
29  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
30  | 
lemma sub_polyfun_alt:  | 
| 56215 | 31  | 
  fixes x :: "'a::{comm_ring,monoid_mult}"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
32  | 
shows "(\<Sum>i\<le>n. a i * x^i) - (\<Sum>i\<le>n. a i * y^i) =  | 
| 56215 | 33  | 
(x - y) * (\<Sum>j<n. \<Sum>k<n-j. a (j+k+1) * y^k * x^j)"  | 
34  | 
proof -  | 
|
35  | 
  { fix j
 | 
|
36  | 
have "(\<Sum>k = Suc j..n. a k * y^(k - Suc j) * x^j) =  | 
|
37  | 
(\<Sum>k <n - j. a (Suc (j + k)) * y^k * x^j)"  | 
|
| 64267 | 38  | 
by (rule sum.reindex_bij_witness[where i="\<lambda>i. i + Suc j" and j="\<lambda>i. i - Suc j"]) auto }  | 
| 56215 | 39  | 
then show ?thesis  | 
40  | 
by (simp add: sub_polyfun)  | 
|
41  | 
qed  | 
|
42  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
43  | 
lemma polyfun_linear_factor:  | 
| 56215 | 44  | 
  fixes a :: "'a::{comm_ring,monoid_mult}"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
45  | 
shows "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) =  | 
| 56215 | 46  | 
(z-a) * (\<Sum>i<n. b i * z^i) + (\<Sum>i\<le>n. c i * a^i)"  | 
47  | 
proof -  | 
|
48  | 
  { fix z
 | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
49  | 
have "(\<Sum>i\<le>n. c i * z^i) - (\<Sum>i\<le>n. c i * a^i) =  | 
| 56215 | 50  | 
(z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)"  | 
| 64267 | 51  | 
by (simp add: sub_polyfun sum_distrib_right)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
52  | 
then have "(\<Sum>i\<le>n. c i * z^i) =  | 
| 56215 | 53  | 
(z - a) * (\<Sum>j<n. (\<Sum>k = Suc j..n. c k * a^(k - Suc j)) * z^j)  | 
54  | 
+ (\<Sum>i\<le>n. c i * a^i)"  | 
|
55  | 
by (simp add: algebra_simps) }  | 
|
56  | 
then show ?thesis  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
57  | 
by (intro exI allI)  | 
| 56215 | 58  | 
qed  | 
59  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
60  | 
lemma polyfun_linear_factor_root:  | 
| 56215 | 61  | 
  fixes a :: "'a::{comm_ring,monoid_mult}"
 | 
62  | 
assumes "(\<Sum>i\<le>n. c i * a^i) = 0"  | 
|
63  | 
shows "\<exists>b. \<forall>z. (\<Sum>i\<le>n. c i * z^i) = (z-a) * (\<Sum>i<n. b i * z^i)"  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
64  | 
using polyfun_linear_factor [of c n a] assms  | 
| 
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
65  | 
by simp  | 
| 56215 | 66  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
67  | 
lemma adhoc_norm_triangle: "a + norm(y) \<le> b ==> norm(x) \<le> a ==> norm(x + y) \<le> b"  | 
| 56215 | 68  | 
by (metis norm_triangle_mono order.trans order_refl)  | 
69  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
70  | 
proposition polyfun_extremal_lemma:  | 
| 56215 | 71  | 
fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"  | 
72  | 
assumes "e > 0"  | 
|
73  | 
shows "\<exists>M. \<forall>z. M \<le> norm z \<longrightarrow> norm(\<Sum>i\<le>n. c i * z^i) \<le> e * norm(z) ^ Suc n"  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
74  | 
proof (induction n)  | 
| 56215 | 75  | 
case 0  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
76  | 
show ?case  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57129 
diff
changeset
 | 
77  | 
by (rule exI [where x="norm (c 0) / e"]) (auto simp: mult.commute pos_divide_le_eq assms)  | 
| 56215 | 78  | 
next  | 
79  | 
case (Suc n)  | 
|
80  | 
then obtain M where M: "\<forall>z. M \<le> norm z \<longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" ..  | 
|
81  | 
show ?case  | 
|
82  | 
proof (rule exI [where x="max 1 (max M ((e + norm(c(Suc n))) / e))"], clarify)  | 
|
83  | 
fix z::'a  | 
|
84  | 
assume "max 1 (max M ((e + norm (c (Suc n))) / e)) \<le> norm z"  | 
|
85  | 
then have norm1: "0 < norm z" "M \<le> norm z" "(e + norm (c (Suc n))) / e \<le> norm z"  | 
|
86  | 
by auto  | 
|
87  | 
then have norm2: "(e + norm (c (Suc n))) \<le> e * norm z" "(norm z * norm z ^ n) > 0"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
88  | 
apply (metis assms less_divide_eq mult.commute not_le)  | 
| 56215 | 89  | 
using norm1 apply (metis mult_pos_pos zero_less_power)  | 
90  | 
done  | 
|
91  | 
have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n)) =  | 
|
92  | 
(e + norm (c (Suc n))) * (norm z * norm z ^ n)"  | 
|
93  | 
by (simp add: norm_mult norm_power algebra_simps)  | 
|
94  | 
also have "... \<le> (e * norm z) * (norm z * norm z ^ n)"  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
95  | 
using norm2 by (metis real_mult_le_cancel_iff1)  | 
| 56215 | 96  | 
also have "... = e * (norm z * (norm z * norm z ^ n))"  | 
97  | 
by (simp add: algebra_simps)  | 
|
98  | 
finally have "e * (norm z * norm z ^ n) + norm (c (Suc n) * (z * z ^ n))  | 
|
99  | 
\<le> e * (norm z * (norm z * norm z ^ n))" .  | 
|
100  | 
then show "norm (\<Sum>i\<le>Suc n. c i * z^i) \<le> e * norm z ^ Suc (Suc n)" using M norm1  | 
|
101  | 
by (drule_tac x=z in spec) (auto simp: intro!: adhoc_norm_triangle)  | 
|
102  | 
qed  | 
|
103  | 
qed  | 
|
104  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
105  | 
lemma norm_lemma_xy: assumes "\<bar>b\<bar> + 1 \<le> norm(y) - a" "norm(x) \<le> a" shows "b \<le> norm(x + y)"  | 
| 
62626
 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 
paulson <lp15@cam.ac.uk> 
parents: 
61945 
diff
changeset
 | 
106  | 
proof -  | 
| 
 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 
paulson <lp15@cam.ac.uk> 
parents: 
61945 
diff
changeset
 | 
107  | 
have "b \<le> norm y - norm x"  | 
| 
 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 
paulson <lp15@cam.ac.uk> 
parents: 
61945 
diff
changeset
 | 
108  | 
using assms by linarith  | 
| 
 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 
paulson <lp15@cam.ac.uk> 
parents: 
61945 
diff
changeset
 | 
109  | 
then show ?thesis  | 
| 
 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 
paulson <lp15@cam.ac.uk> 
parents: 
61945 
diff
changeset
 | 
110  | 
by (metis (no_types) add.commute norm_diff_ineq order_trans)  | 
| 
 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 
paulson <lp15@cam.ac.uk> 
parents: 
61945 
diff
changeset
 | 
111  | 
qed  | 
| 56215 | 112  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
113  | 
proposition polyfun_extremal:  | 
| 56215 | 114  | 
fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra"  | 
115  | 
assumes "\<exists>k. k \<noteq> 0 \<and> k \<le> n \<and> c k \<noteq> 0"  | 
|
116  | 
shows "eventually (\<lambda>z. norm(\<Sum>i\<le>n. c i * z^i) \<ge> B) at_infinity"  | 
|
117  | 
using assms  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
118  | 
proof (induction n)  | 
| 56215 | 119  | 
case 0 then show ?case  | 
120  | 
by simp  | 
|
121  | 
next  | 
|
122  | 
case (Suc n)  | 
|
123  | 
show ?case  | 
|
124  | 
proof (cases "c (Suc n) = 0")  | 
|
125  | 
case True  | 
|
126  | 
with Suc show ?thesis  | 
|
127  | 
by auto (metis diff_is_0_eq diffs0_imp_equal less_Suc_eq_le not_less_eq)  | 
|
128  | 
next  | 
|
129  | 
case False  | 
|
130  | 
with polyfun_extremal_lemma [of "norm(c (Suc n)) / 2" c n]  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
131  | 
obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow>  | 
| 56215 | 132  | 
norm (\<Sum>i\<le>n. c i * z^i) \<le> norm (c (Suc n)) / 2 * norm z ^ Suc n"  | 
133  | 
by auto  | 
|
134  | 
show ?thesis  | 
|
135  | 
unfolding eventually_at_infinity  | 
|
| 61945 | 136  | 
proof (rule exI [where x="max M (max 1 ((\<bar>B\<bar> + 1) / (norm (c (Suc n)) / 2)))"], clarsimp)  | 
| 56215 | 137  | 
fix z::'a  | 
138  | 
assume les: "M \<le> norm z" "1 \<le> norm z" "(\<bar>B\<bar> * 2 + 2) / norm (c (Suc n)) \<le> norm z"  | 
|
139  | 
then have "\<bar>B\<bar> * 2 + 2 \<le> norm z * norm (c (Suc n))"  | 
|
140  | 
by (metis False pos_divide_le_eq zero_less_norm_iff)  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
141  | 
then have "\<bar>B\<bar> * 2 + 2 \<le> norm z ^ (Suc n) * norm (c (Suc n))"  | 
| 60420 | 142  | 
by (metis \<open>1 \<le> norm z\<close> order.trans mult_right_mono norm_ge_zero self_le_power zero_less_Suc)  | 
| 56215 | 143  | 
then show "B \<le> norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * (z * z ^ n))" using M les  | 
144  | 
apply auto  | 
|
145  | 
apply (rule norm_lemma_xy [where a = "norm (c (Suc n)) * norm z ^ (Suc n) / 2"])  | 
|
146  | 
apply (simp_all add: norm_mult norm_power)  | 
|
147  | 
done  | 
|
148  | 
qed  | 
|
149  | 
qed  | 
|
150  | 
qed  | 
|
151  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
152  | 
proposition polyfun_rootbound:  | 
| 56215 | 153  | 
 fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
154  | 
assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"  | 
|
155  | 
   shows "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<and> card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
 | 
|
156  | 
using assms  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
157  | 
proof (induction n arbitrary: c)  | 
| 56215 | 158  | 
case (Suc n) show ?case  | 
159  | 
 proof (cases "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = {}")
 | 
|
160  | 
case False  | 
|
161  | 
then obtain a where a: "(\<Sum>i\<le>Suc n. c i * a^i) = 0"  | 
|
162  | 
by auto  | 
|
163  | 
from polyfun_linear_factor_root [OF this]  | 
|
164  | 
obtain b where "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i< Suc n. b i * z^i)"  | 
|
165  | 
by auto  | 
|
166  | 
then have b: "\<And>z. (\<Sum>i\<le>Suc n. c i * z^i) = (z - a) * (\<Sum>i\<le>n. b i * z^i)"  | 
|
167  | 
by (metis lessThan_Suc_atMost)  | 
|
168  | 
   then have ins_ab: "{z. (\<Sum>i\<le>Suc n. c i * z^i) = 0} = insert a {z. (\<Sum>i\<le>n. b i * z^i) = 0}"
 | 
|
169  | 
by auto  | 
|
170  | 
have c0: "c 0 = - (a * b 0)" using b [of 0]  | 
|
171  | 
by simp  | 
|
| 69508 | 172  | 
then have extr_prem: "\<not> (\<exists>k\<le>n. b k \<noteq> 0) \<Longrightarrow> \<exists>k. k \<noteq> 0 \<and> k \<le> Suc n \<and> c k \<noteq> 0"  | 
| 56215 | 173  | 
by (metis Suc.prems le0 minus_zero mult_zero_right)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
174  | 
have "\<exists>k\<le>n. b k \<noteq> 0"  | 
| 56215 | 175  | 
apply (rule ccontr)  | 
176  | 
using polyfun_extremal [OF extr_prem, of 1]  | 
|
| 
70097
 
4005298550a6
The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
 
paulson <lp15@cam.ac.uk> 
parents: 
69739 
diff
changeset
 | 
177  | 
apply (auto simp: eventually_at_infinity b simp del: sum.atMost_Suc)  | 
| 56215 | 178  | 
apply (drule_tac x="of_real ba" in spec, simp)  | 
179  | 
done  | 
|
180  | 
then show ?thesis using Suc.IH [of b] ins_ab  | 
|
181  | 
by (auto simp: card_insert_if)  | 
|
182  | 
qed simp  | 
|
183  | 
qed simp  | 
|
184  | 
||
| 69739 | 185  | 
corollary  | 
| 56215 | 186  | 
  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
187  | 
assumes "\<exists>k. k \<le> n \<and> c k \<noteq> 0"  | 
|
188  | 
    shows polyfun_rootbound_finite: "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | 
|
189  | 
      and polyfun_rootbound_card:   "card {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<le> n"
 | 
|
190  | 
using polyfun_rootbound [OF assms] by auto  | 
|
191  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
192  | 
proposition polyfun_finite_roots:  | 
| 56215 | 193  | 
  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
194  | 
    shows  "finite {z. (\<Sum>i\<le>n. c i * z^i) = 0} \<longleftrightarrow> (\<exists>k. k \<le> n \<and> c k \<noteq> 0)"
 | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
195  | 
proof (cases " \<exists>k\<le>n. c k \<noteq> 0")  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
196  | 
case True then show ?thesis  | 
| 56215 | 197  | 
by (blast intro: polyfun_rootbound_finite)  | 
198  | 
next  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
60420 
diff
changeset
 | 
199  | 
case False then show ?thesis  | 
| 56215 | 200  | 
by (auto simp: infinite_UNIV_char_0)  | 
201  | 
qed  | 
|
202  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
203  | 
lemma polyfun_eq_0:  | 
| 56215 | 204  | 
  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
205  | 
shows "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0) \<longleftrightarrow> (\<forall>k. k \<le> n \<longrightarrow> c k = 0)"  | 
|
206  | 
proof (cases "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = 0)")  | 
|
207  | 
case True  | 
|
| 69508 | 208  | 
  then have "\<not> finite {z. (\<Sum>i\<le>n. c i * z^i) = 0}"
 | 
| 56215 | 209  | 
by (simp add: infinite_UNIV_char_0)  | 
210  | 
with True show ?thesis  | 
|
211  | 
by (metis (poly_guards_query) polyfun_rootbound_finite)  | 
|
212  | 
next  | 
|
213  | 
case False  | 
|
214  | 
then show ?thesis  | 
|
215  | 
by auto  | 
|
216  | 
qed  | 
|
217  | 
||
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
218  | 
theorem polyfun_eq_const:  | 
| 56215 | 219  | 
  fixes c :: "nat \<Rightarrow> 'a::{comm_ring,real_normed_div_algebra}"
 | 
220  | 
shows "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"  | 
|
| 
69730
 
0c3dcb3a17f6
tagged 5 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69683 
diff
changeset
 | 
221  | 
proof -  | 
| 56215 | 222  | 
  {fix z
 | 
223  | 
have "(\<Sum>i\<le>n. c i * z^i) = (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) + k"  | 
|
224  | 
by (induct n) auto  | 
|
225  | 
} then  | 
|
226  | 
have "(\<forall>z. (\<Sum>i\<le>n. c i * z^i) = k) \<longleftrightarrow> (\<forall>z. (\<Sum>i\<le>n. (if i = 0 then c 0 - k else c i) * z^i) = 0)"  | 
|
227  | 
by auto  | 
|
228  | 
also have "... \<longleftrightarrow> c 0 = k \<and> (\<forall>k. k \<noteq> 0 \<and> k \<le> n \<longrightarrow> c k = 0)"  | 
|
229  | 
by (auto simp: polyfun_eq_0)  | 
|
230  | 
finally show ?thesis .  | 
|
231  | 
qed  | 
|
232  | 
||
233  | 
end  | 
|
234  |