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