src/HOL/Algebra/abstract/NatSum.ML
author paulson
Thu, 11 Nov 1999 10:25:29 +0100
changeset 8006 299127ded09d
parent 7998 3d0c34795831
child 8707 5de763446504
permissions -rw-r--r--
tidied
paulson@7998
     1
(*
paulson@7998
     2
    Sums with naturals as index domain
paulson@7998
     3
    $Id$
paulson@7998
     4
    Author: Clemens Ballarin, started 12 December 1996
paulson@7998
     5
*)
paulson@7998
     6
paulson@7998
     7
section "Basic Properties";
paulson@7998
     8
paulson@7998
     9
Goalw [SUM_def] "SUM 0 f = (f 0::'a::ring)";
paulson@7998
    10
by (Asm_simp_tac 1);
paulson@7998
    11
qed "SUM_0";
paulson@7998
    12
paulson@7998
    13
Goalw [SUM_def]
paulson@7998
    14
  "SUM (Suc n) f = (f (Suc n) + SUM n f::'a::ring)";
paulson@7998
    15
by (Simp_tac 1);
paulson@7998
    16
qed "SUM_Suc";
paulson@7998
    17
paulson@7998
    18
Addsimps [SUM_0, SUM_Suc];
paulson@7998
    19
paulson@7998
    20
Goal
paulson@7998
    21
  "SUM (Suc n) f = (SUM n (%i. f (Suc i)) + f 0::'a::ring)";
paulson@7998
    22
by (nat_ind_tac "n" 1);
paulson@7998
    23
(* Base case *)
paulson@7998
    24
by (Simp_tac 1);
paulson@7998
    25
(* Induction step *)
paulson@7998
    26
by (asm_full_simp_tac (simpset() addsimps [a_assoc]) 1);
paulson@7998
    27
qed "SUM_Suc2";
paulson@7998
    28
paulson@7998
    29
(* Congruence rules *)
paulson@7998
    30
paulson@7998
    31
val [p_equal, p_context] = goal NatSum.thy
paulson@7998
    32
  "[| m = n; !!i. i <= n ==> f i = g i |] ==> SUM m f = (SUM n g::'a::ring)";
paulson@7998
    33
by (simp_tac (simpset() addsimps [p_equal]) 1);
paulson@7998
    34
by (cut_inst_tac [("n", "n")] le_refl 1);
paulson@7998
    35
by (etac rev_mp 1);
paulson@7998
    36
by (res_inst_tac [("P", "%k. k <= n --> SUM k f = SUM k g")] nat_induct 1);
paulson@7998
    37
(* Base case *)
paulson@7998
    38
by (simp_tac (simpset() addsimps [p_context]) 1);
paulson@7998
    39
(* Induction step *)
paulson@7998
    40
by (rtac impI 1);
paulson@7998
    41
by (etac impE 1);
paulson@7998
    42
by (rtac Suc_leD 1);
paulson@7998
    43
by (assume_tac 1);
paulson@7998
    44
by (asm_simp_tac (simpset() addsimps [p_context]) 1);
paulson@7998
    45
qed "SUM_cong";
paulson@7998
    46
paulson@7998
    47
Addcongs [SUM_cong];
paulson@7998
    48
paulson@7998
    49
(* Results needed for the development of polynomials *)
paulson@7998
    50
paulson@7998
    51
section "Ring Properties";
paulson@7998
    52
paulson@7998
    53
Goal "SUM n (%i. <0>) = (<0>::'a::ring)";
paulson@7998
    54
by (nat_ind_tac "n" 1);
paulson@7998
    55
by (Simp_tac 1);
paulson@7998
    56
by (Asm_simp_tac 1);
paulson@7998
    57
qed "SUM_zero";
paulson@7998
    58
paulson@7998
    59
Addsimps [SUM_zero];
paulson@7998
    60
paulson@7998
    61
Goal
paulson@7998
    62
  "!!f::nat=>'a::ring. SUM n (%i. f i + g i) = SUM n f + SUM n g";
paulson@7998
    63
by (nat_ind_tac "n" 1);
paulson@7998
    64
(* Base case *)
paulson@7998
    65
by (Simp_tac 1);
paulson@7998
    66
(* Induction step *)
paulson@7998
    67
by (asm_simp_tac (simpset() addsimps a_ac) 1);
paulson@7998
    68
qed "SUM_add";
paulson@7998
    69
paulson@7998
    70
Goal
paulson@7998
    71
  "!!a::'a::ring. SUM n f * a = SUM n (%i. f i * a)";
