| author | wenzelm | 
| Fri, 13 Jan 2023 14:38:19 +0100 | |
| changeset 76963 | a8566127d43b | 
| parent 74097 | 6d7be1227d02 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 45044 | 1 | (*<*) | 
| 2 | theory Reference | |
| 72514 | 3 | imports "HOL-SPARK.SPARK" | 
| 45044 | 4 | begin | 
| 5 | ||
| 74097 | 6 | unbundle bit_operations_syntax | 
| 7 | ||
| 72514 | 8 | lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int | 
| 9 | by (simp flip: mask_eq_exp_minus_1 take_bit_eq_mask take_bit_eq_mod) | |
| 10 | ||
| 45044 | 11 | syntax (my_constrain output) | 
| 61143 | 12 |   "_constrain" :: "logic => type => logic" ("_ :: _" [4, 0] 3)
 | 
| 45044 | 13 | (*>*) | 
| 14 | ||
| 63167 | 15 | chapter \<open>HOL-\SPARK{} Reference\<close>
 | 
| 45044 | 16 | |
| 63167 | 17 | text \<open> | 
| 45044 | 18 | \label{sec:spark-reference}
 | 
| 19 | This section is intended as a quick reference for the HOL-\SPARK{} verification
 | |
| 20 | environment. In \secref{sec:spark-commands}, we give a summary of the commands
 | |
| 21 | provided by the HOL-\SPARK{}, while \secref{sec:spark-types} contains a description
 | |
| 22 | of how particular types of \SPARK{} and FDL are modelled in Isabelle.
 | |
| 63167 | 23 | \<close> | 
| 45044 | 24 | |
| 63167 | 25 | section \<open>Commands\<close> | 
| 45044 | 26 | |
| 63167 | 27 | text \<open> | 
| 45044 | 28 | \label{sec:spark-commands}
 | 
| 29 | This section describes the syntax and effect of each of the commands provided | |
| 30 | by HOL-\SPARK{}.
 | |
| 69597 | 31 | \<^rail>\<open> | 
| 59938 | 32 |   @'spark_open' name ('(' name ')')?
 | 
| 69597 | 33 | \<close> | 
| 56798 
939e88e79724
Discontinued old spark_open; spark_open_siv is now spark_open
 berghofe parents: 
48168diff
changeset | 34 | Opens a new \SPARK{} verification environment and loads a \texttt{*.siv} file with VCs.
 | 
| 
939e88e79724
Discontinued old spark_open; spark_open_siv is now spark_open
 berghofe parents: 
48168diff
changeset | 35 | Alternatively, \texttt{*.vcg} files can be loaded using \isa{\isacommand{spark\_open\_vcg}}.
 | 
| 
939e88e79724
Discontinued old spark_open; spark_open_siv is now spark_open
 berghofe parents: 
48168diff
changeset | 36 | The corresponding \texttt{*.fdl} and \texttt{*.rls}
 | 
| 45044 | 37 | files must reside in the same directory as the file given as an argument to the command. | 
| 38 | This command also generates records and datatypes for the types specified in the | |
| 39 | \texttt{*.fdl} file, unless they have already been associated with user-defined
 | |
| 40 | Isabelle types (see below). | |
| 41 | Since the full package name currently cannot be determined from the files generated by the | |
| 42 | \SPARK{} Examiner, the command also allows to specify an optional package prefix in the
 | |
| 43 | format \texttt{$p_1$\_\_$\ldots$\_\_$p_n$}. When working with projects consisting of several
 | |
| 44 | packages, this is necessary in order for the verification environment to be able to map proof | |
| 45 | functions and types defined in Isabelle to their \SPARK{} counterparts.
 | |
| 69597 | 46 | \<^rail>\<open> | 
| 59938 | 47 | @'spark_proof_functions' ((name '=' term)+) | 
| 69597 | 48 | \<close> | 
| 45044 | 49 | Associates a proof function with the given name to a term. The name should be the full name | 
| 50 | of the proof function as it appears in the \texttt{*.fdl} file, including the package prefix.
 | |
| 51 | This command can be used both inside and outside a verification environment. The latter | |
| 52 | variant is useful for introducing proof functions that are shared by several procedures | |
| 53 | or packages, whereas the former allows the given term to refer to the types generated | |
| 54 | by \isa{\isacommand{spark\_open}} for record or enumeration types specified in the
 | |
