Theory Isar

(*<*)
theory Isar
imports LaTeXsugar
begin
declare [[quick_and_dirty]]
(*>*)
text‹
Apply-scripts are unreadable and hard to maintain. The language of choice
for larger proofs is \concept{Isar}. The two key features of Isar are:
\begin{itemize}
\item It is structured, not linear.
\item It is readable without its being run because
you need to state what you are proving at any given point.
\end{itemize}
Whereas apply-scripts are like assembly language programs, Isar proofs
are like structured programs with comments. A typical Isar proof looks like this:
›text‹
\begin{tabular}{@ {}l}
\isacom{proof}\\
\quad\isacom{assume} "›$\mathit{formula}_0$"›\\
\quad\isacom{have} "›$\mathit{formula}_1$"› \quad\isacom{by} simp›\\
\quad\vdots\\
\quad\isacom{have} "›$\mathit{formula}_n$"› \quad\isacom{by} blast›\\
\quad\isacom{show} "›$\mathit{formula}_{n+1}$"› \quad\isacom{by} …›\\
\isacom{qed}
\end{tabular}
›text‹
It proves $\mathit{formula}_0 \Longrightarrow \mathit{formula}_{n+1}$
(provided each proof step succeeds).
The intermediate \isacom{have} statements are merely stepping stones
on the way towards the \isacom{show} statement that proves the actual
goal. In more detail, this is the Isar core syntax:
\medskip

\begin{tabular}{@ {}lcl@ {}}
\textit{proof} &=& \indexed{\isacom{by}}{by} \textit{method}\\
      &$\mid$& \indexed{\isacom{proof}}{proof} [\textit{method}] \ \textit{step}$^*$ \ \indexed{\isacom{qed}}{qed}
\end{tabular}
\medskip

\begin{tabular}{@ {}lcl@ {}}
\textit{step} &=& \indexed{\isacom{fix}}{fix} \textit{variables} \\
      &$\mid$& \indexed{\isacom{assume}}{assume} \textit{proposition} \\
      &$\mid$& [\indexed{\isacom{from}}{from} \textit{fact}$^+$] (\indexed{\isacom{have}}{have} $\mid$ \indexed{\isacom{show}}{show}) \ \textit{proposition} \ \textit{proof}
\end{tabular}
\medskip

\begin{tabular}{@ {}lcl@ {}}
\textit{proposition} &=& [\textit{name}:] "›\textit{formula}"›
\end{tabular}
\medskip

\begin{tabular}{@ {}lcl@ {}}
\textit{fact} &=& \textit{name} \ $\mid$ \ \dots
\end{tabular}
\medskip

\noindent A proof can either be an atomic \isacom{by} with a single proof
method which must finish off the statement being proved, for example auto›,  or it can be a \isacom{proof}--\isacom{qed} block of multiple
steps. Such a block can optionally begin with a proof method that indicates
how to start off the proof, e.g., \mbox{(induction xs)›}.

A step either assumes a proposition or states a proposition
together with its proof. The optional \isacom{from} clause
indicates which facts are to be used in the proof.
Intermediate propositions are stated with \isacom{have}, the overall goal
is stated with \isacom{show}. A step can also introduce new local variables with
\isacom{fix}. Logically, \isacom{fix} introduces ⋀›-quantified
variables, \isacom{assume} introduces the assumption of an implication
(⟹›) and \isacom{have}/\isacom{show} introduce the conclusion.

Propositions are optionally named formulas. These names can be referred to in
later \isacom{from} clauses. In the simplest case, a fact is such a name.
But facts can also be composed with OF› and of› as shown in
\autoref{sec:forward-proof} --- hence the \dots\ in the above grammar.  Note
that assumptions, intermediate \isacom{have} statements and global lemmas all
have the same status and are thus collectively referred to as
\conceptidx{facts}{fact}.

Fact names can stand for whole lists of facts. For example, if f› is
defined by command \isacom{fun}, f.simps› refers to the whole list of
recursion equations defining f›. Individual facts can be selected by
writing f.simps(2)›, whole sublists by writing f.simps(2-4)›.


\section{Isar by Example}

We show a number of proofs of Cantor's theorem that a function from a set to
its powerset cannot be surjective, illustrating various features of Isar. The
constant constsurj is predefined.
›

lemma "¬ surj(f :: 'a  'a set)"
proof
  assume 0: "surj f"
  from 0 have 1: "A. a. A = f a" by(simp add: surj_def)
  from 1 have 2: "a. {x. x  f x} = f a" by blast
  from 2 show "False" by blast
qed

text‹
The \isacom{proof} command lacks an explicit method by which to perform
the proof. In such cases Isabelle tries to use some standard introduction
rule, in the above case for ¬›:
\[
\inferrule{
\mbox{@{thm (prem 1) notI}}}
{\mbox{@{thm (concl) notI}}}
\]
In order to prove prop~ P, assume P› and show False›.
Thus we may assume \mbox{\noquotes{@{prop [source] "surj f"}}}. The proof shows that names of propositions
may be (single!) digits --- meaningful names are hard to invent and are often
not necessary. Both \isacom{have} steps are obvious. The second one introduces
the diagonal set term{x. x  f x}, the key idea in the proof.
If you wonder why 2› directly implies False›: from 2›
it follows that propa  f a  a  f a.

\subsection{\indexed{this›}{this}, \indexed{\isacom{then}}{then}, \indexed{\isacom{hence}}{hence} and \indexed{\isacom{thus}}{thus}}

