src/HOL/Isar_Examples/Summation.thy
author wenzelm
Tue, 10 Nov 2015 23:41:20 +0100
changeset 61626 c304402cc3df
parent 61541 846c72206207
child 61932 2e48182cc82c
permissions -rw-r--r--
recovered from a9c0572109af;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33026
8f35633c4922 modernized session Isar_Examples;
wenzelm
parents: 31758
diff changeset
     1
(*  Title:      HOL/Isar_Examples/Summation.thy
7443
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
     3
*)
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
     4
58882
6e2010ab8bd9 modernized header;
wenzelm
parents: 58614
diff changeset
     5
section \<open>Summing natural numbers\<close>
7443
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
     6
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
     7
theory Summation
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
     8
imports Main
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
     9
begin
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
    10
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    11
text \<open>Subsequently, we prove some summation laws of natural numbers
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    12
  (including odds, squares, and cubes). These examples demonstrate how plain
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    13
  natural deduction (including induction) may be combined with calculational
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    14
  proof.\<close>
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
    15
7761
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    16
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
    17
subsection \<open>Summation laws\<close>
7443
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
    18
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    19
text \<open>The sum of natural numbers \<open>0 + \<cdots> + n\<close> equals \<open>n \<times> (n + 1)/2\<close>.
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    20
  Avoiding formal reasoning about division we prove this equation multiplied
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    21
  by \<open>2\<close>.\<close>
7800
8ee919e42174 improved presentation;
wenzelm
parents: 7761
diff changeset
    22
8ee919e42174 improved presentation;
wenzelm
parents: 7761
diff changeset
    23
theorem sum_of_naturals:
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
    24
  "2 * (\<Sum>i::nat=0..n. i) = n * (n + 1)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    25
  (is "?P n" is "?S n = _")
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    26
proof (induct n)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    27
  show "?P 0" by simp
10146
wenzelm
parents: 10007
diff changeset
    28
next
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    29
  fix n have "?S (n + 1) = ?S n + 2 * (n + 1)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    30
    by simp
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    31
  also assume "?S n = n * (n + 1)"
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    32
  also have "\<dots> + 2 * (n + 1) = (n + 1) * (n + 2)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    33
    by simp
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    34
  finally show "?P (Suc n)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    35
    by simp
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    36
qed
7443
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
    37
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    38
text \<open>The above proof is a typical instance of mathematical induction. The
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    39
  main statement is viewed as some \<open>?P n\<close> that is split by the induction
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    40
  method into base case \<open>?P 0\<close>, and step case \<open>?P n \<Longrightarrow> ?P (Suc n)\<close> for
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    41
  arbitrary \<open>n\<close>.
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
    42
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    43
  The step case is established by a short calculation in forward manner.
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    44
  Starting from the left-hand side \<open>?S (n + 1)\<close> of the thesis, the final
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    45
  result is achieved by transformations involving basic arithmetic reasoning
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    46
  (using the Simplifier). The main point is where the induction hypothesis
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    47
  \<open>?S n = n \<times> (n + 1)\<close> is introduced in order to replace a certain subterm.
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    48
  So the ``transitivity'' rule involved here is actual \<^emph>\<open>substitution\<close>. Also
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    49
  note how the occurrence of ``\dots'' in the subsequent step documents the
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    50
  position where the right-hand side of the hypothesis got filled in.
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
    51
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    52
  \<^medskip>
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    53
  A further notable point here is integration of calculations with plain
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    54
  natural deduction. This works so well in Isar for two reasons.
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
    55
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    56
    \<^enum> Facts involved in \isakeyword{also}~/ \isakeyword{finally}
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    57
    calculational chains may be just anything. There is nothing special
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    58
    about \isakeyword{have}, so the natural deduction element
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    59
    \isakeyword{assume} works just as well.
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    60
  
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    61
    \<^enum> There are two \<^emph>\<open>separate\<close> primitives for building natural deduction
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    62
    contexts: \isakeyword{fix}~\<open>x\<close> and \isakeyword{assume}~\<open>A\<close>. Thus it is
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    63
    possible to start reasoning with some new ``arbitrary, but fixed''
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    64
    elements before bringing in the actual assumption. In contrast, natural
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    65
    deduction is occasionally formalized with basic context elements of the
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    66
    form \<open>x:A\<close> instead.
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
    67
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    68
  \<^medskip>
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    69
  We derive further summation laws for odds, squares, and cubes as follows.
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    70
  The basic technique of induction plus calculation is the same as before.\<close>
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
    71
7800
8ee919e42174 improved presentation;
wenzelm
parents: 7761
diff changeset
    72
