| author | blanchet | 
| Thu, 22 Apr 2010 15:01:36 +0200 | |
| changeset 36288 | 156e4f179bb0 | 
| parent 30184 | 37969710e61f | 
| child 39838 | eb47307ab930 | 
| permissions | -rw-r--r-- | 
| 145 | 1 | |
| 104 | 2 | \chapter{Theories, Terms and Types} \label{theories}
 | 
| 30184 
37969710e61f
removed parts of the manual that are clearly obsolete, or covered by
 wenzelm parents: 
19627diff
changeset | 3 | \index{theories|(}
 | 
| 104 | 4 | |
| 6568 | 5 | \section{The theory loader}\label{sec:more-theories}
 | 
| 6 | \index{theories!reading}\index{files!reading}
 | |
| 7 | ||
| 8 | Isabelle's theory loader manages dependencies of the internal graph of theory | |
| 9 | nodes (the \emph{theory database}) and the external view of the file system.
 | |
| 10 | See \S\ref{sec:intro-theories} for its most basic commands, such as
 | |
| 11 | \texttt{use_thy}.  There are a few more operations available.
 | |
| 12 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 13 | \begin{ttbox}
 | 
| 6568 | 14 | use_thy_only : string -> unit | 
| 7136 | 15 | update_thy_only : string -> unit | 
| 6568 | 16 | touch_thy : string -> unit | 
| 6658 | 17 | remove_thy : string -> unit | 
| 8136 | 18 | delete_tmpfiles : bool ref \hfill\textbf{initially true}
 | 
| 286 | 19 | \end{ttbox}
 | 
| 20 | ||
| 324 | 21 | \begin{ttdescription}
 | 
| 6569 | 22 | \item[\ttindexbold{use_thy_only} "$name$";] is similar to \texttt{use_thy},
 | 
| 23 |   but processes the actual theory file $name$\texttt{.thy} only, ignoring
 | |
| 6568 | 24 |   $name$\texttt{.ML}.  This might be useful in replaying proof scripts by hand
 | 
| 25 | from the very beginning, starting with the fresh theory. | |
| 26 | ||
| 7136 | 27 | \item[\ttindexbold{update_thy_only} "$name$";] is similar to
 | 
| 28 |   \texttt{update_thy}, but processes the actual theory file
 | |
| 29 |   $name$\texttt{.thy} only, ignoring $name$\texttt{.ML}.
 | |
| 30 | ||
| 6569 | 31 | \item[\ttindexbold{touch_thy} "$name$";] marks theory node $name$ of the
 | 
| 6568 | 32 | internal graph as outdated. While the theory remains usable, subsequent | 
| 33 |   operations such as \texttt{use_thy} may cause a reload.
 | |
| 34 | ||
| 6658 | 35 | \item[\ttindexbold{remove_thy} "$name$";] deletes theory node $name$,
 | 
| 36 |   including \emph{all} of its descendants.  Beware!  This is a quick way to
 | |
| 37 |   dispose a large number of theories at once.  Note that {\ML} bindings to
 | |
| 38 | theorems etc.\ of removed theories may still persist. | |
| 39 | ||
| 324 | 40 | \end{ttdescription}
 | 
| 138 
9ba8bff1addc
added chapter "Defining Theories" and made changes for new Readthy functions
 clasohm parents: 
104diff
changeset | 41 | |
| 6568 | 42 | \medskip Theory and {\ML} files are located by skimming through the
 | 
| 43 | directories listed in Isabelle's internal load path, which merely contains the | |
| 44 | current directory ``\texttt{.}'' by default.  The load path may be accessed by
 | |
| 45 | the following operations. | |
| 332 | 46 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 47 | \begin{ttbox}
 | 
| 6568 | 48 | show_path: unit -> string list | 
| 49 | add_path: string -> unit | |
| 50 | del_path: string -> unit | |
| 51 | reset_path: unit -> unit | |
| 7440 | 52 | with_path: string -> ('a -> 'b) -> 'a -> 'b
 | 
| 11052 | 53 | no_document: ('a -> 'b) -> 'a -> 'b
 | 
| 286 | 54 | \end{ttbox}
 | 
| 55 | ||
| 324 | 56 | \begin{ttdescription}
 | 
| 6568 | 57 | \item[\ttindexbold{show_path}();] displays the load path components in
 | 
| 6571 | 58 | canonical string representation (which is always according to Unix rules). | 
| 6568 | 59 | |
| 6569 | 60 | \item[\ttindexbold{add_path} "$dir$";] adds component $dir$ to the beginning
 | 
| 61 | of the load path. | |
| 6568 | 62 | |
| 6569 | 63 | \item[\ttindexbold{del_path} "$dir$";] removes any occurrences of component
 | 
| 6568 | 64 | $dir$ from the load path. | 
| 65 | ||
| 66 | \item[\ttindexbold{reset_path}();] resets the load path to ``\texttt{.}''
 | |
| 67 | (current directory) only. | |
| 7440 | 68 | |
| 69 | \item[\ttindexbold{with_path} "$dir$" $f$ $x$;] temporarily adds component
 | |
| 11052 | 70 | $dir$ to the beginning of the load path while executing $(f~x)$. | 
| 71 | ||
| 72 | \item[\ttindexbold{no_document} $f$ $x$;] temporarily disables {\LaTeX}
 | |
| 73 | document generation while executing $(f~x)$. | |
| 7440 | 74 | |
| 324 | 75 | \end{ttdescription}
 | 
| 286 | 76 | |
| 7440 | 77 | Furthermore, in operations referring indirectly to some file (e.g.\ | 
| 78 | \texttt{use_dir}) the argument may be prefixed by a directory that will be
 | |
| 79 | temporarily appended to the load path, too. | |
| 104 | 80 | |
| 81 | ||
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
864diff
changeset | 82 | \section{Basic operations on theories}\label{BasicOperationsOnTheories}
 | 
| 4384 | 83 | |
| 84 | \subsection{*Theory inclusion}
 | |
| 85 | \begin{ttbox}
 | |
| 86 | subthy : theory * theory -> bool | |
| 87 | eq_thy : theory * theory -> bool | |
| 88 | transfer : theory -> thm -> thm | |
| 89 | transfer_sg : Sign.sg -> thm -> thm | |
| 90 | \end{ttbox}
 | |
| 91 | ||
| 92 | Inclusion and equality of theories is determined by unique | |
| 93 | identification stamps that are created when declaring new components. | |
| 94 | Theorems contain a reference to the theory (actually to its signature) | |
| 95 | they have been derived in. Transferring theorems to super theories | |
| 96 | has no logical significance, but may affect some operations in subtle | |
| 97 | ways (e.g.\ implicit merges of signatures when applying rules, or | |
| 98 | pretty printing of theorems). | |
| 99 | ||
| 100 | \begin{ttdescription}
 | |
| 101 | ||
| 102 | \item[\ttindexbold{subthy} ($thy@1$, $thy@2$)] determines if $thy@1$
 | |
| 103 | is included in $thy@2$ wrt.\ identification stamps. | |
| 104 | ||
| 105 | \item[\ttindexbold{eq_thy} ($thy@1$, $thy@2$)] determines if $thy@1$
 | |
| 106 | is exactly the same as $thy@2$. | |
| 107 | ||
| 108 | \item[\ttindexbold{transfer} $thy$ $thm$] transfers theorem $thm$ to
 | |
| 109 | theory $thy$, provided the latter includes the theory of $thm$. | |
| 110 | ||
| 111 | \item[\ttindexbold{transfer_sg} $sign$ $thm$] is similar to
 | |
| 112 |   \texttt{transfer}, but identifies the super theory via its
 | |
| 113 | signature. | |
| 114 | ||
| 115 | \end{ttdescription}
 | |
| 116 | ||
| 117 | ||
| 6571 | 118 | \subsection{*Primitive theories}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 119 | \begin{ttbox}
 | 
| 4317 | 120 | ProtoPure.thy : theory | 
| 3108 | 121 | Pure.thy : theory | 
| 122 | CPure.thy : theory | |
| 286 | 123 | \end{ttbox}
 | 
| 3108 | 124 | \begin{description}
 | 
| 4317 | 125 | \item[\ttindexbold{ProtoPure.thy}, \ttindexbold{Pure.thy},
 | 
| 126 |   \ttindexbold{CPure.thy}] contain the syntax and signature of the
 | |
| 127 | meta-logic. There are basically no axioms: meta-level inferences | |
| 128 |   are carried out by \ML\ functions.  \texttt{Pure} and \texttt{CPure}
 | |
| 129 | just differ in their concrete syntax of prefix function application: | |
| 130 |   $t(u@1, \ldots, u@n)$ in \texttt{Pure} vs.\ $t\,u@1,\ldots\,u@n$ in
 | |
| 131 |   \texttt{CPure}.  \texttt{ProtoPure} is their common parent,
 | |
| 132 | containing no syntax for printing prefix applications at all! | |
| 6571 | 133 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 134 | %% FIXME | 
| 478 | 135 | %\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"} $\cdots$] extends
 | 
| 136 | % the theory $thy$ with new types, constants, etc. $T$ identifies the theory | |
| 137 | % internally. When a theory is redeclared, say to change an incorrect axiom, | |
| 138 | % bindings to the old axiom may persist. Isabelle ensures that the old and | |
| 139 | % new theories are not involved in the same proof. Attempting to combine | |
| 140 | % different theories having the same name $T$ yields the fatal error | |
| 141 | %extend_theory : theory -> string -> \(\cdots\) -> theory | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 142 | %\begin{ttbox}
 | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 143 | %Attempt to merge different versions of theory: \(T\) | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 144 | %\end{ttbox}
 | 
| 3108 | 145 | \end{description}
 | 
| 286 | 146 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 147 | %% FIXME | 
| 275 | 148 | %\item [\ttindexbold{extend_theory} $thy$ {\tt"}$T${\tt"}
 | 
| 149 | % ($classes$, $default$, $types$, $arities$, $consts$, $sextopt$) $rules$] | |
| 150 | %\hfill\break %%% include if line is just too short | |
| 286 | 151 | %is the \ML{} equivalent of the following theory definition:
 | 
| 275 | 152 | %\begin{ttbox}
 | 
| 153 | %\(T\) = \(thy\) + | |
| 154 | %classes \(c\) < \(c@1\),\(\dots\),\(c@m\) | |
| 155 | % \dots | |
| 156 | %default {\(d@1,\dots,d@r\)}
 | |
| 157 | %types \(tycon@1\),\dots,\(tycon@i\) \(n\) | |
| 158 | % \dots | |
| 159 | %arities \(tycon@1'\),\dots,\(tycon@j'\) :: (\(s@1\),\dots,\(s@n\))\(c\) | |
| 160 | % \dots | |
| 161 | %consts \(b@1\),\dots,\(b@k\) :: \(\tau\) | |
| 162 | % \dots | |
| 163 | %rules \(name\) \(rule\) | |
| 164 | % \dots | |
| 165 | %end | |
| 166 | %\end{ttbox}
 | |
| 167 | %where | |
| 168 | %\begin{tabular}[t]{l@{~=~}l}
 | |
| 169 | %$classes$ & \tt[("$c$",["$c@1$",\dots,"$c@m$"]),\dots] \\
 | |
| 170 | %$default$ & \tt["$d@1$",\dots,"$d@r$"]\\ | |
| 171 | %$types$ & \tt[([$tycon@1$,\dots,$tycon@i$], $n$),\dots] \\ | |
| 172 | %$arities$ & \tt[([$tycon'@1$,\dots,$tycon'@j$], ([$s@1$,\dots,$s@n$],$c$)),\dots] | |
| 173 | %\\ | |
| 174 | %$consts$ & \tt[([$b@1$,\dots,$b@k$],$\tau$),\dots] \\ | |
| 175 | %$rules$   & \tt[("$name$",$rule$),\dots]
 | |
| 176 | %\end{tabular}
 | |
| 104 | 177 | |
| 178 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 179 | \subsection{Inspecting a theory}\label{sec:inspct-thy}
 | 
| 104 | 180 | \index{theories!inspecting|bold}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 181 | \begin{ttbox}
 | 
| 4317 | 182 | print_syntax : theory -> unit | 
| 183 | print_theory : theory -> unit | |
| 184 | parents_of : theory -> theory list | |
| 185 | ancestors_of : theory -> theory list | |
| 186 | sign_of : theory -> Sign.sg | |
| 187 | Sign.stamp_names_of : Sign.sg -> string list | |
| 104 | 188 | \end{ttbox}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 189 | These provide means of viewing a theory's components. | 
| 324 | 190 | \begin{ttdescription}
 | 
| 3108 | 191 | \item[\ttindexbold{print_syntax} $thy$] prints the syntax of $thy$
 | 
| 192 | (grammar, macros, translation functions etc., see | |
| 193 |   page~\pageref{pg:print_syn} for more details).
 | |
| 194 | ||
| 195 | \item[\ttindexbold{print_theory} $thy$] prints the logical parts of
 | |
| 196 | $thy$, excluding the syntax. | |
| 4317 | 197 | |
| 198 | \item[\ttindexbold{parents_of} $thy$] returns the direct ancestors
 | |
| 199 | of~$thy$. | |
| 200 | ||
| 201 | \item[\ttindexbold{ancestors_of} $thy$] returns all ancestors of~$thy$
 | |
| 202 | (not including $thy$ itself). | |
| 203 | ||
| 204 | \item[\ttindexbold{sign_of} $thy$] returns the signature associated
 | |
| 205 |   with~$thy$.  It is useful with functions like {\tt
 | |
| 206 | read_instantiate_sg}, which take a signature as an argument. | |
| 207 | ||
| 208 | \item[\ttindexbold{Sign.stamp_names_of} $sg$]\index{signatures}
 | |
| 209 |   returns the names of the identification \rmindex{stamps} of ax
 | |
| 210 | signature. These coincide with the names of its full ancestry | |
| 211 | including that of $sg$ itself. | |
| 104 | 212 | |
| 324 | 213 | \end{ttdescription}
 | 
| 104 | 214 | |
| 1369 | 215 | |
| 11623 | 216 | \section{Terms}\label{sec:terms}
 | 
| 104 | 217 | \index{terms|bold}
 | 
| 324 | 218 | Terms belong to the \ML\ type \mltydx{term}, which is a concrete datatype
 | 
| 3108 | 219 | with six constructors: | 
| 104 | 220 | \begin{ttbox}
 | 
| 221 | type indexname = string * int; | |
| 222 | infix 9 $; | |
| 223 | datatype term = Const of string * typ | |
| 224 | | Free of string * typ | |
| 225 | | Var of indexname * typ | |
| 226 | | Bound of int | |
| 227 | | Abs of string * typ * term | |
| 228 | | op $ of term * term; | |
| 229 | \end{ttbox}
 | |
| 324 | 230 | \begin{ttdescription}
 | 
| 4317 | 231 | \item[\ttindexbold{Const} ($a$, $T$)] \index{constants|bold}
 | 
| 8136 | 232 |   is the \textbf{constant} with name~$a$ and type~$T$.  Constants include
 | 
| 286 | 233 | connectives like $\land$ and $\forall$ as well as constants like~0 | 
| 234 | and~$Suc$. Other constants may be required to define a logic's concrete | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 235 | syntax. | 
| 104 | 236 | |
| 4317 | 237 | \item[\ttindexbold{Free} ($a$, $T$)] \index{variables!free|bold}
 | 
| 8136 | 238 |   is the \textbf{free variable} with name~$a$ and type~$T$.
 | 
| 104 | 239 | |
| 4317 | 240 | \item[\ttindexbold{Var} ($v$, $T$)] \index{unknowns|bold}
 | 
| 8136 | 241 |   is the \textbf{scheme variable} with indexname~$v$ and type~$T$.  An
 | 
| 324 | 242 |   \mltydx{indexname} is a string paired with a non-negative index, or
 | 
| 243 | subscript; a term's scheme variables can be systematically renamed by | |
| 244 | incrementing their subscripts. Scheme variables are essentially free | |
| 245 | variables, but may be instantiated during unification. | |
| 104 | 246 | |
| 324 | 247 | \item[\ttindexbold{Bound} $i$] \index{variables!bound|bold}
 | 
| 8136 | 248 |   is the \textbf{bound variable} with de Bruijn index~$i$, which counts the
 | 
| 324 | 249 | number of lambdas, starting from zero, between a variable's occurrence | 
| 250 | and its binding. The representation prevents capture of variables. For | |
| 251 |   more information see de Bruijn \cite{debruijn72} or
 | |
| 6592 | 252 |   Paulson~\cite[page~376]{paulson-ml2}.
 | 
| 104 | 253 | |
| 4317 | 254 | \item[\ttindexbold{Abs} ($a$, $T$, $u$)]
 | 
| 324 | 255 |   \index{lambda abs@$\lambda$-abstractions|bold}
 | 
| 8136 | 256 |   is the $\lambda$-\textbf{abstraction} with body~$u$, and whose bound
 | 
| 324 | 257 | variable has name~$a$ and type~$T$. The name is used only for parsing | 
| 258 | and printing; it has no logical significance. | |
| 104 | 259 | |
| 324 | 260 | \item[$t$ \$ $u$] \index{$@{\tt\$}|bold} \index{function applications|bold}
 | 
| 8136 | 261 | is the \textbf{application} of~$t$ to~$u$.
 | 
| 324 | 262 | \end{ttdescription}
 | 
| 9695 | 263 | Application is written as an infix operator to aid readability. Here is an | 
| 264 | \ML\ pattern to recognize FOL formulae of the form~$A\imp B$, binding the | |
| 265 | subformulae to~$A$ and~$B$: | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 266 | \begin{ttbox}
 | 
| 104 | 267 | Const("Trueprop",_) $ (Const("op -->",_) $ A $ B)
 | 
| 268 | \end{ttbox}
 | |
| 269 | ||
| 270 | ||
| 4317 | 271 | \section{*Variable binding}
 | 
| 286 | 272 | \begin{ttbox}
 | 
| 273 | loose_bnos : term -> int list | |
| 274 | incr_boundvars : int -> term -> term | |
| 275 | abstract_over : term*term -> term | |
| 276 | variant_abs : string * typ * term -> string * term | |
| 8136 | 277 | aconv          : term * term -> bool\hfill\textbf{infix}
 | 
| 286 | 278 | \end{ttbox}
 | 
| 279 | These functions are all concerned with the de Bruijn representation of | |
| 280 | bound variables. | |
| 324 | 281 | \begin{ttdescription}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 282 | \item[\ttindexbold{loose_bnos} $t$]
 | 
| 286 | 283 | returns the list of all dangling bound variable references. In | 
| 6669 | 284 |   particular, \texttt{Bound~0} is loose unless it is enclosed in an
 | 
| 285 |   abstraction.  Similarly \texttt{Bound~1} is loose unless it is enclosed in
 | |
| 286 | 286 | at least two abstractions; if enclosed in just one, the list will contain | 
| 287 | the number 0. A well-formed term does not contain any loose variables. | |
| 288 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 289 | \item[\ttindexbold{incr_boundvars} $j$]
 | 
| 332 | 290 | increases a term's dangling bound variables by the offset~$j$. This is | 
| 286 | 291 | required when moving a subterm into a context where it is enclosed by a | 
| 292 | different number of abstractions. Bound variables with a matching | |
| 293 | abstraction are unaffected. | |
| 294 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 295 | \item[\ttindexbold{abstract_over} $(v,t)$]
 | 
| 286 | 296 | forms the abstraction of~$t$ over~$v$, which may be any well-formed term. | 
| 6669 | 297 |   It replaces every occurrence of \(v\) by a \texttt{Bound} variable with the
 | 
| 286 | 298 | correct index. | 
| 299 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 300 | \item[\ttindexbold{variant_abs} $(a,T,u)$]
 | 
| 286 | 301 | substitutes into $u$, which should be the body of an abstraction. | 
| 302 | It replaces each occurrence of the outermost bound variable by a free | |
| 303 | variable. The free variable has type~$T$ and its name is a variant | |
| 332 | 304 | of~$a$ chosen to be distinct from all constants and from all variables | 
| 286 | 305 | free in~$u$. | 
| 306 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 307 | \item[$t$ \ttindexbold{aconv} $u$]
 | 
| 286 | 308 | tests whether terms~$t$ and~$u$ are \(\alpha\)-convertible: identical up | 
| 309 | to renaming of bound variables. | |
| 310 | \begin{itemize}
 | |
| 311 | \item | |
| 6669 | 312 |     Two constants, \texttt{Free}s, or \texttt{Var}s are \(\alpha\)-convertible
 | 
| 286 | 313 | if their names and types are equal. | 
| 314 | (Variables having the same name but different types are thus distinct. | |
| 315 | This confusing situation should be avoided!) | |
| 316 | \item | |
| 317 | Two bound variables are \(\alpha\)-convertible | |
| 318 | if they have the same number. | |
| 319 | \item | |
| 320 | Two abstractions are \(\alpha\)-convertible | |
| 321 | if their bodies are, and their bound variables have the same type. | |
| 322 | \item | |
| 323 | Two applications are \(\alpha\)-convertible | |
| 324 | if the corresponding subterms are. | |
| 325 | \end{itemize}
 | |
| 326 | ||
| 324 | 327 | \end{ttdescription}
 | 
| 286 | 328 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 329 | \section{Certified terms}\index{terms!certified|bold}\index{signatures}
 | 
| 8136 | 330 | A term $t$ can be \textbf{certified} under a signature to ensure that every type
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 331 | in~$t$ is well-formed and every constant in~$t$ is a type instance of a | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 332 | constant declared in the signature. The term must be well-typed and its use | 
| 6669 | 333 | of bound variables must be well-formed.  Meta-rules such as \texttt{forall_elim}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 334 | take certified terms as arguments. | 
| 104 | 335 | |
| 324 | 336 | Certified terms belong to the abstract type \mltydx{cterm}.
 | 
| 104 | 337 | Elements of the type can only be created through the certification process. | 
| 338 | In case of error, Isabelle raises exception~\ttindex{TERM}\@.
 | |
| 339 | ||
| 340 | \subsection{Printing terms}
 | |
| 324 | 341 | \index{terms!printing of}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 342 | \begin{ttbox}
 | 
| 275 | 343 | string_of_cterm : cterm -> string | 
| 104 | 344 | Sign.string_of_term : Sign.sg -> term -> string | 
| 345 | \end{ttbox}
 | |
| 324 | 346 | \begin{ttdescription}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 347 | \item[\ttindexbold{string_of_cterm} $ct$]
 | 
| 104 | 348 | displays $ct$ as a string. | 
| 349 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 350 | \item[\ttindexbold{Sign.string_of_term} $sign$ $t$]
 | 
| 104 | 351 | displays $t$ as a string, using the syntax of~$sign$. | 
| 324 | 352 | \end{ttdescription}
 | 
| 104 | 353 | |
| 354 | \subsection{Making and inspecting certified terms}
 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 355 | \begin{ttbox}
 | 
| 8136 | 356 | cterm_of : Sign.sg -> term -> cterm | 
| 357 | read_cterm : Sign.sg -> string * typ -> cterm | |
| 358 | cert_axm : Sign.sg -> string * term -> string * term | |
| 359 | read_axm : Sign.sg -> string * string -> string * term | |
| 360 | rep_cterm      : cterm -> \{T:typ, t:term, sign:Sign.sg, maxidx:int\}
 | |
| 4543 | 361 | Sign.certify_term : Sign.sg -> term -> term * typ * int | 
| 104 | 362 | \end{ttbox}
 | 
| 324 | 363 | \begin{ttdescription}
 | 
| 4543 | 364 | |
| 365 | \item[\ttindexbold{cterm_of} $sign$ $t$] \index{signatures} certifies
 | |
| 366 | $t$ with respect to signature~$sign$. | |
| 367 | ||
| 368 | \item[\ttindexbold{read_cterm} $sign$ ($s$, $T$)] reads the string~$s$
 | |
| 369 | using the syntax of~$sign$, creating a certified term. The term is | |
| 370 | checked to have type~$T$; this type also tells the parser what kind | |
| 371 | of phrase to parse. | |
| 372 | ||
| 373 | \item[\ttindexbold{cert_axm} $sign$ ($name$, $t$)] certifies $t$ with
 | |
| 374 | respect to $sign$ as a meta-proposition and converts all exceptions | |
| 375 | to an error, including the final message | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 376 | \begin{ttbox}
 | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 377 | The error(s) above occurred in axiom "\(name\)" | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 378 | \end{ttbox}
 | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 379 | |
| 4543 | 380 | \item[\ttindexbold{read_axm} $sign$ ($name$, $s$)] similar to {\tt
 | 
| 381 | cert_axm}, but first reads the string $s$ using the syntax of | |
| 382 | $sign$. | |
| 383 | ||
| 384 | \item[\ttindexbold{rep_cterm} $ct$] decomposes $ct$ as a record
 | |
| 385 | containing its type, the term itself, its signature, and the maximum | |
| 386 | subscript of its unknowns. The type and maximum subscript are | |
| 387 | computed during certification. | |
| 388 | ||
| 389 | \item[\ttindexbold{Sign.certify_term}] is a more primitive version of
 | |
| 390 |   \texttt{cterm_of}, returning the internal representation instead of
 | |
| 391 |   an abstract \texttt{cterm}.
 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 392 | |
| 324 | 393 | \end{ttdescription}
 | 
| 104 | 394 | |
| 395 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 396 | \section{Types}\index{types|bold}
 | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 397 | Types belong to the \ML\ type \mltydx{typ}, which is a concrete datatype with
 | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 398 | three constructor functions. These correspond to type constructors, free | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 399 | type variables and schematic type variables. Types are classified by sorts, | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 400 | which are lists of classes (representing an intersection). A class is | 
| 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 401 | represented by a string. | 
| 104 | 402 | \begin{ttbox}
 | 
| 403 | type class = string; | |
| 404 | type sort = class list; | |
| 405 | ||
| 406 | datatype typ = Type of string * typ list | |
| 407 | | TFree of string * sort | |
| 408 | | TVar of indexname * sort; | |
| 409 | ||
| 410 | infixr 5 -->; | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 411 | fun S --> T = Type ("fun", [S, T]);
 | 
| 104 | 412 | \end{ttbox}
 | 
| 324 | 413 | \begin{ttdescription}
 | 
| 4317 | 414 | \item[\ttindexbold{Type} ($a$, $Ts$)] \index{type constructors|bold}
 | 
| 8136 | 415 |   applies the \textbf{type constructor} named~$a$ to the type operand list~$Ts$.
 | 
| 324 | 416 |   Type constructors include~\tydx{fun}, the binary function space
 | 
| 417 |   constructor, as well as nullary type constructors such as~\tydx{prop}.
 | |
| 418 | Other type constructors may be introduced. In expressions, but not in | |
| 419 |   patterns, \hbox{\tt$S$-->$T$} is a convenient shorthand for function
 | |
| 420 | types. | |
| 104 | 421 | |
| 4317 | 422 | \item[\ttindexbold{TFree} ($a$, $s$)] \index{type variables|bold}
 | 
| 8136 | 423 |   is the \textbf{type variable} with name~$a$ and sort~$s$.
 | 
| 104 | 424 | |
| 4317 | 425 | \item[\ttindexbold{TVar} ($v$, $s$)] \index{type unknowns|bold}
 | 
| 8136 | 426 |   is the \textbf{type unknown} with indexname~$v$ and sort~$s$.
 | 
| 324 | 427 | Type unknowns are essentially free type variables, but may be | 
| 428 | instantiated during unification. | |
| 429 | \end{ttdescription}
 | |
| 104 | 430 | |
| 431 | ||
| 432 | \section{Certified types}
 | |
| 433 | \index{types!certified|bold}
 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 434 | Certified types, which are analogous to certified terms, have type | 
| 275 | 435 | \ttindexbold{ctyp}.
 | 
| 104 | 436 | |
| 437 | \subsection{Printing types}
 | |
| 324 | 438 | \index{types!printing of}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 439 | \begin{ttbox}
 | 
| 275 | 440 | string_of_ctyp : ctyp -> string | 
| 104 | 441 | Sign.string_of_typ : Sign.sg -> typ -> string | 
| 442 | \end{ttbox}
 | |
| 324 | 443 | \begin{ttdescription}
 | 
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 444 | \item[\ttindexbold{string_of_ctyp} $cT$]
 | 
| 104 | 445 | displays $cT$ as a string. | 
| 446 | ||
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 447 | \item[\ttindexbold{Sign.string_of_typ} $sign$ $T$]
 | 
| 104 | 448 | displays $T$ as a string, using the syntax of~$sign$. | 
| 324 | 449 | \end{ttdescription}
 | 
| 104 | 450 | |
| 451 | ||
| 452 | \subsection{Making and inspecting certified types}
 | |
| 864 
d63b111b917a
quite a lot of minor and major revisions (inspecting theories, read_axm,
 wenzelm parents: 
478diff
changeset | 453 | \begin{ttbox}
 | 
| 4543 | 454 | ctyp_of : Sign.sg -> typ -> ctyp | 
| 8136 | 455 | rep_ctyp         : ctyp -> \{T: typ, sign: Sign.sg\}
 | 
| 4543 | 456 | Sign.certify_typ : Sign.sg -> typ -> typ | 
| 104 | 457 | \end{ttbox}
 | 
| 324 | 458 | \begin{ttdescription}
 | 
| 4543 | 459 | |
| 460 | \item[\ttindexbold{ctyp_of} $sign$ $T$] \index{signatures} certifies
 | |
| 461 | $T$ with respect to signature~$sign$. | |
| 462 | ||
| 463 | \item[\ttindexbold{rep_ctyp} $cT$] decomposes $cT$ as a record
 | |
| 464 | containing the type itself and its signature. | |
| 465 | ||
| 466 | \item[\ttindexbold{Sign.certify_typ}] is a more primitive version of
 | |
| 467 |   \texttt{ctyp_of}, returning the internal representation instead of
 | |
| 468 |   an abstract \texttt{ctyp}.
 | |
| 104 | 469 | |
| 324 | 470 | \end{ttdescription}
 | 
| 104 | 471 | |
| 1846 | 472 | |
| 104 | 473 | \index{theories|)}
 | 
| 5369 | 474 | |
| 475 | ||
| 476 | %%% Local Variables: | |
| 477 | %%% mode: latex | |
| 478 | %%% TeX-master: "ref" | |
| 479 | %%% End: |