Labels should be avoided. They interrupt the flow of the reader who has to
scan the context for the point where the label was introduced. Ideally, the
proof is a linear flow, where the output of one step becomes the input of the
next step, piping the previously proved fact into the next proof, like
in a UNIX pipe. In such cases the predefined name this› can be used
to refer to the proposition proved in the previous step. This allows us to
eliminate all labels from our proof (we suppress the \isacom{lemma} statement):
›
(*<*)
lemma "¬ surj(f :: 'a  'a set)"
(*>*)
proof
  assume "surj f"
  from this have "a. {x. x  f x} = f a" by(auto simp: surj_def)
  from this show "False" by blast
qed

text‹We have also taken the opportunity to compress the two \isacom{have}
steps into one.

To compact the text further, Isar has a few convenient abbreviations:
\medskip

\begin{tabular}{r@ {\quad=\quad}l}
\isacom{then} & \isacom{from} this›\\
\isacom{thus} & \isacom{then} \isacom{show}\\
\isacom{hence} & \isacom{then} \isacom{have}
\end{tabular}
\medskip

\noindent
With the help of these abbreviations the proof becomes
›
(*<*)
lemma "¬ surj(f :: 'a  'a set)"
(*>*)
proof
  assume "surj f"
  hence "a. {x. x  f x} = f a" by(auto simp: surj_def)
  thus "False" by blast
qed
text‹

There are two further linguistic variations:
\medskip

\begin{tabular}{r@ {\quad=\quad}l}
(\isacom{have}$\mid$\isacom{show}) \ \textit{prop} \ \indexed{\isacom{using}}{using} \ \textit{facts}
&
\isacom{from} \ \textit{facts} \ (\isacom{have}$\mid$\isacom{show}) \ \textit{prop}\\
\indexed{\isacom{with}}{with} \ \textit{facts} & \isacom{from} \ \textit{facts} \isa{this}
\end{tabular}
\medskip

\noindent The \isacom{using} idiom de-emphasizes the used facts by moving them
behind the proposition.

\subsection{Structured Lemma Statements: \indexed{\isacom{fixes}}{fixes}, \indexed{\isacom{assumes}}{assumes}, \indexed{\isacom{shows}}{shows}}
\index{lemma@\isacom{lemma}}
Lemmas can also be stated in a more structured fashion. To demonstrate this
feature with Cantor's theorem, we rephrase \noquotes{@{prop[source]"¬ surj f"}}
a little:
›

lemma
  fixes f :: "'a  'a set"
  assumes s: "surj f"
  shows "False"

txt‹The optional \isacom{fixes} part allows you to state the types of
variables up front rather than by decorating one of their occurrences in the
formula with a type constraint. The key advantage of the structured format is
the \isacom{assumes} part that allows you to name each assumption; multiple
assumptions can be separated by \isacom{and}. The
\isacom{shows} part gives the goal. The actual theorem that will come out of
the proof is \noquotes{@{prop[source]"surj f  False"}}, but during the proof the assumption
\noquotes{@{prop[source]"surj f"}} is available under the name s› like any other fact.
›

proof -
  have " a. {x. x  f x} = f a" using s
    by(auto simp: surj_def)
  thus "False" by blast
qed

text‹
\begin{warn}
Note the hyphen after the \isacom{proof} command.
It is the null method that does nothing to the goal. Leaving it out would be asking
Isabelle to try some suitable introduction rule on the goal constFalse --- but
there is no such rule and \isacom{proof} would fail.
\end{warn}
In the \isacom{have} step the assumption \noquotes{@{prop[source]"surj f"}} is now
referenced by its name s›. The duplication of \noquotes{@{prop[source]"surj f"}} in the
above proofs (once in the statement of the lemma, once in its proof) has been
eliminated.

Stating a lemma with \isacom{assumes}-\isacom{shows} implicitly introduces the
name \indexed{assms›}{assms} that stands for the list of all assumptions. You can refer
to individual assumptions by assms(1)›, assms(2)›, etc.,
thus obviating the need to name them individually.

\section{Proof Patterns}

We show a number of important basic proof patterns. Many of them arise from
the rules of natural deduction that are applied by \isacom{proof} by
default. The patterns are phrased in terms of \isacom{show} but work for
\isacom{have} and \isacom{lemma}, too.

\ifsem\else
\subsection{Logic}
\fi

We start with two forms of \concept{case analysis}:
starting from a formula P› we have the two cases P› and
prop~P, and starting from a fact propP  Q
we have the two cases P› and Q›:
›text_raw‹
\begin{tabular}{@ {}ll@ {}}
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "R" proof-(*>*)
show "R"
proof cases
  assume "P"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "R" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
next
  assume "¬ P"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "R" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)oops(*>*)
text_raw ‹}
\end{minipage}\index{cases@cases›}
&
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "R" proof-(*>*)
have "P  Q" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
then show "R"
proof
  assume "P"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "R" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
next
  assume "Q"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "R" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)oops(*>*)

text_raw ‹}
\end{minipage}
\end{tabular}
\medskip
\begin{isamarkuptext}%
How to prove a logical equivalence:
\end{isamarkuptext}%
\isa{%
›
(*<*)lemma "PQ" proof-(*>*)
show "P  Q"
proof
  assume "P"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "Q" (*<*)sorry(*>*) text_raw‹\ \isasymproof\\›
next
  assume "Q"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "P" (*<*)sorry(*>*) text_raw‹\ \isasymproof\\›
qed(*<*)qed(*>*)
text_raw ‹}
\medskip
\begin{isamarkuptext}%
Proofs by contradiction (@{thm[source] ccontr} stands for ``classical contradiction''):
\end{isamarkuptext}%
\begin{tabular}{@ {}ll@ {}}
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "¬ P" proof-(*>*)
show "¬ P"
proof
  assume "P"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "False" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)oops(*>*)