paulson@7998
    72
by (nat_ind_tac "n" 1);
paulson@7998
    73
(* Base case *)
paulson@7998
    74
by (Simp_tac 1);
paulson@7998
    75
(* Induction step *)
paulson@7998
    76
by (asm_simp_tac (simpset() addsimps [l_distr]) 1);
paulson@7998
    77
qed "SUM_ldistr";
paulson@7998
    78
paulson@7998
    79
Goal
paulson@7998
    80
  "!!a::'a::ring. a * SUM n f = SUM n (%i. a * f i)";
paulson@7998
    81
by (nat_ind_tac "n" 1);
paulson@7998
    82
(* Base case *)
paulson@7998
    83
by (Simp_tac 1);
paulson@7998
    84
(* Induction step *)
paulson@7998
    85
by (asm_simp_tac (simpset() addsimps [r_distr]) 1);
paulson@7998
    86
qed "SUM_rdistr";
paulson@7998
    87
paulson@7998
    88
section "Results for the Construction of Polynomials";
paulson@7998
    89
paulson@7998
    90
goal Main.thy (* could go to Arith *)
paulson@7998
    91
  "!!j::nat. [| m <= j; Suc j <= n |] ==> (n - m) - Suc (j - m) = n - Suc j";
paulson@7998
    92
by (asm_simp_tac (simpset() addsimps [Suc_diff_le RS sym]) 1);
paulson@7998
    93
by (asm_simp_tac (simpset() addsimps [diff_right_cancel, less_imp_le]) 1);
paulson@7998
    94
qed "Suc_diff_lemma";
paulson@7998
    95
paulson@7998
    96
Goal
paulson@7998
    97
  "!!a::nat=>'a::ring. k <= n --> \
paulson@7998
    98
\  SUM k (%j. SUM j (%i. a i * b (j-i)) * c (n-j)) = \
paulson@7998
    99
\  SUM k (%j. a j * SUM (k-j) (%i. b i * c (n-j-i)))";
paulson@7998
   100
by (nat_ind_tac "k" 1);
paulson@7998
   101
(* Base case *)
paulson@7998
   102
by (simp_tac (simpset() addsimps [m_assoc]) 1);
paulson@7998
   103
(* Induction step *)
paulson@7998
   104
by (rtac impI 1);
paulson@7998
   105
by (etac impE 1);
paulson@7998
   106
by (rtac Suc_leD 1);
paulson@7998
   107
by (assume_tac 1);
paulson@7998
   108
by (asm_simp_tac (simpset() addsimps a_ac@[Suc_diff_le, l_distr, r_distr, m_assoc, SUM_add]) 1);
paulson@7998
   109
by (asm_simp_tac (simpset() addsimps a_ac@[Suc_diff_lemma, SUM_ldistr, m_assoc]) 1);
paulson@7998
   110
qed "poly_assoc_lemma1";
paulson@7998
   111
paulson@7998
   112
Goal
paulson@7998
   113
  "!!a::nat=>'a::ring. \
paulson@7998
   114
\  SUM n (%j. SUM j (%i. a i * b (j-i)) * c (n-j)) = \
paulson@7998
   115
\  SUM n (%j. a j * SUM (n-j) (%i. b i * c (n-j-i)))";
paulson@7998
   116
by (rtac (poly_assoc_lemma1 RS mp) 1);
paulson@7998
   117
by (rtac le_refl 1);
paulson@7998
   118
qed "poly_assoc_lemma";
paulson@7998
   119
paulson@7998
   120
goal Main.thy (* could go to Arith *)
paulson@7998
   121
  "!! n. Suc i <= n ==> Suc (a + (n - Suc i)) = a + (n - i)";
paulson@7998
   122
by (asm_simp_tac (simpset() delsimps [add_Suc] addsimps [add_Suc_right RS sym, Suc_diff_Suc]) 1);
paulson@7998
   123
qed "Suc_add_diff_Suc";
paulson@7998
   124
paulson@7998
   125
goal Main.thy (* could go to Arith *)
paulson@7998
   126
  "!! n. [| Suc j <= n; i <= j |] ==> \
paulson@7998
   127
\    n - Suc i - (n - Suc j) = n - i - (n - j)";
paulson@7998
   128
by (res_inst_tac [("m1", "n - Suc i"), ("n1", "n - Suc j")]
paulson@7998
   129
  (diff_Suc_Suc RS subst) 1);
paulson@7998
   130
by (subgoal_tac "Suc i <= n" 1);
paulson@7998
   131
