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