text_raw ‹}
\end{minipage}
&
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "P" proof-(*>*)
show "P"
proof (rule ccontr)
  assume "¬P"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "False" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)oops(*>*)

text_raw ‹}
\end{minipage}
\end{tabular}
\medskip
\begin{isamarkuptext}%
How to prove quantified formulas:
\end{isamarkuptext}%
\begin{tabular}{@ {}ll@ {}}
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "x. P x" proof-(*>*)
show "x. P(x)"
proof
  fix x
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "P(x)" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)oops(*>*)

text_raw ‹}
\end{minipage}
&
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "x. P(x)" proof-(*>*)
show "x. P(x)"
proof
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "P(witness)" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed
(*<*)oops(*>*)

text_raw ‹}
\end{minipage}
\end{tabular}
\medskip
\begin{isamarkuptext}%
In the proof of \noquotes{@{prop[source]"x. P(x)"}},
the step \indexed{\isacom{fix}}{fix}~x› introduces a locally fixed variable x›
into the subproof, the proverbial ``arbitrary but fixed value''.
Instead of x› we could have chosen any name in the subproof.
In the proof of \noquotes{@{prop[source]"x. P(x)"}},
witness› is some arbitrary
term for which we can prove that it satisfies P›.

How to reason forward from \noquotes{@{prop[source] "x. P(x)"}}:
\end{isamarkuptext}%
›
(*<*)lemma True proof- assume 1: "x. P x"(*>*)
have "x. P(x)" (*<*)by(rule 1)(*>*)text_raw‹\ \isasymproof\\›
then obtain x where p: "P(x)" by blast
(*<*)oops(*>*)
text‹
After the \indexed{\isacom{obtain}}{obtain} step, x› (we could have chosen any name)
is a fixed local
variable, and p› is the name of the fact
\noquotes{@{prop[source] "P(x)"}}.
This pattern works for one or more x›.
As an example of the \isacom{obtain} command, here is the proof of
Cantor's theorem in more detail:
›

lemma "¬ surj(f :: 'a  'a set)"
proof
  assume "surj f"
  hence  "a. {x. x  f x} = f a" by(auto simp: surj_def)
  then obtain a where  "{x. x  f x} = f a"  by blast
  hence  "a  f a  a  f a"  by blast
  thus "False" by blast
qed

text_raw‹
\begin{isamarkuptext}%

Finally, how to prove set equality and subset relationship:
\end{isamarkuptext}%
\begin{tabular}{@ {}ll@ {}}
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "A = (B::'a set)" proof-(*>*)
show "A = B"
proof
  show "A  B" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
next
  show "B  A" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)qed(*>*)

text_raw ‹}
\end{minipage}
&
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "A <= (B::'a set)" proof-(*>*)
show "A  B"
proof
  fix x
  assume "x  A"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "x  B" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)qed(*>*)

text_raw ‹}
\end{minipage}
\end{tabular}
\begin{isamarkuptext}%

\ifsem\else
\subsection{Chains of (In)Equations}

In textbooks, chains of equations (and inequations) are often displayed like this:
\begin{quote}
\begin{tabular}{@ {}l@ {\qquad}l@ {}}
$t_1 = t_2$ & \isamath{\,\langle\mathit{justification}\rangle}\\
$\phantom{t_1} = t_3$ & \isamath{\,\langle\mathit{justification}\rangle}\\
\quad $\vdots$\\
$\phantom{t_1} = t_n$ & \isamath{\,\langle\mathit{justification}\rangle}
\end{tabular}
\end{quote}
The Isar equivalent is this:

\begin{samepage}
\begin{quote}
\isacom{have} "t1 = t2"› \isasymproof\\
\isacom{also have} "... = t3"› \isasymproof\\
\quad $\vdots$\\
\isacom{also have} "... = tn"› \isasymproof \\
\isacom{finally show} "t1 = tn"›\ \texttt{.}
\end{quote}
\end{samepage}

\noindent
The ``...›'' and ``.›'' deserve some explanation:
\begin{description}
\item[``...›''] is literally three dots. It is the name of an unknown that Isar
automatically instantiates with the right-hand side of the previous equation.
In general, if this› is the theorem termp t1 t2 then ``...›''
stands for t2.
\item[``.›''] (a single dot) is a proof method that solves a goal by one of the
assumptions. This works here because the result of \isacom{finally}
is the theorem \mbox{t1 = tn},
\isacom{show} "t1 = tn"› states the theorem explicitly,
and ``.›'' proves the theorem with the result of \isacom{finally}.
\end{description}
The above proof template also works for arbitrary mixtures of =›, ≤› and <›,
for example:
\begin{quote}
\isacom{have} "t1 < t2"› \isasymproof\\
\isacom{also have} "... = t3"› \isasymproof\\
\quad $\vdots$\\
\isacom{also have} "... ≤ tn"› \isasymproof \\
\isacom{finally show} "t1 < tn"›\ \texttt{.}
\end{quote}
The relation symbol in the \isacom{finally} step needs to be the most precise one
possible. In the example above, you must not write t1 ≤ tn instead of \mbox{t1 < tn}.

\begin{warn}
Isabelle only supports =›, ≤› and <› but not ≥› and >›
in (in)equation chains (by default).
\end{warn}