by (asm_simp_tac (simpset() delsimps [diff_Suc_Suc] addsimps [Suc_diff_Suc]) 1);
paulson@7998
   132
by (fast_arith_tac 1);
paulson@7998
   133
qed "diff_lemma2";
paulson@7998
   134
paulson@7998
   135
Goal
paulson@7998
   136
  "!! a::nat=>'a::ring. j <= n --> \
paulson@7998
   137
\    SUM j (%i. a i * b (n-i)) = SUM j (%i. a (n-i-(n-j)) * b (i+(n-j)))";
paulson@7998
   138
by (nat_ind_tac "j" 1);
paulson@7998
   139
(* Base case *)
paulson@7998
   140
by (Simp_tac 1);
paulson@7998
   141
(* Induction step *)
paulson@7998
   142
by (rtac impI 1);
paulson@7998
   143
by (etac impE 1);
paulson@7998
   144
by (rtac Suc_leD 1);
paulson@7998
   145
by (assume_tac 1);
paulson@7998
   146
by (stac SUM_Suc2 1);
paulson@7998
   147
by (stac SUM_Suc 1);
paulson@7998
   148
by (asm_simp_tac (simpset()
paulson@7998
   149
    addsimps [a_comm, Suc_add_diff_Suc, diff_diff_cancel, diff_lemma2]) 1);
paulson@7998
   150
qed "poly_comm_lemma1";
paulson@7998
   151
paulson@7998
   152
Goal
paulson@7998
   153
 "!!a::nat=>'a::ring. SUM n (%i. a i * b (n-i)) = SUM n (%i. a (n-i) * b i)";
paulson@7998
   154
by (cut_inst_tac [("j", "n"), ("n", "n")] poly_comm_lemma1 1);
paulson@7998
   155
by (Asm_full_simp_tac 1);
paulson@7998
   156
qed "poly_comm_lemma";
paulson@7998
   157
paulson@7998
   158
section "Changing the Range of Summation";
paulson@7998
   159
paulson@7998
   160
Goal
paulson@7998
   161
  "!! f::(nat=>'a::ring). \
paulson@7998
   162
\    SUM n (%i. if i = x then f i else <0>) = (if x <= n then f x else <0>)";
paulson@7998
   163
by (nat_ind_tac "n" 1);
paulson@7998
   164
(* Base case *)
paulson@8006
   165
by (Simp_tac 1);
paulson@7998
   166
(* Induction step *)
paulson@8006
   167
by (Asm_simp_tac 1);
paulson@7998
   168
by (Clarify_tac 1);
paulson@7998
   169
by (res_inst_tac [("f", "f")] arg_cong 1);
paulson@8006
   170
by (arith_tac 1);
paulson@7998
   171
qed "SUM_if_singleton";
paulson@7998
   172
paulson@7998
   173
Goal
paulson@7998
   174
  "!! f::(nat=>'a::ring). \
paulson@7998
   175
\    m <= n & (ALL i. m < i & i <= n --> f i = <0>) --> \
paulson@7998
   176
\    SUM m f = SUM n f";
paulson@7998
   177
by (nat_ind_tac "n" 1);
paulson@7998
   178
(* Base case *)
paulson@7998
   179
by (Simp_tac 1);
paulson@7998
   180
(* Induction step *)
paulson@7998
   181
by (case_tac "m <= n" 1);
paulson@7998
   182
by (rtac impI 1);
paulson@7998
   183
by (etac impE 1);
paulson@7998
   184
by (SELECT_GOAL Auto_tac 1);
paulson@7998
   185
by (etac conjE 1);
paulson@7998
   186
by (dres_inst_tac [("x", "Suc n")] spec 1);
paulson@7998
   187
by (asm_simp_tac (simpset() addsimps [less_Suc_eq_le]) 1);
paulson@7998
   188
(* case n < m, in fact m = Suc n *)
paulson@7998
   189
by (rtac impI 1);
paulson@7998
   190
by (etac conjE 1);
paulson@7998
   191
by (subgoal_tac "m = Suc n" 1);
paulson@7998
   192
by (Asm_simp_tac 1);
paulson@7998
   193
by (fast_arith_tac 1);
paulson@7998
   194
val SUM_shrink_lemma = result();
paulson@7998
   195
paulson@7998
   196
Goal
paulson@7998
   197
  "!! f::(nat=>'a::ring). \
paulson@7998
   198
\    [| m <= n; !!i. [| m < i; i <= n |] ==> f i = <0>; P (SUM n f) |] ==> \
paulson@7998
   199