| 55 | \texttt{*.fdl} file.
 | |
| 69597 | 56 | \<^rail>\<open> | 
| 59938 | 57 | @'spark_types' ((name '=' type (mapping?))+) | 
| 58 | ; | |
| 62969 | 59 |   mapping: '('((name '=' name)+',')')'
 | 
| 69597 | 60 | \<close> | 
| 45044 | 61 | Associates a \SPARK{} type with the given name with an Isabelle type. This command can
 | 
| 62 | only be used outside a verification environment. The given type must be either a record | |
| 46725 
d34ec0512dfb
Added infrastructure for mapping SPARK field / constructor names
 berghofe parents: 
45044diff
changeset | 63 | or a datatype, where the names of fields or constructors must either match those of the | 
| 
d34ec0512dfb
Added infrastructure for mapping SPARK field / constructor names
 berghofe parents: 
45044diff
changeset | 64 | corresponding \SPARK{} types (modulo casing), or a mapping from \SPARK{} to Isabelle
 | 
| 
d34ec0512dfb
Added infrastructure for mapping SPARK field / constructor names
 berghofe parents: 
45044diff
changeset | 65 | names has to be provided. | 
| 
d34ec0512dfb
Added infrastructure for mapping SPARK field / constructor names
 berghofe parents: 
45044diff
changeset | 66 | This command is useful when having to define | 
| 45044 | 67 | proof functions referring to record or enumeration types that are shared by several | 
| 68 | procedures or packages. First, the types required by the proof functions can be introduced | |
| 69 | using Isabelle's commands for defining records or datatypes. Having introduced the | |
| 70 | types, the proof functions can be defined in Isabelle. Finally, both the proof | |
| 71 | functions and the types can be associated with their \SPARK{} counterparts.
 | |
| 69597 | 72 | \<^rail>\<open> | 
| 59938 | 73 |   @'spark_status' (('(proved)' | '(unproved)')?)
 | 
| 69597 | 74 | \<close> | 
| 45044 | 75 | Outputs the variables declared in the \texttt{*.fdl} file, the rules declared in
 | 
| 76 | the \texttt{*.rls} file, and all VCs, together with their status (proved, unproved).
 | |
| 77 | The output can be restricted to the proved or unproved VCs by giving the corresponding | |
| 78 | option to the command. | |
| 69597 | 79 | \<^rail>\<open> | 
| 59938 | 80 | @'spark_vc' name | 
| 69597 | 81 | \<close> | 
| 45044 | 82 | Initiates the proof of the VC with the given name. Similar to the standard | 
| 83 | \isa{\isacommand{lemma}} or \isa{\isacommand{theorem}} commands, this command
 | |
| 84 | must be followed by a sequence of proof commands. The command introduces the | |
| 85 | hypotheses \texttt{H1} \dots \texttt{H$n$}, as well as the identifiers
 | |
| 86 | \texttt{?C1} \dots \texttt{?C$m$} corresponding to the conclusions of the VC.
 | |
| 69597 | 87 | \<^rail>\<open> | 
| 59938 | 88 | @'spark_end' '(incomplete)'? | 
| 69597 | 89 | \<close> | 
| 48168 | 90 | Closes the current verification environment. Unless the \texttt{incomplete}
 | 
| 91 | option is given, all VCs must have been proved, | |
| 45044 | 92 | otherwise the command issues an error message. As a side effect, the command | 
| 93 | generates a proof review (\texttt{*.prv}) file to inform POGS of the proved
 | |
| 94 | VCs. | |
| 63167 | 95 | \<close> | 
| 45044 | 96 | |
| 63167 | 97 | section \<open>Types\<close> | 
| 45044 | 98 | |
| 63167 | 99 | text \<open> | 
| 45044 | 100 | \label{sec:spark-types}
 | 
| 101 | The main types of FDL are integers, enumeration types, records, and arrays. | |
| 102 | In the following sections, we describe how these types are modelled in | |
| 103 | Isabelle. | |
| 63167 | 104 | \<close> | 
| 45044 | 105 | |
| 63167 | 106 | subsection \<open>Integers\<close> | 
| 45044 | 107 | |
| 63167 | 108 | text \<open> | 
| 69597 | 109 | The FDL type \texttt{integer} is modelled by the Isabelle type \<^typ>\<open>int\<close>.
 | 
