Theory Chapter1_MinMax
text_raw ‹\part{Documentation}›
chapter ‹Quickstart \label{chap:Quickstart}›
text ‹\chapterauthor{Japheth Lim, Rohan Jacob-Rao, David Greenaway \lowercase{and} Norbert Schirmer}›
theory Chapter1_MinMax
imports "AutoCorres2.AutoCorres"
begin
section ‹Introduction›
text ‹
AutoCorres is a tool that attempts to simplify the formal verification of C
programs in the Isabelle/HOL theorem prover. It allows C code
to be automatically abstracted to produce a higher-level functional
specification.
AutoCorres relies on the C-Parser~\cite{CParser-download} developed by Michael Norrish
at NICTA. This tool takes raw C code as input and produces a translation in
SIMPL~\cite{Simpl-AFP}, an imperative language written by Norbert Schirmer on top
of Isabelle. AutoCorres takes this SIMPL code to produce a "monadic"
specification, which is intended to be simpler to reason about in Isabelle.
The composition of these two tools (AutoCorres applied after the C-Parser) can
then be used to reason about C programs.
This guide is written for users of Isabelle/HOL, with some knowledge of C, to
get started proving properties of C programs. Using AutoCorres in conjunction
with the verification condition generator (VCG) @{method runs_to_vcg}, one
should be able to do this without an understanding of SIMPL nor all the details of the monadic
representation produced by AutoCorres. We will see how this is possible in the
next chapter.
›
section ‹A First Proof with AutoCorres›
text ‹
We will now show how to use these tools to prove correctness of some very
simple C functions.
›
subsection ‹Two simple functions: \texttt{min} and \texttt{max}›
text ‹
Consider the following two functions, defined in a file \texttt{minmax.c},
which (we expect) return the minimum and maximum respectively of two unsigned
integers.
\lstinputlisting[language=C, firstline=13]{minmax.c}
It is easy to see that \texttt{min} is correct, but perhaps less obvious why
\texttt{max} is correct. AutoCorres will hopefully allow us to prove these
claims without too much effort.
›
subsection ‹Invoking the C-parser›
text ‹
As mentioned earlier, AutoCorres does not handle C code directly. The first
step is to apply the
C-Parser by using @{command install_C_file} to
obtain a SIMPL translation. We do this using the \texttt{install-C-file}
command in Isabelle, as shown.
›