doc-src/IsarOverview/Isar/Induction.thy
author wenzelm
Sat, 16 Oct 2010 20:27:35 +0100
changeset 39857 ea93e088398d
parent 30649 57753e0ec1d4
permissions -rw-r--r--
more examples;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
     1
(*<*)theory Induction imports Main begin
359b179fc963 updates
nipkow
parents: 23733
diff changeset
     2
fun itrev where
359b179fc963 updates
nipkow
parents: 23733
diff changeset
     3
"itrev [] ys = ys" |
359b179fc963 updates
nipkow
parents: 23733
diff changeset
     4
"itrev (x#xs) ys = itrev xs (x#ys)"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
     5
(*>*)
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     6
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     7
section{*Case distinction and induction \label{sec:Induct}*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     8
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     9
text{* Computer science applications abound with inductively defined
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    10
structures, which is why we treat them in more detail. HOL already
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    11
comes with a datatype of lists with the two constructors @{text Nil}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    12
and @{text Cons}. @{text Nil} is written @{term"[]"} and @{text"Cons x
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    13
xs"} is written @{term"x # xs"}.  *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    14
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    15
subsection{*Case distinction\label{sec:CaseDistinction}*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    16
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    17
text{* We have already met the @{text cases} method for performing
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    18
binary case splits. Here is another example: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    19
lemma "\<not> A \<or> A"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    20
proof cases
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    21
  assume "A" thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    22
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    23
  assume "\<not> A" thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    24
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    25
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    26
text{*\noindent The two cases must come in this order because @{text
27115
0dcafa5c9e3f eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents: 25427
diff changeset
    27
cases} merely abbreviates @{text"(rule case_split)"} where
0dcafa5c9e3f eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents: 25427
diff changeset
    28
@{thm[source] case_split} is @{thm case_split}. If we reverse
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    29
the order of the two cases in the proof, the first case would prove
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    30
@{prop"\<not> A \<Longrightarrow> \<not> A \<or> A"} which would solve the first premise of
27115
0dcafa5c9e3f eliminated obsolete case_split_thm -- use case_split;
wenzelm
parents: 25427
diff changeset
    31
@{thm[source] case_split}, instantiating @{text ?P} with @{term "\<not>
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    32
A"}, thus making the second premise @{prop"\<not> \<not> A \<Longrightarrow> \<not> A \<or> A"}.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    33
Therefore the order of subgoals is not always completely arbitrary.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    34
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    35
The above proof is appropriate if @{term A} is textually small.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    36
However, if @{term A} is large, we do not want to repeat it. This can
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    37
be avoided by the following idiom *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    38
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    39
lemma "\<not> A \<or> A"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    40
proof (cases "A")
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    41
  case True thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    42
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    43
  case False thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    44
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    45
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    46
text{*\noindent which is like the previous proof but instantiates
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    47
@{text ?P} right away with @{term A}. Thus we could prove the two
25412
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    48
cases in any order. The phrase \isakeyword{case}~@{text True}
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    49
abbreviates \isakeyword{assume}~@{text"True: A"} and analogously for
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    50
@{text"False"} and @{prop"\<not>A"}.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    51
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    52
The same game can be played with other datatypes, for example lists,
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    53
where @{term tl} is the tail of a list, and @{text length} returns a
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    54
natural number (remember: $0-1=0$):
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    55
*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    56
(*<*)declare length_tl[simp del](*>*)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    57
lemma "length(tl xs) = length xs - 1"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    58
proof (cases xs)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    59
  case Nil thus ?thesis by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    60
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    61
  case Cons thus ?thesis by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    62
qed
25412
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    63
text{*\noindent Here \isakeyword{case}~@{text Nil} abbreviates
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    64
\isakeyword{assume}~@{text"Nil:"}~@{prop"xs = []"} and
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    65
\isakeyword{case}~@{text Cons} abbreviates \isakeyword{fix}~@{text"? ??"}
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    66
\isakeyword{assume}~@{text"Cons:"}~@{text"xs = ? # ??"},
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    67
where @{text"?"} and @{text"??"}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    68
stand for variable names that have been chosen by the system.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    69
Therefore we cannot refer to them.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    70
Luckily, this proof is simple enough we do not need to refer to them.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    71
However, sometimes one may have to. Hence Isar offers a simple scheme for
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    72
naming those variables: replace the anonymous @{text Cons} by
25412
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    73
@{text"(Cons y ys)"}, which abbreviates \isakeyword{fix}~@{text"y ys"}
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
    74
\isakeyword{assume}~@{text"Cons:"}~@{text"xs = y # ys"}.
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    75
In each \isakeyword{case} the assumption can be
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    76
referred to inside the proof by the name of the constructor. In
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    77
Section~\ref{sec:full-Ind} below we will come across an example
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
    78
of this.
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    79
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
    80
\subsection{Structural induction}
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    81
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
    82
We start with an inductive proof where both cases are proved automatically: *}
16522
nipkow
parents: 16044
diff changeset
    83
lemma "2 * (\<Sum>i::nat\<le>n. i) = n*(n+1)"
25427
8ba39d2d9d0b updated
nipkow
parents: 25412
diff changeset
    84
by (induct n) simp_all
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    85
15909
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    86
text{*\noindent The constraint @{text"::nat"} is needed because all of
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    87
the operations involved are overloaded.
25427
8ba39d2d9d0b updated
nipkow
parents: 25412
diff changeset
    88
This proof also demonstrates that \isakeyword{by} can take two arguments,
8ba39d2d9d0b updated
nipkow
parents: 25412
diff changeset
    89
one to start and one to finish the proof --- the latter is optional.
15909
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    90
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    91
If we want to expose more of the structure of the
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    92
proof, we can use pattern matching to avoid having to repeat the goal
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    93
statement: *}
16522
nipkow
parents: 16044
diff changeset
    94
lemma "2 * (\<Sum>i::nat\<le>n. i) = n*(n+1)" (is "?P n")
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    95
proof (induct n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    96
  show "?P 0" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    97
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    98
  fix n assume "?P n"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    99
  thus "?P(Suc n)" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   100
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   101
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   102
text{* \noindent We could refine this further to show more of the equational
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   103
proof. Instead we explore the same avenue as for case distinctions:
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   104
introducing context via the \isakeyword{case} command: *}
16522
nipkow
parents: 16044
diff changeset
   105
lemma "2 * (\<Sum>i::nat \<le> n. i) = n*(n+1)"
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   106
proof (induct n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   107
  case 0 show ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   108
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   109
  case Suc thus ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   110
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   111
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   112
text{* \noindent The implicitly defined @{text ?case} refers to the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   113
corresponding case to be proved, i.e.\ @{text"?P 0"} in the first case and
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   114
@{text"?P(Suc n)"} in the second case. Context \isakeyword{case}~@{text 0} is
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   115
empty whereas \isakeyword{case}~@{text Suc} assumes @{text"?P n"}. Again we
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   116
have the same problem as with case distinctions: we cannot refer to an anonymous @{term n}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   117
in the induction step because it has not been introduced via \isakeyword{fix}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   118
(in contrast to the previous proof). The solution is the one outlined for
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   119
@{text Cons} above: replace @{term Suc} by @{text"(Suc i)"}: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   120
lemma fixes n::nat shows "n < n*n + 1"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   121
proof (induct n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   122
  case 0 show ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   123
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   124
  case (Suc i) thus "Suc i < Suc i * Suc i + 1" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   125
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   126
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   127
text{* \noindent Of course we could again have written
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   128
\isakeyword{thus}~@{text ?case} instead of giving the term explicitly
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   129
but we wanted to use @{term i} somewhere.
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   130
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   131
\subsection{Generalization via @{text arbitrary}}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   132
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   133
It is frequently necessary to generalize a claim before it becomes
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   134
provable by induction. The tutorial~\cite{LNCS2283} demonstrates this
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   135
with @{prop"itrev xs ys = rev xs @ ys"}, where @{text ys}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   136
needs to be universally quantified before induction succeeds.\footnote{@{thm rev.simps(1)},\quad @{thm rev.simps(2)[no_vars]},\\ @{thm itrev.simps(1)[no_vars]},\quad @{thm itrev.simps(2)[no_vars]}} But
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   137
strictly speaking, this quantification step is already part of the
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   138
proof and the quantifiers should not clutter the original claim. This
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   139
is how the quantification step can be combined with induction: *}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   140
lemma "itrev xs ys = rev xs @ ys"
25427
8ba39d2d9d0b updated
nipkow
parents: 25412
diff changeset
   141
by (induct xs arbitrary: ys) simp_all
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   142
text{*\noindent The annotation @{text"arbitrary:"}~\emph{vars}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   143
universally quantifies all \emph{vars} before the induction.  Hence
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   144
they can be replaced by \emph{arbitrary} values in the proof.
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   145
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 27115
diff changeset
   146
Generalization via @{text"arbitrary"} is particularly convenient
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 27115
diff changeset
   147
if the induction step is a structured proof as opposed to the automatic
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 27115
diff changeset
   148
example above. Then the claim is available in unquantified form but
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   149
with the generalized variables replaced by @{text"?"}-variables, ready
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 27115
diff changeset
   150
for instantiation. In the above example, in the @{const[source] Cons} case the
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 27115
diff changeset
   151
induction hypothesis is @{text"itrev xs ?ys = rev xs @ ?ys"} (available
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 27115
diff changeset
   152
under the name @{const[source] Cons}).
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   153
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   154
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   155
\subsection{Inductive proofs of conditional formulae}
25412
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
   156
\label{sec:full-Ind}
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   157
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   158
Induction also copes well with formulae involving @{text"\<Longrightarrow>"}, for example
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   159
*}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   160
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   161
lemma "xs \<noteq> [] \<Longrightarrow> hd(rev xs) = last xs"
25427
8ba39d2d9d0b updated
nipkow
parents: 25412
diff changeset
   162
by (induct xs) simp_all
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   163
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   164
text{*\noindent This is an improvement over that style the
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   165
tutorial~\cite{LNCS2283} advises, which requires @{text"\<longrightarrow>"}.
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   166
A further improvement is shown in the following proof:
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   167
*}
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   168
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   169
lemma  "map f xs = map f ys \<Longrightarrow> length xs = length ys"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   170
proof (induct ys arbitrary: xs)
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   171
  case Nil thus ?case by simp
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   172
next
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   173
  case (Cons y ys)  note Asm = Cons
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   174
  show ?case
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   175
  proof (cases xs)
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   176
    case Nil
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   177
    hence False using Asm(2) by simp
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   178
    thus ?thesis ..
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   179
  next
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   180
    case (Cons x xs')
25427
8ba39d2d9d0b updated
nipkow
parents: 25412
diff changeset
   181
    with Asm(2) have "map f xs' = map f ys" by simp
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   182
    from Asm(1)[OF this] `xs = x#xs'` show ?thesis by simp
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   183
  qed
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   184
qed
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   185
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   186
text{*\noindent
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   187
The base case is trivial. In the step case Isar assumes
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   188
(under the name @{text Cons}) two propositions:
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   189
\begin{center}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   190
\begin{tabular}{l}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   191
@{text"map f ?xs = map f ys \<Longrightarrow> length ?xs = length ys"}\\
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   192
@{prop"map f xs = map f (y # ys)"}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   193
\end{tabular}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   194
\end{center}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   195
The first is the induction hypothesis, the second, and this is new,
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   196
is the premise of the induction step. The actual goal at this point is merely
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   197
@{prop"length xs = length (y#ys)"}. The assumptions are given the new name
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   198
@{text Asm} to avoid a name clash further down. The proof procedes with a case distinction on @{text xs}. In the case @{prop"xs = []"}, the second of our two
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   199
assumptions (@{text"Asm(2)"}) implies the contradiction @{text"0 = Suc(\<dots>)"}.
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   200
 In the case @{prop"xs = x#xs'"}, we first obtain
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   201
@{prop"map f xs' = map f ys"}, from which a forward step with the first assumption (@{text"Asm(1)[OF this]"}) yields @{prop"length xs' = length ys"}. Together
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   202
with @{prop"xs = x#xs"} this yields the goal
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   203
@{prop"length xs = length (y#ys)"}.
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   204
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   205
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   206
\subsection{Induction formulae involving @{text"\<And>"} or @{text"\<Longrightarrow>"}}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   207
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   208
Let us now consider abstractly the situation where the goal to be proved
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   209
contains both @{text"\<And>"} and @{text"\<Longrightarrow>"}, say @{prop"\<And>x. P x \<Longrightarrow> Q x"}.
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   210
This means that in each case of the induction,
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   211
@{text ?case} would be of the form @{prop"\<And>x. P' x \<Longrightarrow> Q' x"}.  Thus the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   212
first proof steps will be the canonical ones, fixing @{text x} and assuming
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   213
@{prop"P' x"}. To avoid this tedium, induction performs the canonical steps
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   214
automatically: in each step case, the assumptions contain both the
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   215
usual induction hypothesis and @{prop"P' x"}, whereas @{text ?case} is only
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   216
@{prop"Q' x"}.
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   217
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   218
\subsection{Rule induction}
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   219
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   220
HOL also supports inductively defined sets. See \cite{LNCS2283}
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   221
for details. As an example we define our own version of the reflexive
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   222
transitive closure of a relation --- HOL provides a predefined one as well.*}
23733
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 17914
diff changeset
   223
inductive_set
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 17914
diff changeset
   224
  rtc :: "('a \<times> 'a)set \<Rightarrow> ('a \<times> 'a)set"   ("_*" [1000] 999)
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 17914
diff changeset
   225
  for r :: "('a \<times> 'a)set"
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 17914
diff changeset
   226
where
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 17914
diff changeset
   227
  refl:  "(x,x) \<in> r*"
3f8ad7418e55 Adapted to new inductive definition package.
berghofe
parents: 17914
diff changeset
   228
| step:  "\<lbrakk> (x,y) \<in> r; (y,z) \<in> r* \<rbrakk> \<Longrightarrow> (x,z) \<in> r*"
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   229
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   230
text{* \noindent
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   231
First the constant is declared as a function on binary
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   232
relations (with concrete syntax @{term"r*"} instead of @{text"rtc
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   233
r"}), then the defining clauses are given. We will now prove that
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   234
@{term"r*"} is indeed transitive: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   235
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   236
lemma assumes A: "(x,y) \<in> r*" shows "(y,z) \<in> r* \<Longrightarrow> (x,z) \<in> r*"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   237
using A
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   238
proof induct
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   239
  case refl thus ?case .
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   240
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   241
  case step thus ?case by(blast intro: rtc.step)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   242
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   243
text{*\noindent Rule induction is triggered by a fact $(x_1,\dots,x_n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   244
\in R$ piped into the proof, here \isakeyword{using}~@{text A}. The
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   245
proof itself follows the inductive definition very
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   246
closely: there is one case for each rule, and it has the same name as
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   247
the rule, analogous to structural induction.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   248
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   249
However, this proof is rather terse. Here is a more readable version:
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   250
*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   251
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   252
lemma assumes "(x,y) \<in> r*" and "(y,z) \<in> r*" shows "(x,z) \<in> r*"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   253
using assms
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   254
proof induct
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   255
  fix x assume "(x,z) \<in> r*"  -- {*@{text B}[@{text y} := @{text x}]*}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   256
  thus "(x,z) \<in> r*" .
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   257
next
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   258
  fix x' x y
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   259
  assume 1: "(x',x) \<in> r" and
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   260
         IH: "(y,z) \<in> r* \<Longrightarrow> (x,z) \<in> r*" and
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   261
         B:  "(y,z) \<in> r*"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   262
  from 1 IH[OF B] show "(x',z) \<in> r*" by(rule rtc.step)
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   263
qed
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   264
text{*\noindent
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   265
This time, merely for a change, we start the proof with by feeding both
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   266
assumptions into the inductive proof. Only the first assumption is
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   267
``consumed'' by the induction.
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   268
Since the second one is left over we don't just prove @{text ?thesis} but
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   269
@{text"(y,z) \<in> r* \<Longrightarrow> ?thesis"}, just as in the previous proof.
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   270
The base case is trivial. In the assumptions for the induction step we can
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   271
see very clearly how things fit together and permit ourselves the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   272
obvious forward step @{text"IH[OF B]"}.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   273
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   274
The notation \isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   275
is also supported for inductive definitions. The \emph{constructor} is the
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   276
name of the rule and the \emph{vars} fix the free variables in the
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   277
rule; the order of the \emph{vars} must correspond to the
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   278
left-to-right order of the variables as they appear in the rule.
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   279
For example, we could start the above detailed proof of the induction
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   280
with \isakeyword{case}~\isa{(step x' x y)}. In that case we don't need
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   281
to spell out the assumptions but can refer to them by @{text"step(.)"},
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   282
although the resulting text will be quite cryptic.
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   283
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   284
\subsection{More induction}
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   285
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   286
We close the section by demonstrating how arbitrary induction
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   287
rules are applied. As a simple example we have chosen recursion
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   288
induction, i.e.\ induction based on a recursive function
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   289
definition. However, most of what we show works for induction in
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   290
general.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   291
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   292
The example is an unusual definition of rotation: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   293
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   294
fun rot :: "'a list \<Rightarrow> 'a list" where
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   295
"rot [] = []" |
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   296
"rot [x] = [x]" |
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   297
"rot (x#y#zs) = y # rot(x#zs)"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   298
text{*\noindent This yields, among other things, the induction rule
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   299
@{thm[source]rot.induct}: @{thm[display]rot.induct[no_vars]}
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   300
The following proof relies on a default naming scheme for cases: they are
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   301
called 1, 2, etc, unless they have been named explicitly. The latter happens
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   302
only with datatypes and inductively defined sets, but (usually)
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   303
not with recursive functions. *}
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   304
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   305
lemma "xs \<noteq> [] \<Longrightarrow> rot xs = tl xs @ [hd xs]"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   306
proof (induct xs rule: rot.induct)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   307
  case 1 thus ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   308
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   309
  case 2 show ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   310
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   311
  case (3 a b cs)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   312
  have "rot (a # b # cs) = b # rot(a # cs)" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   313
  also have "\<dots> = b # tl(a # cs) @ [hd(a # cs)]" by(simp add:3)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   314
  also have "\<dots> = tl (a # b # cs) @ [hd (a # b # cs)]" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   315
  finally show ?case .
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   316
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   317
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   318
text{*\noindent
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   319
The third case is only shown in gory detail (see \cite{BauerW-TPHOLs01}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   320
for how to reason with chains of equations) to demonstrate that the
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   321
\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)} notation also
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   322
works for arbitrary induction theorems with numbered cases. The order
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   323
of the \emph{vars} corresponds to the order of the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   324
@{text"\<And>"}-quantified variables in each case of the induction
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   325
theorem. For induction theorems produced by \isakeyword{fun} it is
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   326
the order in which the variables appear on the left-hand side of the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   327
equation.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   328
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   329
The proof is so simple that it can be condensed to
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   330
*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   331
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   332
(*<*)lemma "xs \<noteq> [] \<Longrightarrow> rot xs = tl xs @ [hd xs]"(*>*)
25412
6f56f0350f6c updates
nipkow
parents: 25403
diff changeset
   333
by (induct xs rule: rot.induct) simp_all
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   334
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   335
(*<*)end(*>*)
25403
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   336
(*
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   337
lemma assumes A: "(\<And>n. (\<And>m. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   338
  shows "P(n::nat)"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   339
proof (rule A)
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   340
  show "\<And>m. m < n \<Longrightarrow> P m"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   341
  proof (induct n)
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   342
    case 0 thus ?case by simp
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   343
  next
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   344
    case (Suc n)   -- {*\isakeyword{fix} @{term m} \isakeyword{assume} @{text Suc}: @{text[source]"?m < n \<Longrightarrow> P ?m"} @{prop[source]"m < Suc n"}*}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   345
    show ?case    -- {*@{term ?case}*}
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   346
    proof cases
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   347
      assume eq: "m = n"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   348
      from Suc and A have "P n" by blast
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   349
      with eq show "P m" by simp
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   350
    next
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   351
      assume "m \<noteq> n"
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   352
      with Suc have "m < n" by arith
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   353
      thus "P m" by(rule Suc)
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   354
    qed
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   355
  qed
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   356
qed
359b179fc963 updates
nipkow
parents: 23733
diff changeset
   357
*)