| 45044 | 110 | While the FDL \texttt{mod} operator behaves in the same way as its Isabelle
 | 
| 111 | counterpart, this is not the case for the \texttt{div} operator. As has already
 | |
| 112 | been mentioned in \secref{sec:proving-vcs}, the \texttt{div} operator of \SPARK{}
 | |
| 63167 | 113 | always truncates towards zero, whereas the \<open>div\<close> operator of Isabelle | 
| 45044 | 114 | truncates towards minus infinity. Therefore, the FDL \texttt{div} operator is
 | 
| 63167 | 115 | mapped to the \<open>sdiv\<close> operator in Isabelle. The characteristic theorems | 
| 116 | of \<open>sdiv\<close>, in particular those describing the relationship with the standard | |
| 117 | \<open>div\<close> operator, are shown in \figref{fig:sdiv-properties}
 | |
| 45044 | 118 | \begin{figure}
 | 
| 119 | \begin{center}
 | |
| 120 | \small | |
| 121 | \begin{tabular}{ll}
 | |
| 63167 | 122 | \<open>sdiv_def\<close>: & @{thm sdiv_def} \\
 | 
| 123 | \<open>sdiv_minus_dividend\<close>: & @{thm sdiv_minus_dividend} \\
 | |
| 124 | \<open>sdiv_minus_divisor\<close>: & @{thm sdiv_minus_divisor} \\
 | |
| 125 | \<open>sdiv_pos_pos\<close>: & @{thm [mode=no_brackets] sdiv_pos_pos} \\
 | |
| 126 | \<open>sdiv_pos_neg\<close>: & @{thm [mode=no_brackets] sdiv_pos_neg} \\
 | |
| 127 | \<open>sdiv_neg_pos\<close>: & @{thm [mode=no_brackets] sdiv_neg_pos} \\
 | |
| 128 | \<open>sdiv_neg_neg\<close>: & @{thm [mode=no_brackets] sdiv_neg_neg} \\
 | |
| 45044 | 129 | \end{tabular}
 | 
| 130 | \end{center}
 | |
| 63167 | 131 | \caption{Characteristic properties of \<open>sdiv\<close>}
 | 
| 45044 | 132 | \label{fig:sdiv-properties}
 | 
| 133 | \end{figure}
 | |
| 134 | ||
| 135 | \begin{figure}
 | |
| 136 | \begin{center}
 | |
| 137 | \small | |
| 138 | \begin{tabular}{ll}
 | |
| 63167 | 139 | \<open>AND_lower\<close>: & @{thm [mode=no_brackets] AND_lower} \\
 | 
| 140 | \<open>OR_lower\<close>: & @{thm [mode=no_brackets] OR_lower} \\
 | |
| 141 | \<open>XOR_lower\<close>: & @{thm [mode=no_brackets] XOR_lower} \\
 | |
| 142 | \<open>AND_upper1\<close>: & @{thm [mode=no_brackets] AND_upper1} \\
 | |
| 143 | \<open>AND_upper2\<close>: & @{thm [mode=no_brackets] AND_upper2} \\
 | |
| 144 | \<open>OR_upper\<close>: & @{thm [mode=no_brackets] OR_upper} \\
 | |
| 145 | \<open>XOR_upper\<close>: & @{thm [mode=no_brackets] XOR_upper} \\
 | |
| 146 | \<open>AND_mod\<close>: & @{thm [mode=no_brackets] AND_mod}
 | |
| 45044 | 147 | \end{tabular}
 | 
| 148 | \end{center}
 | |
| 149 | \caption{Characteristic properties of bitwise operators}
 | |
| 150 | \label{fig:bitwise}
 | |
| 151 | \end{figure}
 | |
| 152 | The bitwise logical operators of \SPARK{} and FDL are modelled by the operators
 | |
| 63167 | 153 | \<open>AND\<close>, \<open>OR\<close> and \<open>XOR\<close> from Isabelle's \<open>Word\<close> library, | 
| 69597 | 154 | all of which have type \<^typ>\<open>int \<Rightarrow> int \<Rightarrow> int\<close>. A list of properties of these | 
| 45044 | 155 | operators that are useful in proofs about \SPARK{} programs are shown in
 | 
| 156 | \figref{fig:bitwise}
 | |