If you want to go beyond merely using the above proof patterns and want to
understand what \isacom{also} and \isacom{finally} mean, read on.
There is an Isar theorem variable called calculation›, similar to this›.
When the first \isacom{also} in a chain is encountered, Isabelle sets
calculation := this›. In each subsequent \isacom{also} step,
Isabelle composes the theorems calculation› and this› (i.e.\ the two previous
(in)equalities) using some predefined set of rules including transitivity
of =›, ≤› and <› but also mixed rules like prop x  y; y < z   x < z.
The result of this composition is assigned to calculation›. Consider
\begin{quote}
\isacom{have} "t1 ≤ t2"› \isasymproof\\
\isacom{also} \isacom{have} "... < t3"› \isasymproof\\
\isacom{also} \isacom{have} "... = t4"› \isasymproof\\
\isacom{finally show} "t1 < t4"›\ \texttt{.}
\end{quote}
After the first \isacom{also}, calculation› is "t1 ≤ t2"›,
and after the second \isacom{also}, calculation› is "t1 < t3"›.
The command \isacom{finally} is short for \isacom{also from} calculation›.
Therefore the \isacom{also} hidden in \isacom{finally} sets calculation›
to t1 < t4 and the final ``\texttt{.}'' succeeds.

For more information on this style of proof see cite"BauerW-TPHOLs01".
\fi

\section{Streamlining Proofs}

\subsection{Pattern Matching and Quotations}

In the proof patterns shown above, formulas are often duplicated.
This can make the text harder to read, write and maintain. Pattern matching
is an abbreviation mechanism to avoid such duplication. Writing
\begin{quote}
\isacom{show} \ \textit{formula} (›\indexed{\isacom{is}}{is} \textit{pattern})›
\end{quote}
matches the pattern against the formula, thus instantiating the unknowns in
the pattern for later use. As an example, consider the proof pattern for
⟷›:
\end{isamarkuptext}%
›
(*<*)lemma "formula1  formula2" proof-(*>*)
show "formula1  formula2" (is "?L  ?R")
proof
  assume "?L"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "?R" (*<*)sorry(*>*) text_raw‹\ \isasymproof\\›
next
  assume "?R"
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show "?L" (*<*)sorry(*>*) text_raw‹\ \isasymproof\\›
qed(*<*)qed(*>*)

text‹Instead of duplicating formulai in the text, we introduce
the two abbreviations ?L› and ?R› by pattern matching.
Pattern matching works wherever a formula is stated, in particular
with \isacom{have} and \isacom{lemma}.

The unknown \indexed{?thesis›}{thesis} is implicitly matched against any goal stated by
\isacom{lemma} or \isacom{show}. Here is a typical example:›

lemma "formula"
proof -
  text_raw‹\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}›
  show ?thesis (*<*)sorry(*>*) text_raw‹\ \isasymproof\\›
qed

text‹
Unknowns can also be instantiated with \indexed{\isacom{let}}{let} commands
\begin{quote}
\isacom{let} ?t› = "›\textit{some-big-term}"›
\end{quote}
Later proof steps can refer to ?t›:
\begin{quote}
\isacom{have} "›\dots ?t› \dots"›
\end{quote}
\begin{warn}
Names of facts are introduced with name:› and refer to proved
theorems. Unknowns ?X› refer to terms or formulas.
\end{warn}

Although abbreviations shorten the text, the reader needs to remember what
they stand for. Similarly for names of facts. Names like 1›, 2›
and 3› are not helpful and should only be used in short proofs. For
longer proofs, descriptive names are better. But look at this example:
\begin{quote}
\isacom{have} \ x_gr_0: "x > 0"›\\
$\vdots$\\
\isacom{from} x_gr_0› \dots
\end{quote}
The name is longer than the fact it stands for! Short facts do not need names;
one can refer to them easily by quoting them:
\begin{quote}
\isacom{have} \ "x > 0"›\\
$\vdots$\\
\isacom{from} ‹x > 0›› \dots\index{$IMP053@`...`›}
\end{quote}
The outside quotes in ‹x > 0›› are the standard renderings of the symbols \texttt{\textbackslash<open>} and \texttt{\textbackslash<close>}.
They refer to the fact not by name but ``by value''.

\subsection{\indexed{\isacom{moreover}}{moreover}}
\index{ultimately@\isacom{ultimately}}

