doc-src/IsarOverview/Isar/Induction.thy
author nipkow
Mon, 02 May 2005 18:29:29 +0200
changeset 15909 5f0c8a3f0226
parent 13999 454a2ad0c381
child 16044 6887e6d12a94
permissions -rw-r--r--
fixed setsum problem
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     1
(*<*)theory Induction = Main:(*>*)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     2
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     3
section{*Case distinction and induction \label{sec:Induct}*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     4
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     5
text{* Computer science applications abound with inductively defined
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     6
structures, which is why we treat them in more detail. HOL already
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     7
comes with a datatype of lists with the two constructors @{text Nil}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     8
and @{text Cons}. @{text Nil} is written @{term"[]"} and @{text"Cons x
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
     9
xs"} is written @{term"x # xs"}.  *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    10
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    11
subsection{*Case distinction\label{sec:CaseDistinction}*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    12
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    13
text{* We have already met the @{text cases} method for performing
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    14
binary case splits. Here is another example: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    15
lemma "\<not> A \<or> A"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    16
proof cases
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    17
  assume "A" thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    18
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    19
  assume "\<not> A" thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    20
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    21
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    22
text{*\noindent The two cases must come in this order because @{text
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    23
cases} merely abbreviates @{text"(rule case_split_thm)"} where
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    24
@{thm[source] case_split_thm} is @{thm case_split_thm}. If we reverse
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    25
the order of the two cases in the proof, the first case would prove
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    26
@{prop"\<not> A \<Longrightarrow> \<not> A \<or> A"} which would solve the first premise of
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    27
@{thm[source] case_split_thm}, instantiating @{text ?P} with @{term "\<not>
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    28
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
    29
Therefore the order of subgoals is not always completely arbitrary.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    30
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    31
The above proof is appropriate if @{term A} is textually small.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    32
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
    33
be avoided by the following idiom *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    34
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    35
lemma "\<not> A \<or> A"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    36
proof (cases "A")
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    37
  case True thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    38
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    39
  case False thus ?thesis ..
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    40
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    41
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    42
text{*\noindent which is like the previous proof but instantiates
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    43
@{text ?P} right away with @{term A}. Thus we could prove the two
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    44
cases in any order. The phrase `\isakeyword{case}~@{text True}'
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    45
abbreviates `\isakeyword{assume}~@{text"True: A"}' and analogously for
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    46
@{text"False"} and @{prop"\<not>A"}.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    47
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    48
The same game can be played with other datatypes, for example lists,
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    49
where @{term tl} is the tail of a list, and @{text length} returns a
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    50
natural number (remember: $0-1=0$):
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    51
*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    52
(*<*)declare length_tl[simp del](*>*)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    53
lemma "length(tl xs) = length xs - 1"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    54
proof (cases xs)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    55
  case Nil thus ?thesis by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    56
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    57
  case Cons thus ?thesis by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    58
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    59
text{*\noindent Here `\isakeyword{case}~@{text Nil}' abbreviates
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    60
`\isakeyword{assume}~@{text"Nil:"}~@{prop"xs = []"}' and
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    61
`\isakeyword{case}~@{text Cons}'
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    62
abbreviates `\isakeyword{fix}~@{text"? ??"}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    63
\isakeyword{assume}~@{text"Cons:"}~@{text"xs = ? # ??"}'
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    64
where @{text"?"} and @{text"??"}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    65
stand for variable names that have been chosen by the system.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    66
Therefore we cannot refer to them.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    67
Luckily, this proof is simple enough we do not need to refer to them.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    68
However, sometimes one may have to. Hence Isar offers a simple scheme for
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    69
naming those variables: replace the anonymous @{text Cons} by
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    70
@{text"(Cons y ys)"}, which abbreviates `\isakeyword{fix}~@{text"y ys"}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    71
\isakeyword{assume}~@{text"Cons:"}~@{text"xs = y # ys"}'.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    72
In each \isakeyword{case} the assumption can be
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    73
referred to inside the proof by the name of the constructor. In
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    74
Section~\ref{sec:full-Ind} below we will come across an example
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    75
of this. *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    76
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    77
subsection{*Structural induction*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    78
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    79
text{* We start with an inductive proof where both cases are proved automatically: *}
15909
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    80
lemma "2 * (\<Sum>i::nat = 0..<n+1. i) = n*(n+1)"
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    81
by (induct n, simp_all)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    82
15909
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    83
text{*\noindent The constraint @{text"::nat"} is needed because all of
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    84
the operations involved are overloaded.
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    85
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    86
If we want to expose more of the structure of the
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    87
proof, we can use pattern matching to avoid having to repeat the goal
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    88
statement: *}
15909
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
    89
lemma "2 * (\<Sum>i::nat = 0..<n+1. i) = n*(n+1)" (is "?P n")
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    90
proof (induct n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    91
  show "?P 0" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    92
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    93
  fix n assume "?P n"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    94
  thus "?P(Suc n)" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    95
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    96
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    97
text{* \noindent We could refine this further to show more of the equational
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    98
proof. Instead we explore the same avenue as for case distinctions:
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
    99
introducing context via the \isakeyword{case} command: *}
15909
5f0c8a3f0226 fixed setsum problem
nipkow
parents: 13999
diff changeset
   100
lemma "2 * (\<Sum>i::nat = 0..<n+1. i) = n*(n+1)"
13999
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   101
proof (induct n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   102
  case 0 show ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   103
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   104
  case Suc thus ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   105
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   106
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   107
text{* \noindent The implicitly defined @{text ?case} refers to the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   108
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
   109
@{text"?P(Suc n)"} in the second case. Context \isakeyword{case}~@{text 0} is
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   110
empty whereas \isakeyword{case}~@{text Suc} assumes @{text"?P n"}. Again we
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   111
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
   112
in the induction step because it has not been introduced via \isakeyword{fix}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   113
(in contrast to the previous proof). The solution is the one outlined for
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   114
@{text Cons} above: replace @{term Suc} by @{text"(Suc i)"}: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   115
lemma fixes n::nat shows "n < n*n + 1"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   116
proof (induct n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   117
  case 0 show ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   118
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   119
  case (Suc i) thus "Suc i < Suc i * Suc i + 1" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   120
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   121
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   122
text{* \noindent Of course we could again have written
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   123
\isakeyword{thus}~@{text ?case} instead of giving the term explicitly
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   124
but we wanted to use @{term i} somewhere. *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   125
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   126
subsection{*Induction formulae involving @{text"\<And>"} or @{text"\<Longrightarrow>"}\label{sec:full-Ind}*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   127
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   128
text{* Let us now consider the situation where the goal to be proved contains
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   129
@{text"\<And>"} or @{text"\<Longrightarrow>"}, say @{prop"\<And>x. P x \<Longrightarrow> Q x"} --- motivation and a
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   130
real example follow shortly.  This means that in each case of the induction,
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   131
@{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
   132
first proof steps will be the canonical ones, fixing @{text x} and assuming
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   133
@{prop"P' x"}. To avoid this tedium, induction performs these steps
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   134
automatically: for example in case @{text"(Suc n)"}, @{text ?case} is only
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   135
@{prop"Q' x"} whereas the assumptions (named @{term Suc}!) contain both the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   136
usual induction hypothesis \emph{and} @{prop"P' x"}.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   137
It should be clear how this generalises to more complex formulae.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   138
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   139
As an example we will now prove complete induction via
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   140
structural induction. *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   141
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   142
lemma assumes A: "(\<And>n. (\<And>m. m < n \<Longrightarrow> P m) \<Longrightarrow> P n)"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   143
  shows "P(n::nat)"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   144
proof (rule A)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   145
  show "\<And>m. m < n \<Longrightarrow> P m"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   146
  proof (induct n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   147
    case 0 thus ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   148
  next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   149
    case (Suc n)   -- {*\isakeyword{fix} @{term m} \isakeyword{assume} @{text Suc}: @{text[source]"?m < n \<Longrightarrow> P ?m"} @{prop[source]"m < Suc n"}*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   150
    show ?case    -- {*@{term ?case}*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   151
    proof cases
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   152
      assume eq: "m = n"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   153
      from Suc and A have "P n" by blast
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   154
      with eq show "P m" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   155
    next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   156
      assume "m \<noteq> n"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   157
      with Suc have "m < n" by arith
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   158
      thus "P m" by(rule Suc)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   159
    qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   160
  qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   161
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   162
text{* \noindent Given the explanations above and the comments in the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   163
proof text (only necessary for novices), the proof should be quite
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   164
readable.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   165
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   166
The statement of the lemma is interesting because it deviates from the style in
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   167
the Tutorial~\cite{LNCS2283}, which suggests to introduce @{text"\<forall>"} or
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   168
@{text"\<longrightarrow>"} into a theorem to strengthen it for induction. In Isar
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   169
proofs we can use @{text"\<And>"} and @{text"\<Longrightarrow>"} instead. This simplifies the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   170
proof and means we do not have to convert between the two kinds of
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   171
connectives.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   172
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   173
Note that in a nested induction over the same data type, the inner
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   174
case labels hide the outer ones of the same name. If you want to refer
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   175
to the outer ones inside, you need to name them on the outside, e.g.\
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   176
\isakeyword{note}~@{text"outer_IH = Suc"}.  *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   177
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   178
subsection{*Rule induction*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   179
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   180
text{* HOL also supports inductively defined sets. See \cite{LNCS2283}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   181
for details. As an example we define our own version of the reflexive
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   182
transitive closure of a relation --- HOL provides a predefined one as well.*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   183
consts rtc :: "('a \<times> 'a)set \<Rightarrow> ('a \<times> 'a)set"   ("_*" [1000] 999)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   184
inductive "r*"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   185
intros
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   186
refl:  "(x,x) \<in> r*"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   187
step:  "\<lbrakk> (x,y) \<in> r; (y,z) \<in> r* \<rbrakk> \<Longrightarrow> (x,z) \<in> r*"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   188
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   189
text{* \noindent
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   190
First the constant is declared as a function on binary
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   191
relations (with concrete syntax @{term"r*"} instead of @{text"rtc
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   192
r"}), then the defining clauses are given. We will now prove that
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   193
@{term"r*"} is indeed transitive: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   194
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   195
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
   196
using A
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   197
proof induct
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   198
  case refl thus ?case .
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   199
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   200
  case step thus ?case by(blast intro: rtc.step)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   201
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   202
text{*\noindent Rule induction is triggered by a fact $(x_1,\dots,x_n)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   203
\in R$ piped into the proof, here \isakeyword{using}~@{text A}. The
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   204
proof itself follows the inductive definition very
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   205
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
   206
the rule, analogous to structural induction.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   207
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   208
However, this proof is rather terse. Here is a more readable version:
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   209
*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   210
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   211
lemma assumes A: "(x,y) \<in> r*" and B: "(y,z) \<in> r*"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   212
  shows "(x,z) \<in> r*"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   213
proof -
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   214
  from A B show ?thesis
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   215
  proof induct
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   216
    fix x assume "(x,z) \<in> r*"  -- {*@{text B}[@{text y} := @{text x}]*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   217
    thus "(x,z) \<in> r*" .
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   218
  next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   219
    fix x' x y
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   220
    assume 1: "(x',x) \<in> r" and
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   221
           IH: "(y,z) \<in> r* \<Longrightarrow> (x,z) \<in> r*" and
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   222
           B:  "(y,z) \<in> r*"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   223
    from 1 IH[OF B] show "(x',z) \<in> r*" by(rule rtc.step)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   224
  qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   225
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   226
text{*\noindent We start the proof with \isakeyword{from}~@{text"A
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   227
B"}. Only @{text A} is ``consumed'' by the induction step.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   228
Since @{text B} is left over we don't just prove @{text
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   229
?thesis} but @{text"B \<Longrightarrow> ?thesis"}, just as in the previous proof. The
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   230
base case is trivial. In the assumptions for the induction step we can
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   231
see very clearly how things fit together and permit ourselves the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   232
obvious forward step @{text"IH[OF B]"}.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   233
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   234
The notation `\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}'
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   235
is also supported for inductive definitions. The \emph{constructor} is (the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   236
name of) the rule and the \emph{vars} fix the free variables in the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   237
rule; the order of the \emph{vars} must correspond to the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   238
\emph{alphabetical order} of the variables as they appear in the rule.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   239
For example, we could start the above detailed proof of the induction
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   240
with \isakeyword{case}~\isa{(step x' x y)}. However, we can then only
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   241
refer to the assumptions named \isa{step} collectively and not
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   242
individually, as the above proof requires.  *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   243
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   244
subsection{*More induction*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   245
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   246
text{* We close the section by demonstrating how arbitrary induction
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   247
rules are applied. As a simple example we have chosen recursion
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   248
induction, i.e.\ induction based on a recursive function
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   249
definition. However, most of what we show works for induction in
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   250
general.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   251
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   252
The example is an unusual definition of rotation: *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   253
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   254
consts rot :: "'a list \<Rightarrow> 'a list"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   255
recdef rot "measure length"  --"for the internal termination proof"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   256
"rot [] = []"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   257
"rot [x] = [x]"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   258
"rot (x#y#zs) = y # rot(x#zs)"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   259
text{*\noindent This yields, among other things, the induction rule
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   260
@{thm[source]rot.induct}: @{thm[display]rot.induct[no_vars]}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   261
In the following proof we rely on a default naming scheme for cases: they are
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   262
called 1, 2, etc, unless they have been named explicitly. The latter happens
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   263
only with datatypes and inductively defined sets, but not with recursive
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   264
functions. *}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   265
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   266
lemma "xs \<noteq> [] \<Longrightarrow> rot xs = tl xs @ [hd xs]"
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   267
proof (induct xs rule: rot.induct)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   268
  case 1 thus ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   269
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   270
  case 2 show ?case by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   271
next
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   272
  case (3 a b cs)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   273
  have "rot (a # b # cs) = b # rot(a # cs)" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   274
  also have "\<dots> = b # tl(a # cs) @ [hd(a # cs)]" by(simp add:3)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   275
  also have "\<dots> = tl (a # b # cs) @ [hd (a # b # cs)]" by simp
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   276
  finally show ?case .
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   277
qed
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   278
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   279
text{*\noindent
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   280
The third case is only shown in gory detail (see \cite{BauerW-TPHOLs01}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   281
for how to reason with chains of equations) to demonstrate that the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   282
`\isakeyword{case}~\isa{(}\emph{constructor} \emph{vars}\isa{)}' notation also
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   283
works for arbitrary induction theorems with numbered cases. The order
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   284
of the \emph{vars} corresponds to the order of the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   285
@{text"\<And>"}-quantified variables in each case of the induction
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   286
theorem. For induction theorems produced by \isakeyword{recdef} it is
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   287
the order in which the variables appear on the left-hand side of the
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   288
equation.
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   289
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   290
The proof is so simple that it can be condensed to
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   291
*}
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   292
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   293
(*<*)lemma "xs \<noteq> [] \<Longrightarrow> rot xs = tl xs @ [hd xs]"(*>*)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   294
by (induct xs rule: rot.induct, simp_all)
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   295
454a2ad0c381 IsarOverview moved one level up
kleing
parents:
diff changeset
   296
(*<*)end(*>*)