theorem sum_of_odds:
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
    73
  "(\<Sum>i::nat=0..<n. 2 * i + 1) = n^Suc (Suc 0)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    74
  (is "?P n" is "?S n = _")
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    75
proof (induct n)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    76
  show "?P 0" by simp
10146
wenzelm
parents: 10007
diff changeset
    77
next
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    78
  fix n
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    79
  have "?S (n + 1) = ?S n + 2 * n + 1"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    80
    by simp
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10672
diff changeset
    81
  also assume "?S n = n^Suc (Suc 0)"
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    82
  also have "\<dots> + 2 * n + 1 = (n + 1)^Suc (Suc 0)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    83
    by simp
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    84
  finally show "?P (Suc n)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    85
    by simp
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    86
qed
7443
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
    87
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    88
text \<open>Subsequently we require some additional tweaking of Isabelle built-in
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
    89
  arithmetic simplifications, such as bringing in distributivity by hand.\<close>
8814
0a5edcbe0695 adapted to new arithmetic simprocs;
wenzelm
parents: 8659
diff changeset
    90
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    91
lemmas distrib = add_mult_distrib add_mult_distrib2
8814
0a5edcbe0695 adapted to new arithmetic simprocs;
wenzelm
parents: 8659
diff changeset
    92
7761
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    93
theorem sum_of_squares:
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
    94
  "6 * (\<Sum>i::nat=0..n. i^Suc (Suc 0)) = n * (n + 1) * (2 * n + 1)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    95
  (is "?P n" is "?S n = _")
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    96
proof (induct n)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    97
  show "?P 0" by simp
10146
wenzelm
parents: 10007
diff changeset
    98
next
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
    99
  fix n
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   100
  have "?S (n + 1) = ?S n + 6 * (n + 1)^Suc (Suc 0)"
18193
54419506df9e tuned document;
wenzelm
parents: 15912
diff changeset
   101
    by (simp add: distrib)
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   102
  also assume "?S n = n * (n + 1) * (2 * n + 1)"
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   103
  also have "\<dots> + 6 * (n + 1)^Suc (Suc 0) =
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   104
      (n + 1) * (n + 2) * (2 * (n + 1) + 1)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   105
    by (simp add: distrib)
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   106
  finally show "?P (Suc n)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   107
    by simp
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   108
qed
7443
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
   109
7800
8ee919e42174 improved presentation;
wenzelm
parents: 7761
diff changeset
   110
theorem sum_of_cubes:
15561
045a07ac35a7 another reorganization of setsums and intervals
nipkow
parents: 15043
diff changeset
   111
  "4 * (\<Sum>i::nat=0..n. i^3) = (n * (n + 1))^Suc (Suc 0)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   112
  (is "?P n" is "?S n = _")
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   113
proof (induct n)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   114
  show "?P 0" by (simp add: power_eq_if)
10146
wenzelm
parents: 10007
diff changeset
   115
next
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   116
  fix n
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   117
  have "?S (n + 1) = ?S n + 4 * (n + 1)^3"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   118
    by (simp add: power_eq_if distrib)
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10672
diff changeset
   119
  also assume "?S n = (n * (n + 1))^Suc (Suc 0)"
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   120
  also have "\<dots> + 4 * (n + 1)^3 = ((n + 1) * ((n + 1) + 1))^Suc (Suc 0)"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   121
    by (simp add: power_eq_if distrib)
55640
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   122
  finally show "?P (Suc n)"
abc140f21caa tuned proofs;
wenzelm
parents: 50086
diff changeset
   123
    by simp
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   124
qed
7443
e5356e73f57a renamed NatSum to Summation;
wenzelm
parents:
diff changeset
   125
58614
7338eb25226c more cartouches;
wenzelm
parents: 55640
diff changeset
   126
text \<open>Note that in contrast to older traditions of tactical proof
50086
ecffea78d381 tuned -- eliminated obsolete citation of isabelle-ref;
wenzelm
parents: 40880
diff changeset
   127
  scripts, the structured proof applies induction on the original,
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   128
  unsimplified statement. This allows to state the induction cases robustly
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   129
  and conveniently. Simplification (or other automated) methods are then
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   130
  applied in terminal position to solve certain sub-problems completely.
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
   131
61541
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   132
  As a general rule of good proof style, automatic methods such as \<open>simp\<close> or
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   133
  \<open>auto\<close> should normally be never used as initial proof methods with a
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   134
  nested sub-proof to address the automatically produced situation, but only
846c72206207 tuned document;
wenzelm
parents: 58882
diff changeset
   135
  as terminal ones to solve sub-problems.\<close>
7968
964b65b4e433 improved presentation;
wenzelm
parents: 7869
diff changeset
   136
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   137
end