Sometimes one needs a number of facts to enable some deduction. Of course
one can name these facts individually, as shown on the right,
but one can also combine them with \isacom{moreover}, as shown on the left:
›text_raw‹
\begin{tabular}{@ {}ll@ {}}
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "P" proof-(*>*)
have "P1" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
moreover have "P2" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
moreover
text_raw‹\\$\vdots$\\\hspace{-1.4ex}›(*<*)have "True" ..(*>*)
moreover have "Pn" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
ultimately have "P"  (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
(*<*)oops(*>*)

text_raw ‹}
\end{minipage}
&
\qquad
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "P" proof-(*>*)
have lab1: "P1" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
have lab2: "P2" (*<*)sorry(*>*)text_raw‹\ \isasymproof›
text_raw‹\\$\vdots$\\\hspace{-1.4ex}›
have labn: "Pn" (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
from lab1 lab2 text_raw‹\ $\dots$\\›
have "P"  (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
(*<*)oops(*>*)

text_raw ‹}
\end{minipage}
\end{tabular}
\begin{isamarkuptext}%
The \isacom{moreover} version is no shorter but expresses the structure
a bit more clearly and avoids new names.

\subsection{Local Lemmas}

Sometimes one would like to prove some lemma locally within a proof,
a lemma that shares the current context of assumptions but that
has its own assumptions and is generalized over its locally fixed
variables at the end. This is simply an extension of the basic
\indexed{\isacom{have}}{have} construct:
\begin{quote}
\indexed{\isacom{have}}{have} B›\
 \indexed{\isacom{if}}{if} name:› A1 … Am\
 \indexed{\isacom{for}}{for} x1 … xn\\
\isasymproof
\end{quote}
proves ⟦ A1; … ; Am ⟧ ⟹ B›
where all xi have been replaced by unknowns ?xi.
As an example we prove a simple fact about divisibility on integers.
The definition of dvd› is @{thm dvd_def}.
\end{isamarkuptext}%
›

lemma fixes a b :: int assumes "b dvd (a+b)" shows "b dvd a"
proof -
  have "k'. a = b*k'" if asm: "a+b = b*k" for k
  proof
    show "a = b*(k - 1)" using asm by(simp add: algebra_simps)
  qed
  then show ?thesis using assms by(auto simp add: dvd_def)
qed

text‹

\subsection*{Exercises}

\exercise
Give a readable, structured proof of the following lemma:
›
lemma assumes T: "x y. T x y  T y x"
  and A: "x y. A x y  A y x  x = y"
  and TA: "x y. T x y  A x y" and "A x y"
  shows "T x y"
(*<*)oops(*>*)
text‹
\endexercise

\exercise
Give a readable, structured proof of the following lemma:
›
lemma "ys zs. xs = ys @ zs 
            (length ys = length zs  length ys = length zs + 1)"
(*<*)oops(*>*)
text‹
Hint: There are predefined functions @{const_typ take} and @{const_typ drop}
such that take k [x1,…] = [x1,…,xk]› and
drop k [x1,…] = [xk+1,…]›. Let sledgehammer find and apply
the relevant consttake and constdrop lemmas for you.
\endexercise


\section{Case Analysis and Induction}

\subsection{Datatype Case Analysis}
\index{case analysis|(}

We have seen case analysis on formulas. Now we want to distinguish
which form some term takes: is it 0› or of the form termSuc n,
is it term[] or of the form termx#xs, etc. Here is a typical example
proof by case analysis on the form of xs›:
›

lemma "length(tl xs) = length xs - 1"
proof (cases xs)
  assume "xs = []"
  thus ?thesis by simp
next
  fix y ys assume "xs = y#ys"
  thus ?thesis by simp
qed

text‹\index{cases@cases›|(}Function tl› (''tail'') is defined by @{thm list.sel(2)} and
@{thm list.sel(3)}. Note that the result type of constlength is typnat
and prop0 - 1 = (0::nat).

This proof pattern works for any term t› whose type is a datatype.
The goal has to be proved for each constructor C›:
\begin{quote}
\isacom{fix} \ x1 … xn \isacom{assume} "t = C x1 … xn"›
\end{quote}\index{case@\isacom{case}|(}
Each case can be written in a more compact form by means of the \isacom{case}
command:
\begin{quote}
\isacom{case} (C x1 … xn)›
\end{quote}
This is equivalent to the explicit \isacom{fix}-\isacom{assume} line
but also gives the assumption "t = C x1 … xn"› a name: C›,
like the constructor.
Here is the \isacom{case} version of the proof above:
›
(*<*)lemma "length(tl xs) = length xs - 1"(*>*)
proof (cases xs)
  case Nil
  thus ?thesis by simp
next
  case (Cons y ys)
  thus ?thesis by simp
qed

text‹Remember that Nil› and Cons› are the alphanumeric names
for []› and #›. The names of the assumptions
are not used because they are directly piped (via \isacom{thus})
into the proof of the claim.
\index{case analysis|)}

\subsection{Structural Induction}
\index{induction|(}
\index{structural induction|(}

We illustrate structural induction with an example based on natural numbers:
the sum (∑›) of the first n› natural numbers
({0..n::nat}›) is equal to \mbox{termn*(n+1) div 2::nat}.
Never mind the details, just focus on the pattern:
›

lemma "{0..n::nat} = n*(n+1) div 2"
proof (induction n)
  show "{0..0::nat} = 0*(0+1) div 2" by simp
next
  fix n assume "{0..n::nat} = n*(n+1) div 2"
  thus "{0..Suc n} = Suc n*(Suc n+1) div 2" by simp
qed

text‹Except for the rewrite steps, everything is explicitly given. This
makes the proof easily readable, but the duplication means it is tedious to
write and maintain. Here is how pattern
matching can completely avoid any duplication:›

lemma "{0..n::nat} = n*(n+1) div 2" (is "?P n")
proof (induction n)
  show "?P 0" by simp
next
  fix n assume "?P n"
  thus "?P(Suc n)" by simp
qed

text‹The first line introduces an abbreviation ?P n› for the goal.
Pattern matching ?P n› with the goal instantiates ?P› to the
function termλn. {0..n::nat} = n*(n+1) div 2.  Now the proposition to
be proved in the base case can be written as ?P 0›, the induction
hypothesis as ?P n›, and the conclusion of the induction step as
?P(Suc n)›.

Induction also provides the \isacom{case} idiom that abbreviates
the \isacom{fix}-\isacom{assume} step. The above proof becomes
›
(*<*)lemma "{0..n::nat} = n*(n+1) div 2"(*>*)
proof (induction n)
  case 0
  show ?case by simp
next
  case (Suc n)
  thus ?case by simp
qed

text‹
The unknown ?case›\index{case?@?case›|(} is set in each case to the required
claim, i.e., ?P 0› and \mbox{?P(Suc n)›} in the above proof,
without requiring the user to define a ?P›. The general
pattern for induction over typnat is shown on the left-hand side:
›text_raw‹
\begin{tabular}{@ {}ll@ {}}
\begin{minipage}[t]{.4\textwidth}
\isa{%
›
(*<*)lemma "P(n::nat)" proof -(*>*)
show "P(n)"
proof (induction n)
  case 0
  text_raw‹\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}›
  show ?case (*<*)sorry(*>*) text_raw‹\ \isasymproof\\›
next
  case (Suc n)
  text_raw‹\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}›
  show ?case (*<*)sorry(*>*) text_raw‹\ \isasymproof\\›
qed(*<*)qed(*>*)

text_raw ‹}
\end{minipage}
&
\begin{minipage}[t]{.4\textwidth}
~\\
~\\
\isacom{let} ?case = "P(0)"›\\
~\\
~\\
~\\[1ex]
\isacom{fix} n› \isacom{assume} Suc: "P(n)"›\\
\isacom{let} ?case = "P(Suc n)"›\\
\end{minipage}
\end{tabular}
\medskip
›
text‹
On the right side you can see what the \isacom{case} command
on the left stands for.

In case the goal is an implication, induction does one more thing: the
proposition to be proved in each case is not the whole implication but only
its conclusion; the premises of the implication are immediately made
assumptions of that case. That is, if in the above proof we replace
\isacom{show}~"P(n)"› by
\mbox{\isacom{show}~"A(n) ⟹ P(n)"›}
then \isacom{case}~0› stands for
\begin{quote}
\isacom{assume} \ 0: "A(0)"›\\
\isacom{let} ?case = "P(0)"›
\end{quote}
and \isacom{case}~(Suc n)› stands for
\begin{quote}
\isacom{fix} n›\\
\isacom{assume} Suc:›
  \begin{tabular}[t]{l}"A(n) ⟹ P(n)"›\\"A(Suc n)"›\end{tabular}\\
\isacom{let} ?case = "P(Suc n)"›
\end{quote}
The list of assumptions Suc› is actually subdivided
into Suc.IH›, the induction hypotheses (here A(n) ⟹ P(n)›),
and Suc.prems›, the premises of the goal being proved
(here A(Suc n)›).

Induction works for any datatype.
Proving a goal ⟦ A1(x); …; Ak(x) ⟧ ⟹ P(x)›
by induction on x› generates a proof obligation for each constructor
C› of the datatype. The command \isacom{case}~(C x1 … xn)›
performs the following steps:
\begin{enumerate}
\item \isacom{fix} x1 … xn
\item \isacom{assume} the induction hypotheses (calling them C.IH›\index{IH@.IH›})
 and the premises \mbox{Ai(C x1 … xn)›} (calling them C.prems›\index{prems@.prems›})
 and calling the whole list C›
\item \isacom{let} ?case = "P(C x1 … xn)"›
\end{enumerate}
\index{structural induction|)}


\ifsem\else
\subsection{Computation Induction}
\index{rule induction}

In \autoref{sec:recursive-funs} we introduced computation induction and
its realization in Isabelle: the definition
of a recursive function f› via \isacom{fun} proves the corresponding computation
induction rule called f.induct›. Induction with this rule looks like in
\autoref{sec:recursive-funs}, but now with \isacom{proof} instead of \isacom{apply}:
\begin{quote}
\isacom{proof} (induction x1 … xk rule: f.induct›)
\end{quote}
Just as for structural induction, this creates several cases, one for each
defining equation for f›. By default (if the equations have not been named
by the user), the cases are numbered. That is, they are started by
\begin{quote}
\isacom{case} (i x y ...›)
\end{quote}
where i = 1,...,n›, n› is the number of equations defining f›,
and x y ...› are the variables in equation i›. Note the following:
\begin{itemize}
\item
Although i› is an Isar name, i.IH› (or similar) is not. You need
double quotes: "i.IH›". When indexing the name, write "i.IH›"(1),
not "i.IH›(1)".
\item
If defining equations for f› overlap, \isacom{fun} instantiates them to make
them nonoverlapping. This means that one user-provided equation may lead to
several equations and thus to several cases in the induction rule.
These have names of the form "i_j›", where i› is the number of the original
equation and the system-generated j› indicates the subcase.
\end{itemize}
In Isabelle/jEdit, the induction› proof method displays a proof skeleton
with all \isacom{case}s. This is particularly useful for computation induction
and the following rule induction.
\fi


\subsection{Rule Induction}
\index{rule induction|(}

Recall the inductive and recursive definitions of even numbers in
\autoref{sec:inductive-defs}:
›

inductive ev :: "nat  bool" where
ev0: "ev 0" |
evSS: "ev n  ev(Suc(Suc n))"

fun evn :: "nat  bool" where
"evn 0 = True" |
"evn (Suc 0) = False" |
"evn (Suc(Suc n)) = evn n"

text‹We recast the proof of propev n  evn n in Isar. The
left column shows the actual proof text, the right column shows
the implicit effect of the two \isacom{case} commands:›text_raw‹
\begin{tabular}{@ {}l@ {\qquad}l@ {}}
\begin{minipage}[t]{.5\textwidth}
\isa{%
›

lemma "ev n  evn n"
proof(induction rule: ev.induct)
  case ev0
  show ?case by simp
next
  case evSS



  thus ?case by simp
qed

text_raw ‹}
\end{minipage}
&
\begin{minipage}[t]{.5\textwidth}
~\\
~\\
\isacom{let} ?case = "evn 0"›\\
~\\
~\\
\isacom{fix} n›\\
\isacom{assume} evSS:›
  \begin{tabular}[t]{l} "ev n"›\\"evn n"›\end{tabular}\\
\isacom{let} ?case = "evn(Suc(Suc n))"›\\
\end{minipage}
\end{tabular}
\medskip
›
text‹
The proof resembles structural induction, but the induction rule is given
explicitly and the names of the cases are the names of the rules in the
inductive definition.
Let us examine the two assumptions named @{thm[source]evSS}:
propev n is the premise of rule @{thm[source]evSS}, which we may assume
because we are in the case where that rule was used; propevn n
is the induction hypothesis.
\begin{warn}
Because each \isacom{case} command introduces a list of assumptions
named like the case name, which is the name of a rule of the inductive
definition, those rules now need to be accessed with a qualified name, here
@{thm[source] ev.ev0} and @{thm[source] ev.evSS}.
\end{warn}

In the case @{thm[source]evSS} of the proof above we have pretended that the
system fixes a variable n›.  But unless the user provides the name
n›, the system will just invent its own name that cannot be referred
to.  In the above proof, we do not need to refer to it, hence we do not give
it a specific name. In case one needs to refer to it one writes
\begin{quote}
\isacom{case} (evSS m)›
\end{quote}
like \isacom{case}~(Suc n)› in earlier structural inductions.
The name m› is an arbitrary choice. As a result,
case @{thm[source] evSS} is derived from a renamed version of
rule @{thm[source] evSS}: ev m ⟹ ev(Suc(Suc m))›.
Here is an example with a (contrived) intermediate step that refers to m›:
›

lemma "ev n  evn n"
proof(induction rule: ev.induct)
  case ev0 show ?case by simp
next
  case (evSS m)
  have "evn(Suc(Suc m)) = evn m" by simp
  thus ?case using `evn m` by blast
qed

text‹
\indent
In general, let I› be a (for simplicity unary) inductively defined
predicate and let the rules in the definition of I›
be called rule1, \dots, rulen. A proof by rule
induction follows this pattern:\index{inductionrule@induction ... rule:›}
›

(*<*)
inductive I where rule1: "I()" |  rule2: "I()" |  rulen: "I()"
lemma "I x  P x" proof-(*>*)
show "I x  P x"
proof(induction rule: I.induct)
  case rule1
  text_raw‹\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}›
  show ?case (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
next
  text_raw‹\\[-.4ex]$\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}›
(*<*)
  case rule2
  show ?case sorry
(*>*)
next
  case rulen
  text_raw‹\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}›
  show ?case (*<*)sorry(*>*)text_raw‹\ \isasymproof\\›
qed(*<*)qed(*>*)

text‹
One can provide explicit variable names by writing
\isacom{case}~(rulei x1 … xk)›, thus renaming the first k›
free variables in rule i› to x1 … xk,
going through rule i› from left to right.

\subsection{Assumption Naming}
\label{sec:assm-naming}

In any induction, \isacom{case}~name› sets up a list of assumptions
also called name›, which is subdivided into three parts:
\begin{description}
\item[name.IH›]\index{IH@.IH›} contains the induction hypotheses.
\item[name.hyps›]\index{hyps@.hyps›} contains all the other hypotheses of this case in the
induction rule. For rule inductions these are the hypotheses of rule
name›, for structural inductions these are empty.
\item[name.prems›]\index{prems@.prems›} contains the (suitably instantiated) premises
of the statement being proved, i.e., the Ai when
proving ⟦ A1; …; An ⟧ ⟹ A›.
\end{description}
\begin{warn}
Proof method induct› differs from induction›
only in this naming policy: induct› does not distinguish
IH› from hyps› but subsumes IH› under hyps›.
\end{warn}

More complicated inductive proofs than the ones we have seen so far
often need to refer to specific assumptions --- just name› or even
name.prems› and name.IH› can be too unspecific.
This is where the indexing of fact lists comes in handy, e.g.,
name.IH(2)› or name.prems(1-2)›.

\subsection{Rule Inversion}
\label{sec:rule-inversion}
\index{rule inversion|(}

Rule inversion is case analysis of which rule could have been used to
derive some fact. The name \conceptnoidx{rule inversion} emphasizes that we are
reasoning backwards: by which rules could some given fact have been proved?
For the inductive definition of constev, rule inversion can be summarized
like this:
@{prop[display]"ev n  n = 0  (k. n = Suc(Suc k)  ev k)"}
The realisation in Isabelle is a case analysis.
A simple example is the proof that propev n  ev (n - 2). We
already went through the details informally in \autoref{sec:Logic:even}. This
is the Isar proof:
›
(*<*)
notepad
begin fix n
(*>*)
  assume "ev n"
  from this have "ev(n - 2)"
  proof cases
    case ev0 thus "ev(n - 2)" by (simp add: ev.ev0)
  next
    case (evSS k) thus "ev(n - 2)" by (simp add: ev.evSS)
  qed
(*<*)
end
(*>*)

text‹The key point here is that a case analysis over some inductively
defined predicate is triggered by piping the given fact
(here: \isacom{from}~this›) into a proof by cases›.
Let us examine the assumptions available in each case. In case ev0›
we have n = 0› and in case evSS› we have propn = Suc(Suc k)
and propev k. In each case the assumptions are available under the name
of the case; there is no fine-grained naming schema like there is for induction.

Sometimes some rules could not have been used to derive the given fact
because constructors clash. As an extreme example consider
rule inversion applied to propev(Suc 0): neither rule ev0› nor
rule evSS› can yield propev(Suc 0) because Suc 0› unifies
neither with 0› nor with termSuc(Suc n). Impossible cases do not
have to be proved. Hence we can prove anything from propev(Suc 0):
›
(*<*)
notepad begin fix P
(*>*)
  assume "ev(Suc 0)" then have P by cases
(*<*)
end
(*>*)

text‹That is, propev(Suc 0) is simply not provable:›

lemma "¬ ev(Suc 0)"
proof
  assume "ev(Suc 0)" then show False by cases
qed

text‹Normally not all cases will be impossible. As a simple exercise,
prove that \mbox{prop¬ ev(Suc(Suc(Suc 0))).}

\subsection{Advanced Rule Induction}
\label{sec:advanced-rule-induction}

So far, rule induction was always applied to goals of the form I x y z ⟹ …›
where I› is some inductively defined predicate and x›, y›, z›
are variables. In some rare situations one needs to deal with an assumption where
not all arguments r›, s›, t› are variables:
\begin{isabelle}
\isacom{lemma} "I r s t ⟹ …"›
\end{isabelle}
Applying the standard form of
rule induction in such a situation will lead to strange and typically unprovable goals.
We can easily reduce this situation to the standard one by introducing
new variables x›, y›, z› and reformulating the goal like this:
\begin{isabelle}
\isacom{lemma} "I x y z ⟹ x = r ⟹ y = s ⟹ z = t ⟹ …"›
\end{isabelle}
Standard rule induction will work fine now, provided the free variables in
r›, s›, t› are generalized via arbitrary›.

However, induction can do the above transformation for us, behind the curtains, so we never
need to see the expanded version of the lemma. This is what we need to write:
\begin{isabelle}
\isacom{lemma} "I r s t ⟹ …"›\isanewline
\isacom{proof}(induction "r" "s" "t" arbitrary: … rule: I.induct)›\index{inductionrule@induction ... rule:›}\index{arbitrary@arbitrary:›}
\end{isabelle}
Like for rule inversion, cases that are impossible because of constructor clashes
will not show up at all. Here is a concrete example:›

lemma "ev (Suc m)  ¬ ev m"
proof(induction "Suc m" arbitrary: m rule: ev.induct)
  fix n assume IH: "m. n = Suc m  ¬ ev m"
  show "¬ ev (Suc n)"
  proof ― ‹contradiction›
    assume "ev(Suc n)"
    thus False
    proof cases ― ‹rule inversion›
      fix k assume "n = Suc k" "ev k"
      thus False using IH by auto
    qed
  qed
qed

text‹
Remarks:
\begin{itemize}
\item 
Instead of the \isacom{case} and ?case› magic we have spelled all formulas out.
This is merely for greater clarity.
\item
We only need to deal with one case because the @{thm[source] ev0} case is impossible.
\item
The form of the IH› shows us that internally the lemma was expanded as explained
above: \noquotes{@{prop[source]"ev x  x = Suc m  ¬ ev m"}}.
\item
The goal prop¬ ev (Suc n) may surprise. The expanded version of the lemma
would suggest that we have a \isacom{fix} m› \isacom{assume} propSuc(Suc n) = Suc m
and need to show prop¬ ev m. What happened is that Isabelle immediately
simplified propSuc(Suc n) = Suc m to propSuc n = m and could then eliminate
m›. Beware of such nice surprises with this advanced form of induction.
\end{itemize}
\begin{warn}
This advanced form of induction does not support the IH›
naming schema explained in \autoref{sec:assm-naming}:
the induction hypotheses are instead found under the name hyps›,
as they are for the simpler
induct› method.
\end{warn}
\index{induction|)}
\index{cases@cases›|)}
\index{case@\isacom{case}|)}
\index{case?@?case›|)}
\index{rule induction|)}
\index{rule inversion|)}