| 63167 | 157 | \<close> | 
| 45044 | 158 | |
| 63167 | 159 | subsection \<open>Enumeration types\<close> | 
| 45044 | 160 | |
| 63167 | 161 | text \<open> | 
| 45044 | 162 | The FDL enumeration type | 
| 163 | \begin{alltt}
 | |
| 164 | type \(t\) = (\(e\sb{1}\), \(e\sb{2}\), \dots, \(e\sb{n}\));
 | |
| 165 | \end{alltt}
 | |
| 166 | is modelled by the Isabelle datatype | |
| 167 | \begin{isabelle}
 | |
| 168 | \normalsize | |
| 169 | \isacommand{datatype}\ $t$\ =\ $e_1$\ $\mid$\ $e_2$\ $\mid$\ \dots\ $\mid$\ $e_n$
 | |
| 170 | \end{isabelle}
 | |
| 69597 | 171 | The HOL-\SPARK{} environment defines a type class \<^class>\<open>spark_enum\<close> that captures
 | 
| 45044 | 172 | the characteristic properties of all enumeration types. It provides the following | 
| 63167 | 173 | polymorphic functions and constants for all types \<open>'a\<close> of this type class: | 
| 45044 | 174 | \begin{flushleft}
 | 
| 175 | @{term_type [mode=my_constrain] pos} \\
 | |
| 176 | @{term_type [mode=my_constrain] val} \\
 | |
| 177 | @{term_type [mode=my_constrain] succ} \\
 | |
| 178 | @{term_type [mode=my_constrain] pred} \\
 | |
| 179 | @{term_type [mode=my_constrain] first_el} \\
 | |
| 180 | @{term_type [mode=my_constrain] last_el}
 | |
| 181 | \end{flushleft}
 | |
| 69597 | 182 | In addition, \<^class>\<open>spark_enum\<close> is a subclass of the \<^class>\<open>linorder\<close> type class, | 
| 63167 | 183 | which allows the comparison operators \<open><\<close> and \<open>\<le>\<close> to be used on | 
| 45044 | 184 | enumeration types. The polymorphic operations shown above enjoy a number of | 
| 185 | generic properties that hold for all enumeration types. These properties are | |
| 186 | listed in \figref{fig:enum-generic-properties}.
 | |
| 187 | Moreover, \figref{fig:enum-specific-properties} shows a list of properties
 | |
| 188 | that are specific to each enumeration type $t$, such as the characteristic | |
| 69597 | 189 | equations for \<^term>\<open>val\<close> and \<^term>\<open>pos\<close>. | 
| 45044 | 190 | \begin{figure}[t]
 | 
| 191 | \begin{center}
 | |
| 192 | \small | |
| 193 | \begin{tabular}{ll}
 | |
| 63167 | 194 | \<open>range_pos\<close>: & @{thm range_pos} \\
 | 
| 195 | \<open>less_pos\<close>: & @{thm less_pos} \\
 | |
| 196 | \<open>less_eq_pos\<close>: & @{thm less_eq_pos} \\
 | |
| 197 | \<open>val_def\<close>: & @{thm val_def} \\
 | |
| 198 | \<open>succ_def\<close>: & @{thm succ_def} \\
 | |
| 199 | \<open>pred_def\<close>: & @{thm pred_def} \\
 | |
| 200 | \<open>first_el_def\<close>: & @{thm first_el_def} \\
 | |
| 201 | \<open>last_el_def\<close>: & @{thm last_el_def} \\
 | |
| 202 | \<open>inj_pos\<close>: & @{thm inj_pos} \\
 | |
| 203 | \<open>val_pos\<close>: & @{thm val_pos} \\
 | |
| 204 | \<open>pos_val\<close>: & @{thm pos_val} \\
 | |
| 205 | \<open>first_el_smallest\<close>: & @{thm first_el_smallest} \\
 | |
| 206 | \<open>last_el_greatest\<close>: & @{thm last_el_greatest} \\
 | |
| 207 | \<open>pos_succ\<close>: & @{thm pos_succ} \\
 | |
| 208 | \<open>pos_pred\<close>: & @{thm pos_pred} \\
 | |
| 209 | \<open>succ_val\<close>: & @{thm succ_val} \\
 | |
| 210 | \<open>pred_val\<close>: & @{thm pred_val}
 | |
| 45044 | 211 | \end{tabular}
 | 
| 212 | \end{center}
 | |
