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