\subsection*{Exercises}


\exercise
Give a structured proof by rule inversion:
›

lemma assumes a: "ev(Suc(Suc n))" shows "ev n"
(*<*)oops(*>*)

text‹
\endexercise

\begin{exercise}
Give a structured proof of prop¬ ev(Suc(Suc(Suc 0)))
by rule inversions. If there are no cases to be proved you can close
a proof immediately with \isacom{qed}.
\end{exercise}

\begin{exercise}
Recall predicate star› from \autoref{sec:star} and iter›
from Exercise~\ref{exe:iter}. Prove propiter r n x y  star r x y
in a structured style; do not just sledgehammer each case of the
required induction.
\end{exercise}

\begin{exercise}
Define a recursive function elems ::› typ'a list  'a set
and prove propx  elems xs  ys zs. xs = ys @ x # zs  x  elems ys.
\end{exercise}

\begin{exercise}
Extend Exercise~\ref{exe:cfg} with a function that checks if some
\mbox{alpha list›} is a balanced
string of parentheses. More precisely, define a \mbox{recursive} function
balanced :: nat ⇒ alpha list ⇒ bool› such that termbalanced n w
is true iff (informally) S (an @ w)›. Formally, prove that
propbalanced n w  S (replicate n a @ w) where
constreplicate ::› typnat  'a  'a list is predefined
and termreplicate n x yields the list [x, …, x]› of length n›.
\end{exercise}
›

(*<*)
end
(*>*)