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