src/HOL/ex/Summation.thy
author nipkow
Mon, 13 Sep 2010 11:13:15 +0200
changeset 39302 d7728f65b353
parent 39198 f967a16dfcdd
permissions -rw-r--r--
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     2
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     3
header {* Some basic facts about discrete summation *}
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     4
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     5
theory Summation
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     6
imports Main
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     7
begin
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     8
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
     9
text {* Auxiliary. *}
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    10
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    11
lemma add_setsum_orient:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    12
  "setsum f {k..<j} + setsum f {l..<k} = setsum f {l..<k} + setsum f {k..<j}"
35732
3b17dff14c4f tuned prefix of ac rules
haftmann
parents: 35267
diff changeset
    13
  by (fact add.commute)
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    14
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    15
lemma add_setsum_int:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    16
  fixes j k l :: int
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    17
  shows "j < k \<Longrightarrow> k < l \<Longrightarrow> setsum f {j..<k} + setsum f {k..<l} = setsum f {j..<l}"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    18
  by (simp_all add: setsum_Un_Int [symmetric] ivl_disj_un)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    19
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    20
text {* The shift operator. *}
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    21
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    22
definition \<Delta> :: "(int \<Rightarrow> 'a\<Colon>ab_group_add) \<Rightarrow> int \<Rightarrow> 'a" where
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    23
  "\<Delta> f k = f (k + 1) - f k"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    24
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    25
lemma \<Delta>_shift:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    26
  "\<Delta> (\<lambda>k. l + f k) = \<Delta> f"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    27
  by (simp add: \<Delta>_def fun_eq_iff)
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    28
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    29
lemma \<Delta>_same_shift:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    30
  assumes "\<Delta> f = \<Delta> g"
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35163
diff changeset
    31
  shows "\<exists>l. plus l \<circ> f = g"
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    32
proof -
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    33
  fix k
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    34
  from assms have "\<And>k. \<Delta> f k = \<Delta> g k" by simp
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    35
  then have k_incr: "\<And>k. f (k + 1) - g (k + 1) = f k - g k"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    36
    by (simp add: \<Delta>_def algebra_simps)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    37
  then have "\<And>k. f ((k - 1) + 1) - g ((k - 1) + 1) = f (k - 1) - g (k - 1)"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    38
    by blast
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    39
  then have k_decr: "\<And>k. f (k - 1) - g (k - 1) = f k - g k"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    40
    by simp
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    41
  have "\<And>k. f k - g k = f 0 - g 0"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    42
  proof -
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    43
    fix k
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    44
    show "f k - g k = f 0 - g 0"
36811
4ab4aa5bee1c renamed former Int.int_induct to Int.int_of_nat_induct, former Presburger.int_induct to Int.int_induct: is more conservative and more natural than the intermediate solution
haftmann
parents: 36808
diff changeset
    45
      by (induct k rule: int_induct) (simp_all add: k_incr k_decr)
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    46
  qed
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35163
diff changeset
    47
  then have "\<And>k. (plus (g 0 - f 0) \<circ> f) k = g k"
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    48
    by (simp add: algebra_simps)
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35163
diff changeset
    49
  then have "plus (g 0 - f 0) \<circ> f = g" ..
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    50
  then show ?thesis ..
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    51
qed
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    52
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    53
text {* The formal sum operator. *}
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    54
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    55
definition \<Sigma> :: "(int \<Rightarrow> 'a\<Colon>ab_group_add) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a" where
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    56
  "\<Sigma> f j l = (if j < l then setsum f {j..<l}
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    57
    else if j > l then - setsum f {l..<j}
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    58
    else 0)"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    59
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    60
lemma \<Sigma>_same [simp]:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    61
  "\<Sigma> f j j = 0"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    62
  by (simp add: \<Sigma>_def)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    63
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    64
lemma \<Sigma>_positive:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    65
  "j < l \<Longrightarrow> \<Sigma> f j l = setsum f {j..<l}"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    66
  by (simp add: \<Sigma>_def)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    67
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    68
lemma \<Sigma>_negative:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    69
  "j > l \<Longrightarrow> \<Sigma> f j l = - \<Sigma> f l j"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    70
  by (simp add: \<Sigma>_def)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    71
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    72
lemma add_\<Sigma>:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    73
  "\<Sigma> f j k + \<Sigma> f k l = \<Sigma> f j l"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    74
  by (simp add: \<Sigma>_def algebra_simps add_setsum_int)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    75
   (simp_all add: add_setsum_orient [of f k j l]
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    76
      add_setsum_orient [of f j l k]
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    77
      add_setsum_orient [of f j k l] add_setsum_int)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    78
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    79
lemma \<Sigma>_incr_upper:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    80
  "\<Sigma> f j (l + 1) = \<Sigma> f j l + f l"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    81
proof -
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    82
  have "{l..<l+1} = {l}" by auto
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    83
  then have "\<Sigma> f l (l + 1) = f l" by (simp add: \<Sigma>_def)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    84
  moreover have "\<Sigma> f j (l + 1) = \<Sigma> f j l + \<Sigma> f l (l + 1)" by (simp add: add_\<Sigma>)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    85
  ultimately show ?thesis by simp
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    86
qed
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    87
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    88
text {* Fundamental lemmas: The relation between @{term \<Delta>} and @{term \<Sigma>}. *}
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    89
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    90
lemma \<Delta>_\<Sigma>:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    91
  "\<Delta> (\<Sigma> f j) = f"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    92
proof
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    93
  fix k
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    94
  show "\<Delta> (\<Sigma> f j) k = f k"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    95
    by (simp add: \<Delta>_def \<Sigma>_incr_upper)
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    96
qed
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    97
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    98
lemma \<Sigma>_\<Delta>:
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
    99
  "\<Sigma> (\<Delta> f) j l = f l - f j"
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
   100
proof -
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
   101
  from \<Delta>_\<Sigma> have "\<Delta> (\<Sigma> (\<Delta> f) j) = \<Delta> f" .
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35163
diff changeset
   102
  then obtain k where "plus k \<circ> \<Sigma> (\<Delta> f) j = f" by (blast dest: \<Delta>_same_shift)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
   103
  then have "\<And>q. f q = k + \<Sigma> (\<Delta> f) j q" by (simp add: fun_eq_iff)
35163
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
   104
  then show ?thesis by simp
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
   105
qed
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
   106
2e0966d6f951 added simple theory about discrete summation
haftmann
parents:
diff changeset
   107
end