| 213 | \caption{Generic properties of functions on enumeration types}
 | |
| 214 | \label{fig:enum-generic-properties}
 | |
| 215 | \end{figure}
 | |
| 216 | \begin{figure}[t]
 | |
| 217 | \begin{center}
 | |
| 218 | \small | |
| 219 | \begin{tabular}{ll@ {\hspace{2cm}}ll}
 | |
| 220 | \texttt{$t$\_val}: & \isa{val\ $0$\ =\ $e_1$} & \texttt{$t$\_pos}: & pos\ $e_1$\ =\ $0$ \\
 | |
| 221 |                    & \isa{val\ $1$\ =\ $e_2$} &                    & pos\ $e_2$\ =\ $1$ \\
 | |
| 222 |                    & \hspace{1cm}\vdots       &                    & \hspace{1cm}\vdots \\
 | |
| 223 |                    & \isa{val\ $(n-1)$\ =\ $e_n$} &                & pos\ $e_n$\ =\ $n-1$
 | |
| 224 | \end{tabular} \\[3ex]
 | |
| 225 | \begin{tabular}{ll}
 | |
| 226 | \texttt{$t$\_card}: & \isa{card($t$)\ =\ $n$} \\
 | |
| 227 | \texttt{$t$\_first\_el}: & \isa{first\_el\ =\ $e_1$} \\
 | |
| 228 | \texttt{$t$\_last\_el}: & \isa{last\_el\ =\ $e_n$}
 | |
| 229 | \end{tabular}
 | |
| 230 | \end{center}
 | |
| 231 | \caption{Type-specific properties of functions on enumeration types}
 | |
| 232 | \label{fig:enum-specific-properties}
 | |
| 233 | \end{figure}
 | |
| 63167 | 234 | \<close> | 
| 45044 | 235 | |
| 63167 | 236 | subsection \<open>Records\<close> | 
| 45044 | 237 | |
| 63167 | 238 | text \<open> | 
| 45044 | 239 | The FDL record type | 
| 240 | \begin{alltt}
 | |
| 241 | type \(t\) = record | |
| 242 |       \(f\sb{1}\) : \(t\sb{1}\);
 | |
| 243 | \(\vdots\) | |
| 244 |       \(f\sb{n}\) : \(t\sb{n}\)
 | |
| 245 | end; | |
| 246 | \end{alltt}
 | |
| 247 | is modelled by the Isabelle record type | |
| 248 | \begin{isabelle}
 | |
| 249 | \normalsize | |
| 250 | \isacommand{record}\ t\ = \isanewline
 | |
| 251 | \ \ $f_1$\ ::\ $t_1$ \isanewline | |
| 252 | \ \ \ \vdots \isanewline | |
| 253 | \ \ $f_n$\ ::\ $t_n$ | |
| 254 | \end{isabelle}
 | |
| 255 | Records are constructed using the notation | |
| 256 | \isa{\isasymlparr$f_1$\ =\ $v_1$,\ $\ldots$,\ $f_n$\ =\ $v_n$\isasymrparr},
 | |
| 257 | a field $f_i$ of a record $r$ is selected using the notation $f_i~r$, and the | |
| 258 | fields $f$ and $f'$ of a record $r$ can be updated using the notation | |
| 259 | \mbox{\isa{$r$\ \isasymlparr$f$\ :=\ $v$,\ $f'$\ :=\ $v'$\isasymrparr}}.
 | |
| 63167 | 260 | \<close> | 
| 45044 | 261 | |
| 63167 | 262 | subsection \<open>Arrays\<close> | 
| 45044 | 263 | |
| 63167 | 264 | text \<open> | 
| 45044 | 265 | The FDL array type | 
| 266 | \begin{alltt}
 | |
| 267 | type \(t\) = array [\(t\sb{1}\), \(\ldots\), \(t\sb{n}\)] of \(u\);
 | |
| 268 | \end{alltt}
 | |
| 269 | is modelled by the Isabelle function type $t_1 \times \cdots \times t_n \Rightarrow u$. | |
| 270 | Array updates are written as \isa{$A$($x_1$\ := $y_1$,\ \dots,\ $x_n$\ :=\ $y_n$)}.
 | |
| 271 | To allow updating an array at a set of indices, HOL-\SPARK{} provides the notation
 | |
