author | wenzelm |
Tue, 24 Sep 2024 21:31:20 +0200 | |
changeset 80946 | b76f64d7d493 |
parent 76987 | 4c275405faae |
permissions | -rw-r--r-- |
45044 | 1 |
(*<*) |
2 |
theory Example_Verification |
|
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
64246
diff
changeset
|
3 |
imports "HOL-SPARK-Examples.Greatest_Common_Divisor" Simple_Greatest_Common_Divisor |
45044 | 4 |
begin |
5 |
(*>*) |
|
6 |
||
63167 | 7 |
chapter \<open>Verifying an Example Program\<close> |
45044 | 8 |
|
63167 | 9 |
text \<open> |
45044 | 10 |
\label{sec:example-verification} |
11 |
\begin{figure} |
|
12 |
\lstinputlisting{Gcd.ads} |
|
13 |
\lstinputlisting{Gcd.adb} |
|
14 |
\caption{\SPARK{} program for computing the greatest common divisor} |
|
15 |
\label{fig:gcd-prog} |
|
16 |
\end{figure} |
|
17 |
||
18 |
\begin{figure} |
|
19 |
\input{Greatest_Common_Divisor} |
|
20 |
\caption{Correctness proof for the greatest common divisor program} |
|
21 |
\label{fig:gcd-proof} |
|
22 |
\end{figure} |
|
23 |
We will now explain the usage of the \SPARK{} verification environment by proving |
|
24 |
the correctness of an example program. As an example, we use a program for computing |
|
25 |
the \emph{greatest common divisor} of two natural numbers shown in \figref{fig:gcd-prog}, |
|
76987 | 26 |
which has been taken from the book about \SPARK{} by Barnes \<^cite>\<open>\<open>\S 11.6\<close> in Barnes\<close>. |
63167 | 27 |
\<close> |
45044 | 28 |
|
63167 | 29 |
section \<open>Importing \SPARK{} VCs into Isabelle\<close> |
45044 | 30 |
|
63167 | 31 |
text \<open> |
45044 | 32 |
In order to specify that the \SPARK{} procedure \texttt{G\_C\_D} behaves like its |
33 |
mathematical counterpart, Barnes introduces a \emph{proof function} \texttt{Gcd} |
|
34 |
in the package specification. Invoking the \SPARK{} Examiner and Simplifier on |
|
35 |
this program yields a file \texttt{g\_c\_d.siv} containing the simplified VCs, |
|
36 |
as well as files \texttt{g\_c\_d.fdl} and \texttt{g\_c\_d.rls}, containing FDL |
|
37 |
declarations and rules, respectively. The files generated by \SPARK{} are assumed to reside in the |
|
38 |
subdirectory \texttt{greatest\_common\_divisor}. For \texttt{G\_C\_D} the |
|
39 |
Examiner generates ten VCs, eight of which are proved automatically by |
|
40 |
the Simplifier. We now show how to prove the remaining two VCs |
|
41 |
interactively using HOL-\SPARK{}. For this purpose, we create a \emph{theory} |
|
42 |
\texttt{Greatest\_Common\_Divisor}, which is shown in \figref{fig:gcd-proof}. |
|
43 |
A theory file always starts with the keyword \isa{\isacommand{theory}} followed |
|
44 |
by the name of the theory, which must be the same as the file name. The theory |
|
45 |
name is followed by the keyword \isa{\isacommand{imports}} and a list of theories |
|
46 |
imported by the current theory. All theories using the HOL-\SPARK{} verification |
|
47 |
environment must import the theory \texttt{SPARK}. In addition, we also include |
|
48 |
the \texttt{GCD} theory. The list of imported theories is followed by the |
|
49 |
\isa{\isacommand{begin}} keyword. In order to interactively process the theory |
|
50 |
shown in \figref{fig:gcd-proof}, we start Isabelle with the command |
|
51 |
\begin{verbatim} |
|
68407 | 52 |
isabelle jedit -l HOL-SPARK Greatest_Common_Divisor.thy |
45044 | 53 |
\end{verbatim} |
54 |
The option ``\texttt{-l HOL-SPARK}'' instructs Isabelle to load the right |
|
55 |
object logic image containing the verification environment. Each proof function |
|
56 |
occurring in the specification of a \SPARK{} program must be linked with a |
|
57 |
corresponding Isabelle function. This is accomplished by the command |
|
58 |
\isa{\isacommand{spark\_proof\_functions}}, which expects a list of equations |
|
59 |
of the form \emph{name}\texttt{\ =\ }\emph{term}, where \emph{name} is the |
|
60 |
name of the proof function and \emph{term} is the corresponding Isabelle term. |
|
61 |
In the case of \texttt{gcd}, both the \SPARK{} proof function and its Isabelle |
|
62 |
counterpart happen to have the same name. Isabelle checks that the type of the |
|
63 |
term linked with a proof function agrees with the type of the function declared |
|
64 |
in the \texttt{*.fdl} file. |
|
65 |
It is worth noting that the |
|
66 |
\isa{\isacommand{spark\_proof\_functions}} command can be invoked both outside, |
|
67 |
i.e.\ before \isa{\isacommand{spark\_open}}, and inside the environment, i.e.\ after |
|
68 |
\isa{\isacommand{spark\_open}}, but before any \isa{\isacommand{spark\_vc}} command. The |
|
69 |
former variant is useful when having to declare proof functions that are shared by several |
|
70 |
procedures, whereas the latter has the advantage that the type of the proof function |
|
71 |
can be checked immediately, since the VCs, and hence also the declarations of proof |
|
72 |
functions in the \texttt{*.fdl} file have already been loaded. |
|
73 |
\begin{figure} |
|
74 |
\begin{flushleft} |
|
75 |
\tt |
|
76 |
Context: \\ |
|
77 |
\ \\ |
|
78 |
\begin{tabular}{ll} |
|
63167 | 79 |
fixes & \<open>m ::\<close>\ "\<open>int\<close>" \\ |
80 |
and & \<open>n ::\<close>\ "\<open>int\<close>" \\ |
|
81 |
and & \<open>c ::\<close>\ "\<open>int\<close>" \\ |
|
82 |
and & \<open>d ::\<close>\ "\<open>int\<close>" \\ |
|
83 |
assumes & \<open>g_c_d_rules1:\<close>\ "\<open>0 \<le> integer__size\<close>" \\ |
|
84 |
and & \<open>g_c_d_rules6:\<close>\ "\<open>0 \<le> natural__size\<close>" \\ |
|
45044 | 85 |
\multicolumn{2}{l}{notes definition} \\ |
63167 | 86 |
\multicolumn{2}{l}{\hspace{2ex}\<open>defns =\<close>\ `\<open>integer__first = - 2147483648\<close>`} \\ |
87 |
\multicolumn{2}{l}{\hspace{4ex}`\<open>integer__last = 2147483647\<close>`} \\ |
|
45044 | 88 |
\multicolumn{2}{l}{\hspace{4ex}\dots} |
89 |
\end{tabular}\ \\[1.5ex] |
|
90 |
\ \\ |
|
91 |
Definitions: \\ |
|
92 |
\ \\ |
|
93 |
\begin{tabular}{ll} |
|
63167 | 94 |
\<open>g_c_d_rules2:\<close> & \<open>integer__first = - 2147483648\<close> \\ |
95 |
\<open>g_c_d_rules3:\<close> & \<open>integer__last = 2147483647\<close> \\ |
|
45044 | 96 |
\dots |
97 |
\end{tabular}\ \\[1.5ex] |
|
98 |
\ \\ |
|
99 |
Verification conditions: \\ |
|
100 |
\ \\ |
|
101 |
path(s) from assertion of line 10 to assertion of line 10 \\ |
|
102 |
\ \\ |
|
63167 | 103 |
\<open>procedure_g_c_d_4\<close>\ (unproved) \\ |
45044 | 104 |
\ \ \begin{tabular}{ll} |
63167 | 105 |
assumes & \<open>H1:\<close>\ "\<open>0 \<le> c\<close>" \\ |
106 |
and & \<open>H2:\<close>\ "\<open>0 < d\<close>" \\ |
|
107 |
and & \<open>H3:\<close>\ "\<open>gcd c d = gcd m n\<close>" \\ |
|
45044 | 108 |
\dots \\ |
63167 | 109 |
shows & "\<open>0 < c - c sdiv d * d\<close>" \\ |
110 |
and & "\<open>gcd d (c - c sdiv d * d) = gcd m n\<close> |
|
45044 | 111 |
\end{tabular}\ \\[1.5ex] |
112 |
\ \\ |
|
113 |
path(s) from assertion of line 10 to finish \\ |
|
114 |
\ \\ |
|
63167 | 115 |
\<open>procedure_g_c_d_11\<close>\ (unproved) \\ |
45044 | 116 |
\ \ \begin{tabular}{ll} |
63167 | 117 |
assumes & \<open>H1:\<close>\ "\<open>0 \<le> c\<close>" \\ |
118 |
and & \<open>H2:\<close>\ "\<open>0 < d\<close>" \\ |
|
119 |
and & \<open>H3:\<close>\ "\<open>gcd c d = gcd m n\<close>" \\ |
|
45044 | 120 |
\dots \\ |
63167 | 121 |
shows & "\<open>d = gcd m n\<close>" |
45044 | 122 |
\end{tabular} |
123 |
\end{flushleft} |
|
124 |
\caption{Output of \isa{\isacommand{spark\_status}} for \texttt{g\_c\_d.siv}} |
|
125 |
\label{fig:gcd-status} |
|
126 |
\end{figure} |
|
127 |
We now instruct Isabelle to open |
|
128 |
a new verification environment and load a set of VCs. This is done using the |
|
129 |
command \isa{\isacommand{spark\_open}}, which must be given the name of a |
|
56798
939e88e79724
Discontinued old spark_open; spark_open_siv is now spark_open
berghofe
parents:
45044
diff
changeset
|
130 |
\texttt{*.siv} file as an argument. Behind the scenes, Isabelle |
45044 | 131 |
parses this file and the corresponding \texttt{*.fdl} and \texttt{*.rls} files, |
132 |
and converts the VCs to Isabelle terms. Using the command \isa{\isacommand{spark\_status}}, |
|
133 |
the user can display the current VCs together with their status (proved, unproved). |
|
134 |
The variants \isa{\isacommand{spark\_status}\ (proved)} |
|
135 |
and \isa{\isacommand{spark\_status}\ (unproved)} show only proved and unproved |
|
136 |
VCs, respectively. For \texttt{g\_c\_d.siv}, the output of |
|
137 |
\isa{\isacommand{spark\_status}} is shown in \figref{fig:gcd-status}. |
|
138 |
To minimize the number of assumptions, and hence the size of the VCs, |
|
139 |
FDL rules of the form ``\dots\ \texttt{may\_be\_replaced\_by}\ \dots'' are |
|
140 |
turned into native Isabelle definitions, whereas other rules are modelled |
|
141 |
as assumptions. |
|
63167 | 142 |
\<close> |
45044 | 143 |
|
63167 | 144 |
section \<open>Proving the VCs\<close> |
45044 | 145 |
|
63167 | 146 |
text \<open> |
45044 | 147 |
\label{sec:proving-vcs} |
63167 | 148 |
The two open VCs are \<open>procedure_g_c_d_4\<close> and \<open>procedure_g_c_d_11\<close>, |
149 |
both of which contain the \<open>gcd\<close> proof function that the \SPARK{} Simplifier |
|
45044 | 150 |
does not know anything about. The proof of a particular VC can be started with |
151 |
the \isa{\isacommand{spark\_vc}} command, which is similar to the standard |
|
152 |
\isa{\isacommand{lemma}} and \isa{\isacommand{theorem}} commands, with the |
|
153 |
difference that it only takes a name of a VC but no formula as an argument. |
|
154 |
A VC can have several conclusions that can be referenced by the identifiers |
|
63167 | 155 |
\<open>?C1\<close>, \<open>?C2\<close>, etc. If there is just one conclusion, it can |
156 |
also be referenced by \<open>?thesis\<close>. It is important to note that the |
|
157 |
\texttt{div} operator of FDL behaves differently from the \<open>div\<close> operator |
|
45044 | 158 |
of Isabelle/HOL on negative numbers. The former always truncates towards zero, |
159 |
whereas the latter truncates towards minus infinity. This is why the FDL |
|
63167 | 160 |
\texttt{div} operator is mapped to the \<open>sdiv\<close> operator in Isabelle/HOL, |
45044 | 161 |
which is defined as |
162 |
@{thm [display] sdiv_def} |
|
163 |
For example, we have that |
|
164 |
@{lemma "-5 sdiv 4 = -1" by (simp add: sdiv_neg_pos)}, but |
|
165 |
@{lemma "(-5::int) div 4 = -2" by simp}. |
|
63167 | 166 |
For non-negative dividend and divisor, \<open>sdiv\<close> is equivalent to \<open>div\<close>, |
167 |
as witnessed by theorem \<open>sdiv_pos_pos\<close>: |
|
45044 | 168 |
@{thm [display,mode=no_brackets] sdiv_pos_pos} |
169 |
In contrast, the behaviour of the FDL \texttt{mod} operator is equivalent to |
|
170 |
the one of Isabelle/HOL. Moreover, since FDL has no counterpart of the \SPARK{} |
|
171 |
operator \textbf{rem}, the \SPARK{} expression \texttt{c}\ \textbf{rem}\ \texttt{d} |
|
63167 | 172 |
just becomes \<open>c - c sdiv d * d\<close> in Isabelle. The first conclusion of |
173 |
\<open>procedure_g_c_d_4\<close> requires us to prove that the remainder of \<open>c\<close> |
|
174 |
and \<open>d\<close> is greater than \<open>0\<close>. To do this, we use the theorem |
|
64246 | 175 |
\<open>minus_div_mult_eq_mod [symmetric]\<close> describing the correspondence between \<open>div\<close> |
63167 | 176 |
and \<open>mod\<close> |
64246 | 177 |
@{thm [display] minus_div_mult_eq_mod [symmetric]} |
63167 | 178 |
together with the theorem \<open>pos_mod_sign\<close> saying that the result of the |
179 |
\<open>mod\<close> operator is non-negative when applied to a non-negative divisor: |
|
45044 | 180 |
@{thm [display] pos_mod_sign} |
63167 | 181 |
We will also need the aforementioned theorem \<open>sdiv_pos_pos\<close> in order for |
182 |
the standard Isabelle/HOL theorems about \<open>div\<close> to be applicable |
|
183 |
to the VC, which is formulated using \<open>sdiv\<close> rather that \<open>div\<close>. |
|
184 |
Note that the proof uses \texttt{`\<open>0 \<le> c\<close>`} and \texttt{`\<open>0 < d\<close>`} |
|
185 |
rather than \<open>H1\<close> and \<open>H2\<close> to refer to the hypotheses of the current |
|
45044 | 186 |
VC. While the latter variant seems more compact, it is not particularly robust, |
187 |
since the numbering of hypotheses can easily change if the corresponding |
|
188 |
program is modified, making the proof script hard to adjust when there are many hypotheses. |
|
63167 | 189 |
Moreover, proof scripts using abbreviations like \<open>H1\<close> and \<open>H2\<close> |
45044 | 190 |
are hard to read without assistance from Isabelle. |
63167 | 191 |
The second conclusion of \<open>procedure_g_c_d_4\<close> requires us to prove that |
192 |
the \<open>gcd\<close> of \<open>d\<close> and the remainder of \<open>c\<close> and \<open>d\<close> |
|
193 |
is equal to the \<open>gcd\<close> of the original input values \<open>m\<close> and \<open>n\<close>, |
|
45044 | 194 |
which is the actual \emph{invariant} of the procedure. This is a consequence |
63167 | 195 |
of theorem \<open>gcd_non_0_int\<close> |
45044 | 196 |
@{thm [display] gcd_non_0_int} |
64246 | 197 |
Again, we also need theorems \<open>minus_div_mult_eq_mod [symmetric]\<close> and \<open>sdiv_pos_pos\<close> |
45044 | 198 |
to justify that \SPARK{}'s \textbf{rem} operator is equivalent to Isabelle's |
63167 | 199 |
\<open>mod\<close> operator for non-negative operands. |
200 |
The VC \<open>procedure_g_c_d_11\<close> says that if the loop invariant holds before |
|
45044 | 201 |
the last iteration of the loop, the postcondition of the procedure will hold |
202 |
after execution of the loop body. To prove this, we observe that the remainder |
|
63167 | 203 |
of \<open>c\<close> and \<open>d\<close>, and hence \<open>c mod d\<close> is \<open>0\<close> when exiting |
204 |
the loop. This implies that \<open>gcd c d = d\<close>, since \<open>c\<close> is divisible |
|
205 |
by \<open>d\<close>, so the conclusion follows using the assumption \<open>gcd c d = gcd m n\<close>. |
|
45044 | 206 |
This concludes the proofs of the open VCs, and hence the \SPARK{} verification |
207 |
environment can be closed using the command \isa{\isacommand{spark\_end}}. |
|
208 |
This command checks that all VCs have been proved and issues an error message |
|
209 |
if there are remaining unproved VCs. Moreover, Isabelle checks that there is |
|
210 |
no open \SPARK{} verification environment when the final \isa{\isacommand{end}} |
|
211 |
command of a theory is encountered. |
|
63167 | 212 |
\<close> |
45044 | 213 |
|
63167 | 214 |
section \<open>Optimizing the proof\<close> |
45044 | 215 |
|
63167 | 216 |
text \<open> |
45044 | 217 |
\begin{figure} |
218 |
\lstinputlisting{Simple_Gcd.adb} |
|
219 |
\input{Simple_Greatest_Common_Divisor} |
|
220 |
\caption{Simplified greatest common divisor program and proof} |
|
221 |
\label{fig:simple-gcd-proof} |
|
222 |
\end{figure} |
|
223 |
When looking at the program from \figref{fig:gcd-prog} once again, several |
|
224 |
optimizations come to mind. First of all, like the input parameters of the |
|
225 |
procedure, the local variables \texttt{C}, \texttt{D}, and \texttt{R} can |
|
226 |
be declared as \texttt{Natural} rather than \texttt{Integer}. Since natural |
|
227 |
numbers are non-negative by construction, the values computed by the algorithm |
|
228 |
are trivially proved to be non-negative. Since we are working with non-negative |
|
229 |
numbers, we can also just use \SPARK{}'s \textbf{mod} operator instead of |
|
64246 | 230 |
\textbf{rem}, which spares us an application of theorems \<open>minus_div_mult_eq_mod [symmetric]\<close> |
76987 | 231 |
and \<open>sdiv_pos_pos\<close>. Finally, as noted by Barnes \<^cite>\<open>\<open>\S 11.5\<close> in Barnes\<close>, |
45044 | 232 |
we can simplify matters by placing the \textbf{assert} statement between |
233 |
\textbf{while} and \textbf{loop} rather than directly after the \textbf{loop}. |
|
234 |
In the former case, the loop invariant has to be proved only once, whereas in |
|
235 |
the latter case, it has to be proved twice: since the \textbf{assert} occurs after |
|
236 |
the check of the exit condition, the invariant has to be proved for the path |
|
237 |
from the \textbf{assert} statement to the \textbf{assert} statement, and for |
|
238 |
the path from the \textbf{assert} statement to the postcondition. In the case |
|
239 |
of the \texttt{G\_C\_D} procedure, this might not seem particularly problematic, |
|
240 |
since the proof of the invariant is very simple, but it can unnecessarily |
|
241 |
complicate matters if the proof of the invariant is non-trivial. The simplified |
|
242 |
program for computing the greatest common divisor, together with its correctness |
|
243 |
proof, is shown in \figref{fig:simple-gcd-proof}. Since the package specification |
|
244 |
has not changed, we only show the body of the packages. The two VCs can now be |
|
63167 | 245 |
proved by a single application of Isabelle's proof method \<open>simp\<close>. |
246 |
\<close> |
|
45044 | 247 |
|
248 |
(*<*) |
|
249 |
end |
|
250 |
(*>*) |