7998
|
1 |
(*
|
|
2 |
Sums with naturals as index domain
|
|
3 |
$Id$
|
|
4 |
Author: Clemens Ballarin, started 12 December 1996
|
|
5 |
*)
|
|
6 |
|
|
7 |
section "Basic Properties";
|
|
8 |
|
|
9 |
Goalw [SUM_def] "SUM 0 f = (f 0::'a::ring)";
|
|
10 |
by (Asm_simp_tac 1);
|
|
11 |
qed "SUM_0";
|
|
12 |
|
|
13 |
Goalw [SUM_def]
|
|
14 |
"SUM (Suc n) f = (f (Suc n) + SUM n f::'a::ring)";
|
|
15 |
by (Simp_tac 1);
|
|
16 |
qed "SUM_Suc";
|
|
17 |
|
|
18 |
Addsimps [SUM_0, SUM_Suc];
|
|
19 |
|
|
20 |
Goal
|
|
21 |
"SUM (Suc n) f = (SUM n (%i. f (Suc i)) + f 0::'a::ring)";
|
8707
|
22 |
by (induct_tac "n" 1);
|
7998
|
23 |
(* Base case *)
|
|
24 |
by (Simp_tac 1);
|
|
25 |
(* Induction step *)
|
|
26 |
by (asm_full_simp_tac (simpset() addsimps [a_assoc]) 1);
|
|
27 |
qed "SUM_Suc2";
|
|
28 |
|
|
29 |
(* Congruence rules *)
|
|
30 |
|
|
31 |
val [p_equal, p_context] = goal NatSum.thy
|
|
32 |
"[| m = n; !!i. i <= n ==> f i = g i |] ==> SUM m f = (SUM n g::'a::ring)";
|
|
33 |
by (simp_tac (simpset() addsimps [p_equal]) 1);
|
|
34 |
by (cut_inst_tac [("n", "n")] le_refl 1);
|
|
35 |
by (etac rev_mp 1);
|
|
36 |
by (res_inst_tac [("P", "%k. k <= n --> SUM k f = SUM k g")] nat_induct 1);
|
|
37 |
(* Base case *)
|
|
38 |
by (simp_tac (simpset() addsimps [p_context]) 1);
|
|
39 |
(* Induction step *)
|
|
40 |
by (asm_simp_tac (simpset() addsimps [p_context]) 1);
|
|
41 |
qed "SUM_cong";
|
|
42 |
|
|
43 |
Addcongs [SUM_cong];
|
|
44 |
|
|
45 |
(* Results needed for the development of polynomials *)
|
|
46 |
|
|
47 |
section "Ring Properties";
|
|
48 |
|
|
49 |
Goal "SUM n (%i. <0>) = (<0>::'a::ring)";
|
8707
|
50 |
by (induct_tac "n" 1);
|
7998
|
51 |
by (Simp_tac 1);
|
|
52 |
by (Asm_simp_tac 1);
|
|
53 |
qed "SUM_zero";
|
|
54 |
|
|
55 |
Addsimps [SUM_zero];
|
|
56 |
|
|
57 |
Goal
|
|
58 |
"!!f::nat=>'a::ring. SUM n (%i. f i + g i) = SUM n f + SUM n g";
|
8707
|
59 |
by (induct_tac "n" 1);
|
7998
|
60 |
(* Base case *)
|
|
61 |
by (Simp_tac 1);
|
|
62 |
(* Induction step *)
|
|
63 |
by (asm_simp_tac (simpset() addsimps a_ac) 1);
|
|
64 |
qed "SUM_add";
|
|
65 |
|
|
66 |
Goal
|
|
67 |
"!!a::'a::ring. SUM n f * a = SUM n (%i. f i * a)";
|
8707
|
68 |
by (induct_tac "n" 1);
|
7998
|
69 |
(* Base case *)
|
|
70 |
by (Simp_tac 1);
|
|
71 |
(* Induction step *)
|
|
72 |
by (asm_simp_tac (simpset() addsimps [l_distr]) 1);
|
|
73 |
qed "SUM_ldistr";
|
|
74 |
|
|
75 |
Goal
|
|
76 |
"!!a::'a::ring. a * SUM n f = SUM n (%i. a * f i)";
|
8707
|
77 |
by (induct_tac "n" 1);
|
7998
|
78 |
(* Base case *)
|
|
79 |
by (Simp_tac 1);
|
|
80 |
(* Induction step *)
|
|
81 |
by (asm_simp_tac (simpset() addsimps [r_distr]) 1);
|
|
82 |
qed "SUM_rdistr";
|
|
83 |
|
|
84 |
section "Results for the Construction of Polynomials";
|
|
85 |
|
|
86 |
Goal
|
|
87 |
"!!a::nat=>'a::ring. k <= n --> \
|
|
88 |
\ SUM k (%j. SUM j (%i. a i * b (j-i)) * c (n-j)) = \
|
|
89 |
\ SUM k (%j. a j * SUM (k-j) (%i. b i * c (n-j-i)))";
|
8707
|
90 |
by (induct_tac "k" 1);
|
7998
|
91 |
(* Base case *)
|
|
92 |
by (simp_tac (simpset() addsimps [m_assoc]) 1);
|
|
93 |
(* Induction step *)
|
|
94 |
by (rtac impI 1);
|
|
95 |
by (etac impE 1);
|
8863
|
96 |
by (arith_tac 1);
|
8738
|
97 |
by (asm_simp_tac
|
|
98 |
(simpset() addsimps a_ac@
|
|
99 |
[Suc_diff_le, l_distr, r_distr, m_assoc, SUM_add]) 1);
|
8863
|
100 |
by (asm_simp_tac (simpset() addsimps a_ac@ [SUM_ldistr, m_assoc]) 1);
|
|
101 |
qed_spec_mp "poly_assoc_lemma1";
|
7998
|
102 |
|
|
103 |
Goal
|
|
104 |
"!!a::nat=>'a::ring. \
|
|
105 |
\ SUM n (%j. SUM j (%i. a i * b (j-i)) * c (n-j)) = \
|
8863
|
106 |
\ SUM n (%j. a j * SUM (n-j) (%i. b i * c (n-(j+i))))";
|
|
107 |
by (asm_simp_tac (simpset() addsimps [poly_assoc_lemma1]) 1);
|
7998
|
108 |
qed "poly_assoc_lemma";
|
|
109 |
|
|
110 |
Goal
|
|
111 |
"!! a::nat=>'a::ring. j <= n --> \
|
|
112 |
\ SUM j (%i. a i * b (n-i)) = SUM j (%i. a (n-i-(n-j)) * b (i+(n-j)))";
|
8707
|
113 |
by (induct_tac "j" 1);
|
7998
|
114 |
(* Base case *)
|
|
115 |
by (Simp_tac 1);
|
|
116 |
(* Induction step *)
|
|
117 |
by (stac SUM_Suc2 1);
|
8863
|
118 |
by (asm_simp_tac (simpset() addsimps [Suc_diff_Suc, a_comm]) 1);
|
7998
|
119 |
qed "poly_comm_lemma1";
|
|
120 |
|
|
121 |
Goal
|
|
122 |
"!!a::nat=>'a::ring. SUM n (%i. a i * b (n-i)) = SUM n (%i. a (n-i) * b i)";
|
|
123 |
by (cut_inst_tac [("j", "n"), ("n", "n")] poly_comm_lemma1 1);
|
|
124 |
by (Asm_full_simp_tac 1);
|
|
125 |
qed "poly_comm_lemma";
|
|
126 |
|
|
127 |
section "Changing the Range of Summation";
|
|
128 |
|
|
129 |
Goal
|
|
130 |
"!! f::(nat=>'a::ring). \
|
|
131 |
\ SUM n (%i. if i = x then f i else <0>) = (if x <= n then f x else <0>)";
|
8707
|
132 |
by (induct_tac "n" 1);
|
7998
|
133 |
(* Base case *)
|
8006
|
134 |
by (Simp_tac 1);
|
7998
|
135 |
(* Induction step *)
|
8006
|
136 |
by (Asm_simp_tac 1);
|
7998
|
137 |
by (Clarify_tac 1);
|
|
138 |
by (res_inst_tac [("f", "f")] arg_cong 1);
|
8006
|
139 |
by (arith_tac 1);
|
7998
|
140 |
qed "SUM_if_singleton";
|
|
141 |
|
|
142 |
Goal
|
|
143 |
"!! f::(nat=>'a::ring). \
|
|
144 |
\ m <= n & (ALL i. m < i & i <= n --> f i = <0>) --> \
|
|
145 |
\ SUM m f = SUM n f";
|
8707
|
146 |
by (induct_tac "n" 1);
|
7998
|
147 |
(* Base case *)
|
|
148 |
by (Simp_tac 1);
|
|
149 |
(* Induction step *)
|
|
150 |
by (case_tac "m <= n" 1);
|
8707
|
151 |
by Auto_tac;
|
7998
|
152 |
by (subgoal_tac "m = Suc n" 1);
|
|
153 |
by (Asm_simp_tac 1);
|
|
154 |
by (fast_arith_tac 1);
|
|
155 |
val SUM_shrink_lemma = result();
|
|
156 |
|
|
157 |
Goal
|
|
158 |
"!! f::(nat=>'a::ring). \
|
|
159 |
\ [| m <= n; !!i. [| m < i; i <= n |] ==> f i = <0>; P (SUM n f) |] ==> \
|
|
160 |
\ P (SUM m f)";
|
|
161 |
by (cut_inst_tac [("m", "m"), ("n", "n"), ("f", "f")] SUM_shrink_lemma 1);
|
|
162 |
by (Asm_full_simp_tac 1);
|
|
163 |
qed "SUM_shrink";
|
|
164 |
|
|
165 |
Goal
|
|
166 |
"!! f::(nat=>'a::ring). \
|
|
167 |
\ [| m <= n; !!i. [| m < i; i <= n |] ==> f i = <0>; P (SUM m f) |] ==> \
|
|
168 |
\ P (SUM n f)";
|
|
169 |
by (cut_inst_tac [("m", "m"), ("n", "n"), ("f", "f")] SUM_shrink_lemma 1);
|
|
170 |
by (Asm_full_simp_tac 1);
|
|
171 |
qed "SUM_extend";
|
|
172 |
|
|
173 |
Goal
|
|
174 |
"!! f::(nat=>'a::ring). \
|
|
175 |
\ (ALL i. i < m --> f i = <0>) --> SUM d (%i. f (i+m)) = SUM (m+d) f";
|
8707
|
176 |
by (induct_tac "d" 1);
|
7998
|
177 |
(* Base case *)
|
8707
|
178 |
by (induct_tac "m" 1);
|
7998
|
179 |
by (Simp_tac 1);
|
8707
|
180 |
by (Force_tac 1);
|
7998
|
181 |
(* Induction step *)
|
|
182 |
by (asm_simp_tac (simpset() addsimps add_ac) 1);
|
|
183 |
val SUM_shrink_below_lemma = result();
|
|
184 |
|
|
185 |
Goal
|
|
186 |
"!! f::(nat=>'a::ring). \
|
|
187 |
\ [| m <= n; !!i. i < m ==> f i = <0>; P (SUM (n-m) (%i. f (i+m))) |] ==> \
|
|
188 |
\ P (SUM n f)";
|
8707
|
189 |
by (asm_full_simp_tac
|
|
190 |
(simpset() addsimps [SUM_shrink_below_lemma, add_diff_inverse, leD]) 1);
|
7998
|
191 |
qed "SUM_extend_below";
|
|
192 |
|
|
193 |
section "Result for the Univeral Property of Polynomials";
|
|
194 |
|
|
195 |
Goal
|
|
196 |
"!!f::nat=>'a::ring. j <= n + m --> \
|
|
197 |
\ SUM j (%k. SUM k (%i. f i * g (k - i))) = \
|
|
198 |
\ SUM j (%k. SUM (j - k) (%i. f k * g i))";
|
8707
|
199 |
by (induct_tac "j" 1);
|
7998
|
200 |
(* Base case *)
|
|
201 |
by (Simp_tac 1);
|
|
202 |
(* Induction step *)
|
8863
|
203 |
by (simp_tac (simpset() addsimps [Suc_diff_le, SUM_add]) 1);
|
7998
|
204 |
by (asm_simp_tac (simpset() addsimps a_ac) 1);
|
|
205 |
val DiagSum_lemma = result();
|
|
206 |
|
|
207 |
Goal
|
|
208 |
"!!f::nat=>'a::ring. \
|
|
209 |
\ SUM (n + m) (%k. SUM k (%i. f i * g (k - i))) = \
|
|
210 |
\ SUM (n + m) (%k. SUM (n + m - k) (%i. f k * g i))";
|
|
211 |
by (rtac (DiagSum_lemma RS mp) 1);
|
|
212 |
by (rtac le_refl 1);
|
|
213 |
qed "DiagSum";
|
|
214 |
|
|
215 |
|
|
216 |
|
|
217 |
|