| 272 | \isa{\dots\ [:=]\ \dots}, which can be combined with \isa{\dots\ :=\ \dots} and has
 | |
| 273 | the properties | |
| 274 | @{thm [display,mode=no_brackets] fun_upds_in fun_upds_notin upds_singleton}
 | |
| 275 | Thus, we can write expressions like | |
| 276 | @{term [display] "(A::int\<Rightarrow>int) ({0..9} [:=] 42, 15 := 99, {20..29} [:=] 0)"}
 | |
| 277 | that would be cumbersome to write using single updates. | |
| 63167 | 278 | \<close> | 
| 45044 | 279 | |
| 63167 | 280 | section \<open>User-defined proof functions and types\<close> | 
| 45044 | 281 | |
| 63167 | 282 | text \<open> | 
| 45044 | 283 | To illustrate the interplay between the commands for introducing user-defined proof | 
| 284 | functions and types mentioned in \secref{sec:spark-commands}, we now discuss a larger
 | |
| 285 | example involving the definition of proof functions on complex types. Assume we would | |
| 286 | like to define an array type, whose elements are records that themselves contain | |
| 287 | arrays. Moreover, assume we would like to initialize all array elements and record | |
| 288 | fields of type \texttt{Integer} in an array of this type with the value \texttt{0}.
 | |
| 289 | The specification of package \texttt{Complex\_Types} containing the definition of
 | |
| 290 | the array type, which we call \texttt{Array\_Type2}, is shown in \figref{fig:complex-types}.
 | |
| 291 | It also contains the declaration of a proof function \texttt{Initialized} that is used
 | |
| 292 | to express that the array has been initialized. The two other proof functions | |
| 293 | \texttt{Initialized2} and \texttt{Initialized3} are used to reason about the
 | |
| 294 | initialization of the inner array. Since the array types and proof functions | |
| 295 | may be used by several packages, such as the one shown in \figref{fig:complex-types-app},
 | |
| 296 | it is advantageous to define the proof functions in a central theory that can | |
| 297 | be included by other theories containing proofs about packages using \texttt{Complex\_Types}.
 | |
| 298 | We show this theory in \figref{fig:complex-types-thy}. Since the proof functions
 | |
| 299 | refer to the enumeration and record types defined in \texttt{Complex\_Types},
 | |
| 300 | we need to define the Isabelle counterparts of these types using the | |
| 301 | \isa{\isacommand{datatype}} and \isa{\isacommand{record}} commands in order
 | |
| 302 | to be able to write down the definition of the proof functions. These types are | |
| 303 | linked to the corresponding \SPARK{} types using the \isa{\isacommand{spark\_types}}
 | |
| 304 | command. Note that we have to specify the full name of the \SPARK{} functions
 | |
| 305 | including the package prefix. Using the logic of Isabelle, we can then define | |
| 306 | functions involving the enumeration and record types introduced above, and link | |
| 307 | them to the corresponding \SPARK{} proof functions. It is important that the
 | |
| 308 | \isa{\isacommand{definition}} commands are preceeded by the \isa{\isacommand{spark\_types}}
 | |
| 63167 | 309 | command, since the definition of \<open>initialized3\<close> uses the \<open>val\<close> | 
| 310 | function for enumeration types that is only available once that \<open>day\<close> | |
| 45044 | 311 | has been declared as a \SPARK{} type.
 | 
| 312 | \begin{figure}
 | |
| 313 | \lstinputlisting{complex_types.ads}
 | |
| 314 | \caption{Nested array and record types}
 | |
| 315 | \label{fig:complex-types}
 | |
| 316 | \end{figure}
 | |
| 317 | \begin{figure}
 | |
| 318 | \lstinputlisting{complex_types_app.ads}
 | |
| 319 | \lstinputlisting{complex_types_app.adb}
 | |
| 320 | \caption{Application of \texttt{Complex\_Types} package}
 | |
| 321 | \label{fig:complex-types-app}
 | |
| 322 | \end{figure}
 | |
| 323 | \begin{figure}
 | |
| 324 | \input{Complex_Types}
 | |
| 325 | \caption{Theory defining proof functions for complex types}
 | |
| 326 | \label{fig:complex-types-thy}
 | |
| 327 | \end{figure}
 | |
| 63167 | 328 | \<close> | 
| 45044 | 329 | |
| 330 | (*<*) | |
| 331 | end | |
| 332 | (*>*) |