\      P (SUM m f)";
paulson@7998
   200
by (cut_inst_tac [("m", "m"), ("n", "n"), ("f", "f")] SUM_shrink_lemma 1);
paulson@7998
   201
by (Asm_full_simp_tac 1);
paulson@7998
   202
qed "SUM_shrink";
paulson@7998
   203
paulson@7998
   204
Goal
paulson@7998
   205
  "!! f::(nat=>'a::ring). \
paulson@7998
   206
\    [| m <= n; !!i. [| m < i; i <= n |] ==> f i = <0>; P (SUM m f) |] ==> \
paulson@7998
   207
\      P (SUM n f)";
paulson@7998
   208
by (cut_inst_tac [("m", "m"), ("n", "n"), ("f", "f")] SUM_shrink_lemma 1);
paulson@7998
   209
by (Asm_full_simp_tac 1);
paulson@7998
   210
qed "SUM_extend";
paulson@7998
   211
paulson@7998
   212
Goal
paulson@7998
   213
  "!! f::(nat=>'a::ring). \
paulson@7998
   214
\    (ALL i. i < m --> f i = <0>) --> SUM d (%i. f (i+m)) = SUM (m+d) f";
paulson@7998
   215
by (nat_ind_tac "d" 1);
paulson@7998
   216
(* Base case *)
paulson@7998
   217
by (nat_ind_tac "m" 1);
paulson@7998
   218
by (Simp_tac 1);
paulson@7998
   219
by (rtac impI 1);
paulson@7998
   220
by (etac impE 1);
paulson@7998
   221
by (Asm_simp_tac 1);
paulson@7998
   222
by (subgoal_tac "SUM m f = <0>" 1);
paulson@7998
   223
by (Asm_simp_tac 1);
paulson@7998
   224
by (Asm_full_simp_tac 1);
paulson@7998
   225
(* Induction step *)
paulson@7998
   226
by (asm_simp_tac (simpset() addsimps add_ac) 1);
paulson@7998
   227
val SUM_shrink_below_lemma = result();
paulson@7998
   228
paulson@7998
   229
Goal
paulson@7998
   230
  "!! f::(nat=>'a::ring). \
paulson@7998
   231
\    [| m <= n; !!i. i < m ==> f i = <0>; P (SUM (n-m) (%i. f (i+m))) |] ==> \
paulson@7998
   232
\    P (SUM n f)";
paulson@7998
   233
by (asm_full_simp_tac (simpset() addsimps
paulson@7998
   234
  [SUM_shrink_below_lemma, add_diff_inverse, leD]) 1);
paulson@7998
   235
qed "SUM_extend_below";
paulson@7998
   236
paulson@7998
   237
section "Result for the Univeral Property of Polynomials";
paulson@7998
   238
paulson@7998
   239
Goal
paulson@7998
   240
  "!!f::nat=>'a::ring. j <= n + m --> \
paulson@7998
   241
\    SUM j (%k. SUM k (%i. f i * g (k - i))) = \
paulson@7998
   242
\    SUM j (%k. SUM (j - k) (%i. f k * g i))";
paulson@7998
   243
by (nat_ind_tac "j" 1);
paulson@7998
   244
(* Base case *)
paulson@7998
   245
by (Simp_tac 1);
paulson@7998
   246
(* Induction step *)
paulson@7998
   247
by (simp_tac (simpset() addsimps [Suc_diff_le]) 1);
paulson@7998
   248
by (simp_tac (simpset() addsimps [SUM_add]) 1);
paulson@7998
   249
by (rtac impI 1);
paulson@7998
   250
by (etac impE 1);
paulson@7998
   251
by (dtac Suc_leD 1);
paulson@7998
   252
by (assume_tac 1);
paulson@7998
   253
by (asm_simp_tac (simpset() addsimps a_ac) 1);
paulson@7998
   254
val DiagSum_lemma = result();
paulson@7998
   255
paulson@7998
   256
Goal
paulson@7998
   257
  "!!f::nat=>'a::ring. \
paulson@7998
   258
\    SUM (n + m) (%k. SUM k (%i. f i * g (k - i))) = \
paulson@7998
   259
\    SUM (n + m) (%k. SUM (n + m - k) (%i. f k * g i))";
paulson@7998
   260
by (rtac (DiagSum_lemma RS mp) 1);
paulson@7998
   261
by (rtac le_refl 1);
paulson@7998
   262
qed "DiagSum";
paulson@7998
   263
paulson@7998
   264
paulson@7998
   265
paulson@7998
   266