| author | avigad | 
| Tue, 12 Jul 2005 17:56:03 +0200 | |
| changeset 16775 | c1b87ef4a1c3 | 
| parent 16168 | adb83939177f | 
| child 17043 | d3e52c3bfb07 | 
| permissions | -rw-r--r-- | 
| 7135 | 1 | |
| 13048 | 2 | \chapter{Generic tools and packages}\label{ch:gen-tools}
 | 
| 7167 | 3 | |
| 12621 | 4 | \section{Theory specification commands}
 | 
| 12618 | 5 | |
| 6 | \subsection{Axiomatic type classes}\label{sec:axclass}
 | |
| 7167 | 7 | |
| 8517 | 8 | \indexisarcmd{axclass}\indexisarcmd{instance}\indexisarmeth{intro-classes}
 | 
| 7167 | 9 | \begin{matharray}{rcl}
 | 
| 8517 | 10 |   \isarcmd{axclass} & : & \isartrans{theory}{theory} \\
 | 
| 11 |   \isarcmd{instance} & : & \isartrans{theory}{proof(prove)} \\
 | |
| 12 | intro_classes & : & \isarmeth \\ | |
| 7167 | 13 | \end{matharray}
 | 
| 14 | ||
| 8517 | 15 | Axiomatic type classes are provided by Isabelle/Pure as a \emph{definitional}
 | 
| 16 | interface to type classes (cf.~\S\ref{sec:classes}).  Thus any object logic
 | |
| 8547 | 17 | may make use of this light-weight mechanism of abstract theories | 
| 8901 | 18 | \cite{Wenzel:1997:TPHOL}.  There is also a tutorial on using axiomatic type
 | 
| 13024 | 19 | classes in Isabelle \cite{isabelle-axclass} that is part of the standard
 | 
| 8901 | 20 | Isabelle documentation. | 
| 8517 | 21 | |
| 7167 | 22 | \begin{rail}
 | 
| 12879 | 23 | 'axclass' classdecl (axmdecl prop +) | 
| 8517 | 24 | ; | 
| 14605 
9de4d64eee3b
'instance' and intro_classes now handle general sorts;
 wenzelm parents: 
14212diff
changeset | 25 |   'instance' (nameref ('<' | subseteq) nameref | nameref '::' arity)
 | 
| 7167 | 26 | ; | 
| 27 | \end{rail}
 | |
| 28 | ||
| 29 | \begin{descr}
 | |
| 13041 | 30 | |
| 13024 | 31 | \item [$\AXCLASS~c \subseteq \vec c~~axms$] defines an axiomatic type class as | 
| 11100 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
11095diff
changeset | 32 | the intersection of existing classes, with additional axioms holding. Class | 
| 10223 | 33 | axioms may not contain more than one type variable. The class axioms (with | 
| 34 | implicit sort constraints added) are bound to the given names. Furthermore | |
| 12976 | 35 |   a class introduction rule is generated (being bound as $c{.}intro$); this
 | 
| 36 | rule is employed by method $intro_classes$ to support instantiation proofs | |
| 37 | of this class. | |
| 13041 | 38 | |
| 12976 | 39 | The ``axioms'' are stored as theorems according to the given name | 
| 13039 | 40 | specifications, adding the class name $c$ as name space prefix; the same | 
| 41 |   facts are also stored collectively as $c{\dtt}axioms$.
 | |
| 14605 
9de4d64eee3b
'instance' and intro_classes now handle general sorts;
 wenzelm parents: 
14212diff
changeset | 42 | |
| 
9de4d64eee3b
'instance' and intro_classes now handle general sorts;
 wenzelm parents: 
14212diff
changeset | 43 | \item [$\INSTANCE~c@1 \subseteq c@2$ and $\INSTANCE~t :: (\vec s)s$] setup a | 
| 11100 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
11095diff
changeset | 44 | goal stating a class relation or type arity. The proof would usually | 
| 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
11095diff
changeset | 45 | proceed by $intro_classes$, and then establish the characteristic theorems | 
| 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
11095diff
changeset | 46 | of the type classes involved. After finishing the proof, the theory will be | 
| 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
11095diff
changeset | 47 | augmented by a type signature declaration corresponding to the resulting | 
| 
34d58b1818f4
\<subseteq> syntax for classes/classrel/axclass/instance;
 wenzelm parents: 
11095diff
changeset | 48 | theorem. | 
| 13041 | 49 | |
| 8517 | 50 | \item [$intro_classes$] repeatedly expands all class introduction rules of | 
| 10858 | 51 | this theory. Note that this method usually needs not be named explicitly, | 
| 13040 | 52 | as it is already included in the default proof step (of $\PROOFNAME$ etc.). | 
| 53 | In particular, instantiation of trivial (syntactic) classes may be performed | |
| 54 | by a single ``$\DDOT$'' proof step. | |
| 13027 | 55 | |
| 7167 | 56 | \end{descr}
 | 
| 57 | ||
| 7315 | 58 | |
| 12618 | 59 | \subsection{Locales and local contexts}\label{sec:locale}
 | 
| 60 | ||
| 13040 | 61 | Locales are named local contexts, consisting of a list of declaration elements | 
| 13041 | 62 | that are modeled after the Isar proof context commands (cf.\ | 
| 13040 | 63 | \S\ref{sec:proof-context}).
 | 
| 12976 | 64 | |
| 13048 | 65 | |
| 12976 | 66 | \subsubsection{Localized commands}
 | 
| 12618 | 67 | |
| 12976 | 68 | Existing locales may be augmented later on by adding new facts. Note that the | 
| 69 | actual context definition may not be changed! Several theory commands that | |
| 70 | produce facts in some way are available in ``localized'' versions, referring | |
| 71 | to a named locale instead of the global theory context. | |
| 12967 | 72 | |
| 12976 | 73 | \indexouternonterm{locale}
 | 
| 12967 | 74 | \begin{rail}
 | 
| 75 |   locale: '(' 'in' name ')'
 | |
| 76 | ; | |
| 12976 | 77 | \end{rail}
 | 
| 12967 | 78 | |
| 12976 | 79 | Emerging facts of localized commands are stored in two versions, both in the | 
| 80 | target locale and the theory (after export). The latter view produces a | |
| 81 | qualified binding, using the locale name as a name space prefix. | |
| 82 | ||
| 83 | For example, ``$\LEMMAS~(\IN~loc)~a = \vec b$'' retrieves facts $\vec b$ from | |
| 84 | the locale context of $loc$ and augments its body by an appropriate | |
| 85 | ``$\isarkeyword{notes}$'' element (see below).  The exported view of $a$,
 | |
| 86 | after discharging the locale context, is stored as $loc{.}a$ within the global
 | |
| 13041 | 87 | theory. A localized goal ``$\LEMMANAME~(\IN~loc)~a:~\phi$'' works similarly, | 
| 88 | only that the fact emerges through the subsequent proof, which may refer to | |
| 89 | the full infrastructure of the locale context (covering local parameters with | |
| 90 | typing and concrete syntax, assumptions, definitions etc.). Most notably, | |
| 13411 | 91 | fact declarations of the locale are active during the proof as well (e.g.\ | 
| 13041 | 92 | local $simp$ rules). | 
| 12976 | 93 | |
| 13411 | 94 | As a general principle, results exported from a locale context acquire | 
| 95 | additional premises according to the specification. Usually this is only a | |
| 96 | single predicate according to the standard ``closed'' view of locale | |
| 97 | specifications. | |
| 98 | ||
| 12976 | 99 | |
| 100 | \subsubsection{Locale specifications}
 | |
| 101 | ||
| 102 | \indexisarcmd{locale}\indexisarcmd{print-locale}\indexisarcmd{print-locales}
 | |
| 103 | \begin{matharray}{rcl}
 | |
| 104 |   \isarcmd{locale} & : & \isarkeep{theory} \\
 | |
| 105 |   \isarcmd{print_locale}^* & : & \isarkeep{theory~|~proof} \\
 | |
| 106 |   \isarcmd{print_locales}^* & : & \isarkeep{theory~|~proof} \\
 | |
| 107 | \end{matharray}
 | |
| 108 | ||
| 109 | \indexouternonterm{contextexpr}\indexouternonterm{contextelem}
 | |
| 110 | ||
| 111 | \railalias{printlocale}{print\_locale}
 | |
| 112 | \railterm{printlocale}
 | |
| 113 | ||
| 114 | \begin{rail}
 | |
| 13411 | 115 |   'locale' ('(open)')? name ('=' localeexpr)?
 | 
| 12976 | 116 | ; | 
| 117 | printlocale localeexpr | |
| 118 | ; | |
| 119 | localeexpr: ((contextexpr '+' (contextelem+)) | contextexpr | (contextelem+)) | |
| 120 | ; | |
| 121 | ||
| 122 |   contextexpr: nameref | '(' contextexpr ')' |
 | |
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
16010diff
changeset | 123 | (contextexpr (name mixfix? +)) | (contextexpr + '+') | 
| 12976 | 124 | ; | 
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 125 | contextelem: fixes | constrains | assumes | defines | notes | includes | 
| 12976 | 126 | ; | 
| 127 |   fixes: 'fixes' (name ('::' type)? structmixfix? + 'and')
 | |
| 128 | ; | |
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 129 | constrains: 'constrains' (name '::' type + 'and') | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 130 | ; | 
| 12976 | 131 | assumes: 'assumes' (thmdecl? props + 'and') | 
| 132 | ; | |
| 133 | defines: 'defines' (thmdecl? prop proppat? + 'and') | |
| 134 | ; | |
| 135 | notes: 'notes' (thmdef? thmrefs + 'and') | |
| 136 | ; | |
| 137 | includes: 'includes' contextexpr | |
| 138 | ; | |
| 12967 | 139 | \end{rail}
 | 
| 12618 | 140 | |
| 12976 | 141 | \begin{descr}
 | 
| 13411 | 142 | |
| 143 | \item [$\LOCALE~loc~=~import~+~body$] defines a new locale $loc$ as a context | |
| 12976 | 144 | consisting of a certain view of existing locales ($import$) plus some | 
| 145 | additional elements ($body$). Both $import$ and $body$ are optional; the | |
| 13024 | 146 | degenerate form $\LOCALE~loc$ defines an empty locale, which may still be | 
| 147 | useful to collect declarations of facts later on. Type-inference on locale | |
| 12976 | 148 | expressions automatically takes care of the most general typing that the | 
| 149 | combined context elements may acquire. | |
| 13041 | 150 | |
| 12976 | 151 | The $import$ consists of a structured context expression, consisting of | 
| 152 | references to existing locales, renamed contexts, or merged contexts. | |
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
16010diff
changeset | 153 | Renaming uses positional notation: $c~\vec x$ means that (a prefix of) the | 
| 12976 | 154 | fixed parameters of context $c$ are named according to $\vec x$; a | 
| 16102 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
16010diff
changeset | 155 |   ``\texttt{_}'' (underscore) \indexisarthm{_@\texttt{_}} means to skip that
 | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
16010diff
changeset | 156 | position. Renaming by default deletes existing syntax. Optionally, | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
16010diff
changeset | 157 | new syntax may by specified with a mixfix annotation. Note that the | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
16010diff
changeset | 158 | special syntax declared with ``$(structure)$'' (see below) is | 
| 
c5f6726d9bb1
Locale expressions: rename with optional mixfix syntax.
 ballarin parents: 
16010diff
changeset | 159 | neither deleted nor can it be changed. | 
| 13041 | 160 | Merging proceeds from left-to-right, suppressing any duplicates stemming | 
| 161 | from different paths through the import hierarchy. | |
| 162 | ||
| 12976 | 163 | The $body$ consists of basic context elements, further context expressions | 
| 164 | may be included as well. | |
| 165 | ||
| 166 |   \begin{descr}
 | |
| 13041 | 167 | |
| 12976 | 168 |   \item [$\FIXES{~x::\tau~(mx)}$] declares a local parameter of type $\tau$
 | 
| 169 | and mixfix annotation $mx$ (both are optional). The special syntax | |
| 13027 | 170 | declaration ``$(structure)$'' means that $x$ may be referenced | 
| 171 | implicitly in this context. | |
| 13041 | 172 | |
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 173 |   \item [$\CONSTRAINS{~x::\tau}$] introduces a type constraint $\tau$
 | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 174 | on the local parameter $x$. | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 175 | |
| 12976 | 176 |   \item [$\ASSUMES{a}{\vec\phi}$] introduces local premises, similar to
 | 
| 177 |     $\ASSUMENAME$ within a proof (cf.\ \S\ref{sec:proof-context}).
 | |
| 13041 | 178 | |
| 12976 | 179 |   \item [$\DEFINES{a}{x \equiv t}$] defines a previously declared parameter.
 | 
| 13041 | 180 | This is close to $\DEFNAME$ within a proof (cf.\ | 
| 12976 | 181 |     \S\ref{sec:proof-context}), but $\DEFINESNAME$ takes an equational
 | 
| 13041 | 182 | proposition instead of variable-term pair. The left-hand side of the | 
| 183 |     equation may have additional arguments, e.g.\ ``$\DEFINES{}{f~\vec x
 | |
| 184 | \equiv t}$''. | |
| 185 | ||
| 12976 | 186 |   \item [$\NOTES{a}{\vec b}$] reconsiders facts within a local context.  Most
 | 
| 187 | notably, this may include arbitrary declarations in any attribute | |
| 188 | specifications included here, e.g.\ a local $simp$ rule. | |
| 13041 | 189 | |
| 12976 | 190 |   \item [$\INCLUDES{c}$] copies the specified context in a statically scoped
 | 
| 15763 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 191 |     manner.  Only available in the long goal format of \S\ref{sec:goals}.
 | 
| 13041 | 192 | |
| 12976 | 193 | In contrast, the initial $import$ specification of a locale expression | 
| 194 | maintains a dynamic relation to the locales being referenced (benefiting | |
| 195 | from any later fact declarations in the obvious manner). | |
| 196 |   \end{descr}
 | |
| 13411 | 197 | |
| 13041 | 198 |   Note that ``$\IS{p}$'' patterns given in the syntax of $\ASSUMESNAME$ and
 | 
| 13411 | 199 | $\DEFINESNAME$ above are illegal in locale definitions. In the long goal | 
| 200 |   format of \S\ref{sec:goals}, term bindings may be included as expected,
 | |
| 201 | though. | |
| 202 | ||
| 203 | \medskip By default, locale specifications are ``closed up'' by turning the | |
| 204 | given text into a predicate definition $loc_axioms$ and deriving the | |
| 205 | original assumptions as local lemmas (modulo local definitions). The | |
| 206 | predicate statement covers only the newly specified assumptions, omitting | |
| 207 | the content of included locale expressions. The full cumulative view is | |
| 208 | only provided on export, involving another predicate $loc$ that refers to | |
| 209 | the complete specification text. | |
| 210 | ||
| 211 | In any case, the predicate arguments are those locale parameters that | |
| 212 | actually occur in the respective piece of text. Also note that these | |
| 213 | predicates operate at the meta-level in theory, but the locale packages | |
| 214 | attempts to internalize statements according to the object-logic setup | |
| 215 | (e.g.\ replacing $\Forall$ by $\forall$, and $\Imp$ by $\imp$ in HOL; see | |
| 216 |   also \S\ref{sec:object-logic}).  Separate introduction rules
 | |
| 217 | $loc_axioms.intro$ and $loc.intro$ are declared as well. | |
| 218 | ||
| 219 | The $(open)$ option of a locale specification prevents both the current | |
| 220 | $loc_axioms$ and cumulative $loc$ predicate constructions. Predicates are | |
| 221 | also omitted for empty specification texts. | |
| 12976 | 222 | |
| 223 | \item [$\isarkeyword{print_locale}~import~+~body$] prints the specified locale
 | |
| 224 | expression in a flattened form. The notable special case | |
| 225 |   $\isarkeyword{print_locale}~loc$ just prints the contents of the named
 | |
| 226 | locale, but keep in mind that type-inference will normalize type variables | |
| 227 | according to the usual alphabetical order. | |
| 13041 | 228 | |
| 12976 | 229 | \item [$\isarkeyword{print_locales}$] prints the names of all locales of the
 | 
| 230 | current theory. | |
| 231 | ||
| 232 | \end{descr}
 | |
| 233 | ||
| 12621 | 234 | |
| 15763 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 235 | \subsubsection{Interpretation of locales}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 236 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 237 | Locale expressions (more precisely, \emph{context expressions}) may be
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 238 | instantiated, and the instantiated facts added to the current context. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 239 | This requires a proof of the instantiated specification and is called | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 240 | \emph{locale interpretation}.  Interpretation is possible in theories
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 241 | ($\isarcmd{interpretation}$) and proof contexts
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 242 | ($\isarcmd{interpret}$).
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 243 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 244 | \indexisarcmd{interpretation}\indexisarcmd{interpret}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 245 | \indexisarcmd{print-interps}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 246 | \begin{matharray}{rcl}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 247 |   \isarcmd{interpretation} & : & \isartrans{theory}{proof(prove)} \\
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 248 |   \isarcmd{interpret} & : & \isartrans{proof(state) ~|~ proof(chain)}{proof(prove)} \\
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 249 |   \isarcmd{print_interps}^* & : &  \isarkeep{theory~|~proof} \\
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 250 | \end{matharray}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 251 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 252 | \indexouternonterm{interp}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 253 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 254 | \railalias{printinterps}{print\_interps}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 255 | \railterm{printinterps}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 256 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 257 | \begin{rail}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 258 | 'interpretation' interp | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 259 | ; | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 260 | 'interpret' interp | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 261 | ; | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 262 | printinterps name | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 263 | ; | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 264 |   interp: thmdecl? contextexpr ('[' (inst+) ']')?
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 265 | ; | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 266 | \end{rail}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 267 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 268 | \begin{descr}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 269 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 270 | \item [$\isarcmd{interpretation}~expr~insts$]
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 271 | interprets $expr$ in the theory. The instantiation is positional. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 272 | All parameters must receive an instantiation term --- with the | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 273 | exception of defined parameters. These are, if omitted, derived | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 274 | from the defining equation and other instantiations. Use ``\_'' to | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 275 | omit an instantiation term. Free variables are automatically | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 276 | generalized. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 277 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 278 | The context expression may be preceded by a name and/or attributes. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 279 | These take effect in the post-processing of facts. The name is used | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 280 | to prefix fact names, for example to avoid accidental hiding of | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 281 | other facts. Attributes are applied after attributes of the | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 282 | interpreted facts. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 283 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 284 | The command is aware of interpretations already active in the | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 285 | theory. No proof obligations are generated for those, neither is | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 286 | post-processing applied to their facts. This avoids duplication of | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 287 | interpreted facts, in particular. Note that, in the case of a | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 288 | locale with import, parts of the interpretation may already be | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 289 | active. The command will only generate proof obligations and add | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 290 | facts for new parts. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 291 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 292 | Adding facts to locales has the | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 293 | effect of adding interpreted facts to the theory for all active | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 294 | interpretations also. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 295 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 296 | \item [$\isarcmd{interpret}~expr~insts$]
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 297 | interprets $expr$ in the proof context and is otherwise similar to | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 298 | the previous command. Free variables in instantiations are not | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 299 | generalized, however. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 300 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 301 | \item [$\isarcmd{print_interps}~loc$]
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 302 | prints the interpretations of a particular locale $loc$ that are | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 303 | active in the current context, either theory or proof context. | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 304 | |
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 305 | \end{descr}
 | 
| 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 306 | |
| 15837 | 307 | \begin{warn}
 | 
| 308 | Since attributes are applied to interpreted theorems, interpretation | |
| 309 | may modify the current simpset and claset. Take this into | |
| 310 | account when choosing attributes for local theorems. | |
| 311 | \end{warn}
 | |
| 312 | ||
| 16168 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 313 | \begin{warn}
 | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 314 | An interpretation may subsume previous interpretations. A warning | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 315 | is issued, since it is likely that these could have been generalized | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 316 | in the first place. The locale package does not attempt to remove | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 317 | subsumed interpretations. This situation is normally harmless, but | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 318 | note that $blast$ gets confused by the presence of multiple axclass | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 319 | instances a rule. | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 320 | \end{warn}
 | 
| 
adb83939177f
Locales: new element constrains, parameter renaming with syntax,
 ballarin parents: 
16102diff
changeset | 321 | |
| 15763 
b901a127ac73
Interpretation supports statically scoped attributes; documentation.
 ballarin parents: 
14605diff
changeset | 322 | |
| 12621 | 323 | \section{Derived proof schemes}
 | 
| 324 | ||
| 325 | \subsection{Generalized elimination}\label{sec:obtain}
 | |
| 326 | ||
| 327 | \indexisarcmd{obtain}
 | |
| 328 | \begin{matharray}{rcl}
 | |
| 329 |   \isarcmd{obtain} & : & \isartrans{proof(state)}{proof(prove)} \\
 | |
| 330 | \end{matharray}
 | |
| 331 | ||
| 332 | Generalized elimination means that additional elements with certain properties | |
| 13041 | 333 | may be introduced in the current context, by virtue of a locally proven | 
| 12621 | 334 | ``soundness statement''. Technically speaking, the $\OBTAINNAME$ language | 
| 335 | element is like a declaration of $\FIXNAME$ and $\ASSUMENAME$ (see also see | |
| 336 | \S\ref{sec:proof-context}), together with a soundness proof of its additional
 | |
| 337 | claim. According to the nature of existential reasoning, assumptions get | |
| 338 | eliminated from any result exported from the context later, provided that the | |
| 339 | corresponding parameters do \emph{not} occur in the conclusion.
 | |
| 340 | ||
| 341 | \begin{rail}
 | |
| 12879 | 342 | 'obtain' (vars + 'and') 'where' (props + 'and') | 
| 12621 | 343 | ; | 
| 344 | \end{rail}
 | |
| 12618 | 345 | |
| 12621 | 346 | $\OBTAINNAME$ is defined as a derived Isar command as follows, where $\vec b$ | 
| 347 | shall refer to (optional) facts indicated for forward chaining. | |
| 348 | \begin{matharray}{l}
 | |
| 349 | \langle facts~\vec b\rangle \\ | |
| 350 |   \OBTAIN{\vec x}{a}{\vec \phi}~~\langle proof\rangle \equiv {} \\[1ex]
 | |
| 13041 | 351 |   \quad \HAVE{}{\All{thesis} (\All{\vec x} \vec\phi \Imp thesis) \Imp thesis} \\
 | 
| 352 |   \quad \PROOF{succeed} \\
 | |
| 12621 | 353 |   \qquad \FIX{thesis} \\
 | 
| 13041 | 354 |   \qquad \ASSUME{that~[intro?]}{\All{\vec x} \vec\phi \Imp thesis} \\
 | 
| 13042 | 355 |   \qquad \THUS{}{thesis} \\
 | 
| 356 |   \quad\qquad \APPLY{-} \\
 | |
| 13041 | 357 |   \quad\qquad \USING{\vec b}~~\langle proof\rangle \\
 | 
| 358 |   \quad \QED{} \\
 | |
| 12621 | 359 |   \quad \FIX{\vec x}~\ASSUMENAME^\ast~a\colon~\vec\phi \\
 | 
| 360 | \end{matharray}
 | |
| 361 | ||
| 362 | Typically, the soundness proof is relatively straight-forward, often just by | |
| 13048 | 363 | canonical automated tools such as ``$\BY{simp}$'' or ``$\BY{blast}$''.
 | 
| 364 | Accordingly, the ``$that$'' reduction above is declared as simplification and | |
| 365 | introduction rule. | |
| 12621 | 366 | |
| 367 | \medskip | |
| 368 | ||
| 369 | In a sense, $\OBTAINNAME$ represents at the level of Isar proofs what would be | |
| 370 | meta-logical existential quantifiers and conjunctions. This concept has a | |
| 13041 | 371 | broad range of useful applications, ranging from plain elimination (or | 
| 12621 | 372 | introduction) of object-level existentials and conjunctions, to elimination | 
| 373 | over results of symbolic evaluation of recursive definitions, for example. | |
| 374 | Also note that $\OBTAINNAME$ without parameters acts much like $\HAVENAME$, | |
| 13041 | 375 | where the result is treated as a genuine assumption. | 
| 12621 | 376 | |
| 377 | ||
| 378 | \subsection{Calculational reasoning}\label{sec:calculation}
 | |
| 7315 | 379 | |
| 8619 | 380 | \indexisarcmd{also}\indexisarcmd{finally}
 | 
| 381 | \indexisarcmd{moreover}\indexisarcmd{ultimately}
 | |
| 12976 | 382 | \indexisarcmd{print-trans-rules}
 | 
| 383 | \indexisaratt{trans}\indexisaratt{sym}\indexisaratt{symmetric}
 | |
| 7315 | 384 | \begin{matharray}{rcl}
 | 
| 385 |   \isarcmd{also} & : & \isartrans{proof(state)}{proof(state)} \\
 | |
| 386 |   \isarcmd{finally} & : & \isartrans{proof(state)}{proof(chain)} \\
 | |
| 8619 | 387 |   \isarcmd{moreover} & : & \isartrans{proof(state)}{proof(state)} \\
 | 
| 388 |   \isarcmd{ultimately} & : & \isartrans{proof(state)}{proof(chain)} \\
 | |
| 10154 | 389 |   \isarcmd{print_trans_rules}^* & : & \isarkeep{theory~|~proof} \\
 | 
| 7315 | 390 | trans & : & \isaratt \\ | 
| 12976 | 391 | sym & : & \isaratt \\ | 
| 392 | symmetric & : & \isaratt \\ | |
| 7315 | 393 | \end{matharray}
 | 
| 394 | ||
| 395 | Calculational proof is forward reasoning with implicit application of | |
| 11332 | 396 | transitivity rules (such those of $=$, $\leq$, $<$). Isabelle/Isar maintains | 
| 7391 | 397 | an auxiliary register $calculation$\indexisarthm{calculation} for accumulating
 | 
| 7897 | 398 | results obtained by transitivity composed with the current result. Command | 
| 399 | $\ALSO$ updates $calculation$ involving $this$, while $\FINALLY$ exhibits the | |
| 400 | final $calculation$ by forward chaining towards the next goal statement. Both | |
| 401 | commands require valid current facts, i.e.\ may occur only after commands that | |
| 402 | produce theorems such as $\ASSUMENAME$, $\NOTENAME$, or some finished proof of | |
| 8619 | 403 | $\HAVENAME$, $\SHOWNAME$ etc. The $\MOREOVER$ and $\ULTIMATELY$ commands are | 
| 404 | similar to $\ALSO$ and $\FINALLY$, but only collect further results in | |
| 405 | $calculation$ without applying any rules yet. | |
| 7315 | 406 | |
| 13041 | 407 | Also note that the implicit term abbreviation ``$\dots$'' has its canonical | 
| 408 | application with calculational proofs. It refers to the argument of the | |
| 409 | preceding statement. (The argument of a curried infix expression happens to be | |
| 410 | its right-hand side.) | |
| 7315 | 411 | |
| 412 | Isabelle/Isar calculations are implicitly subject to block structure in the | |
| 413 | sense that new threads of calculational reasoning are commenced for any new | |
| 414 | block (as opened by a local goal, for example). This means that, apart from | |
| 415 | being able to nest calculations, there is no separate \emph{begin-calculation}
 | |
| 416 | command required. | |
| 417 | ||
| 8619 | 418 | \medskip | 
| 419 | ||
| 13041 | 420 | The Isar calculation proof commands may be defined as follows:\footnote{We
 | 
| 421 | suppress internal bookkeeping such as proper handling of block-structure.} | |
| 8619 | 422 | \begin{matharray}{rcl}
 | 
| 423 |   \ALSO@0 & \equiv & \NOTE{calculation}{this} \\
 | |
| 9606 | 424 |   \ALSO@{n+1} & \equiv & \NOTE{calculation}{trans~[OF~calculation~this]} \\[0.5ex]
 | 
| 8619 | 425 |   \FINALLY & \equiv & \ALSO~\FROM{calculation} \\
 | 
| 426 |   \MOREOVER & \equiv & \NOTE{calculation}{calculation~this} \\
 | |
| 427 |   \ULTIMATELY & \equiv & \MOREOVER~\FROM{calculation} \\
 | |
| 428 | \end{matharray}
 | |
| 429 | ||
| 7315 | 430 | \begin{rail}
 | 
| 13024 | 431 |   ('also' | 'finally') ('(' thmrefs ')')?
 | 
| 8619 | 432 | ; | 
| 8507 | 433 | 'trans' (() | 'add' | 'del') | 
| 7315 | 434 | ; | 
| 435 | \end{rail}
 | |
| 436 | ||
| 437 | \begin{descr}
 | |
| 13041 | 438 | |
| 8547 | 439 | \item [$\ALSO~(\vec a)$] maintains the auxiliary $calculation$ register as | 
| 7315 | 440 | follows. The first occurrence of $\ALSO$ in some calculational thread | 
| 7905 | 441 | initializes $calculation$ by $this$. Any subsequent $\ALSO$ on the same | 
| 7335 | 442 | level of block-structure updates $calculation$ by some transitivity rule | 
| 7458 | 443 | applied to $calculation$ and $this$ (in that order). Transitivity rules are | 
| 11095 | 444 | picked from the current context, unless alternative rules are given as | 
| 445 | explicit arguments. | |
| 9614 | 446 | |
| 8547 | 447 | \item [$\FINALLY~(\vec a)$] maintaining $calculation$ in the same way as | 
| 7315 | 448 | $\ALSO$, and concludes the current calculational thread. The final result | 
| 449 | is exhibited as fact for forward chaining towards the next goal. Basically, | |
| 7987 | 450 |   $\FINALLY$ just abbreviates $\ALSO~\FROM{calculation}$.  Note that
 | 
| 451 |   ``$\FINALLY~\SHOW{}{\Var{thesis}}~\DOT$'' and
 | |
| 452 |   ``$\FINALLY~\HAVE{}{\phi}~\DOT$'' are typical idioms for concluding
 | |
| 453 | calculational proofs. | |
| 9614 | 454 | |
| 8619 | 455 | \item [$\MOREOVER$ and $\ULTIMATELY$] are analogous to $\ALSO$ and $\FINALLY$, | 
| 456 | but collect results only, without applying rules. | |
| 13041 | 457 | |
| 13024 | 458 | \item [$\isarkeyword{print_trans_rules}$] prints the list of transitivity
 | 
| 459 | rules (for calculational commands $\ALSO$ and $\FINALLY$) and symmetry rules | |
| 460 | (for the $symmetric$ operation and single step elimination patters) of the | |
| 461 | current context. | |
| 13041 | 462 | |
| 8547 | 463 | \item [$trans$] declares theorems as transitivity rules. | 
| 13041 | 464 | |
| 13024 | 465 | \item [$sym$] declares symmetry rules. | 
| 13041 | 466 | |
| 12976 | 467 | \item [$symmetric$] resolves a theorem with some rule declared as $sym$ in the | 
| 468 |   current context.  For example, ``$\ASSUME{[symmetric]}{x = y}$'' produces a
 | |
| 469 | swapped fact derived from that assumption. | |
| 13041 | 470 | |
| 13024 | 471 | In structured proof texts it is often more appropriate to use an explicit | 
| 472 |   single-step elimination proof, such as ``$\ASSUME{}{x = y}~\HENCE{}{y =
 | |
| 13041 | 473 | x}~\DDOT$''. The very same rules known to $symmetric$ are declared as | 
| 474 | $elim?$ as well. | |
| 13027 | 475 | |
| 7315 | 476 | \end{descr}
 | 
| 477 | ||
| 478 | ||
| 13041 | 479 | \section{Proof tools}
 | 
| 8517 | 480 | |
| 12618 | 481 | \subsection{Miscellaneous methods and attributes}\label{sec:misc-meth-att}
 | 
| 8517 | 482 | |
| 9606 | 483 | \indexisarmeth{unfold}\indexisarmeth{fold}\indexisarmeth{insert}
 | 
| 8517 | 484 | \indexisarmeth{erule}\indexisarmeth{drule}\indexisarmeth{frule}
 | 
| 485 | \indexisarmeth{fail}\indexisarmeth{succeed}
 | |
| 486 | \begin{matharray}{rcl}
 | |
| 487 | unfold & : & \isarmeth \\ | |
| 10741 | 488 | fold & : & \isarmeth \\ | 
| 489 | insert & : & \isarmeth \\[0.5ex] | |
| 8517 | 490 | erule^* & : & \isarmeth \\ | 
| 491 | drule^* & : & \isarmeth \\ | |
| 13024 | 492 | frule^* & : & \isarmeth \\ | 
| 8517 | 493 | succeed & : & \isarmeth \\ | 
| 494 | fail & : & \isarmeth \\ | |
| 495 | \end{matharray}
 | |
| 7135 | 496 | |
| 497 | \begin{rail}
 | |
| 10741 | 498 |   ('fold' | 'unfold' | 'insert') thmrefs
 | 
| 499 | ; | |
| 500 |   ('erule' | 'drule' | 'frule') ('('nat')')? thmrefs
 | |
| 7135 | 501 | ; | 
| 502 | \end{rail}
 | |
| 503 | ||
| 7167 | 504 | \begin{descr}
 | 
| 13041 | 505 | |
| 13024 | 506 | \item [$unfold~\vec a$ and $fold~\vec a$] expand (or fold back again) the | 
| 507 | given meta-level definitions throughout all goals; any chained facts | |
| 508 | provided are inserted into the goal and subject to rewriting as well. | |
| 13041 | 509 | |
| 10741 | 510 | \item [$insert~\vec a$] inserts theorems as facts into all goals of the proof | 
| 511 | state. Note that current facts indicated for forward chaining are ignored. | |
| 13024 | 512 | |
| 8547 | 513 | \item [$erule~\vec a$, $drule~\vec a$, and $frule~\vec a$] are similar to the | 
| 514 |   basic $rule$ method (see \S\ref{sec:pure-meth-att}), but apply rules by
 | |
| 8517 | 515 | elim-resolution, destruct-resolution, and forward-resolution, respectively | 
| 10741 | 516 |   \cite{isabelle-ref}.  The optional natural number argument (default $0$)
 | 
| 13041 | 517 | specifies additional assumption steps to be performed here. | 
| 518 | ||
| 10741 | 519 | Note that these methods are improper ones, mainly serving for | 
| 520 | experimentation and tactic script emulation. Different modes of basic rule | |
| 521 | application are usually expressed in Isar at the proof language level, | |
| 522 | rather than via implicit proof state manipulations. For example, a proper | |
| 13041 | 523 | single-step elimination would be done using the plain $rule$ method, with | 
| 10741 | 524 | forward chaining of current facts. | 
| 13024 | 525 | |
| 8517 | 526 | \item [$succeed$] yields a single (unchanged) result; it is the identity of | 
| 527 |   the ``\texttt{,}'' method combinator (cf.\ \S\ref{sec:syn-meth}).
 | |
| 13024 | 528 | |
| 8517 | 529 | \item [$fail$] yields an empty result sequence; it is the identity of the | 
| 530 |   ``\texttt{|}'' method combinator (cf.\ \S\ref{sec:syn-meth}).
 | |
| 13024 | 531 | |
| 7167 | 532 | \end{descr}
 | 
| 7135 | 533 | |
| 10318 | 534 | \indexisaratt{tagged}\indexisaratt{untagged}
 | 
| 9614 | 535 | \indexisaratt{THEN}\indexisaratt{COMP}
 | 
| 14175 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 536 | \indexisaratt{unfolded}\indexisaratt{folded}
 | 
| 13027 | 537 | \indexisaratt{standard}\indexisarattof{Pure}{elim-format}
 | 
| 13024 | 538 | \indexisaratt{no-vars}
 | 
| 8517 | 539 | \begin{matharray}{rcl}
 | 
| 9905 | 540 | tagged & : & \isaratt \\ | 
| 541 | untagged & : & \isaratt \\[0.5ex] | |
| 9614 | 542 | THEN & : & \isaratt \\ | 
| 8517 | 543 | COMP & : & \isaratt \\[0.5ex] | 
| 9905 | 544 | unfolded & : & \isaratt \\ | 
| 545 | folded & : & \isaratt \\[0.5ex] | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9936diff
changeset | 546 | elim_format & : & \isaratt \\ | 
| 13041 | 547 | standard^* & : & \isaratt \\ | 
| 9936 | 548 | no_vars^* & : & \isaratt \\ | 
| 8517 | 549 | \end{matharray}
 | 
| 550 | ||
| 551 | \begin{rail}
 | |
| 9905 | 552 | 'tagged' (nameref+) | 
| 8517 | 553 | ; | 
| 9905 | 554 | 'untagged' name | 
| 8517 | 555 | ; | 
| 10154 | 556 |   ('THEN' | 'COMP') ('[' nat ']')? thmref
 | 
| 8517 | 557 | ; | 
| 9905 | 558 |   ('unfolded' | 'folded') thmrefs
 | 
| 8517 | 559 | ; | 
| 560 | \end{rail}
 | |
| 561 | ||
| 562 | \begin{descr}
 | |
| 13041 | 563 | |
| 9905 | 564 | \item [$tagged~name~args$ and $untagged~name$] add and remove $tags$ of some | 
| 8517 | 565 | theorem. Tags may be any list of strings that serve as comment for some | 
| 566 | tools (e.g.\ $\LEMMANAME$ causes the tag ``$lemma$'' to be added to the | |
| 567 | result). The first string is considered the tag name, the rest its | |
| 568 | arguments. Note that untag removes any tags of the same name. | |
| 13041 | 569 | |
| 570 | \item [$THEN~a$ and $COMP~a$] compose rules by resolution. $THEN$ resolves | |
| 571 | with the first premise of $a$ (an alternative position may be also | |
| 572 | specified); the $COMP$ version skips the automatic lifting process that is | |
| 573 |   normally intended (cf.\ \texttt{RS} and \texttt{COMP} in
 | |
| 8547 | 574 |   \cite[\S5]{isabelle-ref}).
 | 
| 13041 | 575 | |
| 9905 | 576 | \item [$unfolded~\vec a$ and $folded~\vec a$] expand and fold back again the | 
| 577 | given meta-level definitions throughout a rule. | |
| 13041 | 578 | |
| 13027 | 579 | \item [$elim_format$] turns a destruction rule into elimination rule format, | 
| 580 | by resolving with the rule $\PROP A \Imp (\PROP A \Imp \PROP B) \Imp \PROP | |
| 581 | B$. | |
| 13048 | 582 | |
| 583 |   Note that the Classical Reasoner (\S\ref{sec:classical}) provides its own
 | |
| 584 | version of this operation. | |
| 13041 | 585 | |
| 586 | \item [$standard$] puts a theorem into the standard form of object-rules at | |
| 587 | the outermost theory level. Note that this operation violates the local | |
| 588 | proof context (including active locales). | |
| 589 | ||
| 9232 | 590 | \item [$no_vars$] replaces schematic variables by free ones; this is mainly | 
| 591 | for tuning output of pretty printed theorems. | |
| 13027 | 592 | |
| 8517 | 593 | \end{descr}
 | 
| 7135 | 594 | |
| 595 | ||
| 12621 | 596 | \subsection{Further tactic emulations}\label{sec:tactics}
 | 
| 9606 | 597 | |
| 598 | The following improper proof methods emulate traditional tactics. These admit | |
| 599 | direct access to the goal state, which is normally considered harmful! In | |
| 600 | particular, this may involve both numbered goal addressing (default 1), and | |
| 601 | dynamic instantiation within the scope of some subgoal. | |
| 602 | ||
| 603 | \begin{warn}
 | |
| 14175 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 604 | Dynamic instantiations refer to universally quantified parameters of | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 605 | a subgoal (the dynamic context) rather than fixed variables and term | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 606 | abbreviations of a (static) Isar context. | 
| 9606 | 607 | \end{warn}
 | 
| 608 | ||
| 14175 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 609 | Tactic emulation methods, unlike their ML counterparts, admit | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 610 | simultaneous instantiation from both dynamic and static contexts. If | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 611 | names occur in both contexts goal parameters hide locally fixed | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 612 | variables. Likewise, schematic variables refer to term abbreviations, | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 613 | if present in the static context. Otherwise the schematic variable is | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 614 | interpreted as a schematic variable and left to be solved by unification | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 615 | with certain parts of the subgoal. | 
| 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 616 | |
| 9606 | 617 | Note that the tactic emulation proof methods in Isabelle/Isar are consistently | 
| 14175 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 618 | named $foo_tac$. Note also that variable names occurring on left hand sides | 
| 14212 | 619 | of instantiations must be preceded by a question mark if they coincide with | 
| 620 | a keyword or contain dots. | |
| 14175 
dbd16ebaf907
Method rule_tac understands Isar contexts: documentation.
 ballarin parents: 
13622diff
changeset | 621 | This is consistent with the attribute $where$ (see \S\ref{sec:pure-meth-att}).
 | 
| 9606 | 622 | |
| 623 | \indexisarmeth{rule-tac}\indexisarmeth{erule-tac}
 | |
| 624 | \indexisarmeth{drule-tac}\indexisarmeth{frule-tac}
 | |
| 625 | \indexisarmeth{cut-tac}\indexisarmeth{thin-tac}
 | |
| 9642 | 626 | \indexisarmeth{subgoal-tac}\indexisarmeth{rename-tac}
 | 
| 9614 | 627 | \indexisarmeth{rotate-tac}\indexisarmeth{tactic}
 | 
| 9606 | 628 | \begin{matharray}{rcl}
 | 
| 629 | rule_tac^* & : & \isarmeth \\ | |
| 630 | erule_tac^* & : & \isarmeth \\ | |
| 631 | drule_tac^* & : & \isarmeth \\ | |
| 632 | frule_tac^* & : & \isarmeth \\ | |
| 633 | cut_tac^* & : & \isarmeth \\ | |
| 634 | thin_tac^* & : & \isarmeth \\ | |
| 635 | subgoal_tac^* & : & \isarmeth \\ | |
| 9614 | 636 | rename_tac^* & : & \isarmeth \\ | 
| 637 | rotate_tac^* & : & \isarmeth \\ | |
| 9606 | 638 | tactic^* & : & \isarmeth \\ | 
| 639 | \end{matharray}
 | |
| 640 | ||
| 641 | \railalias{ruletac}{rule\_tac}
 | |
| 642 | \railterm{ruletac}
 | |
| 643 | ||
| 644 | \railalias{eruletac}{erule\_tac}
 | |
| 645 | \railterm{eruletac}
 | |
| 646 | ||
| 647 | \railalias{druletac}{drule\_tac}
 | |
| 648 | \railterm{druletac}
 | |
| 649 | ||
| 650 | \railalias{fruletac}{frule\_tac}
 | |
| 651 | \railterm{fruletac}
 | |
| 652 | ||
| 653 | \railalias{cuttac}{cut\_tac}
 | |
| 654 | \railterm{cuttac}
 | |
| 655 | ||
| 656 | \railalias{thintac}{thin\_tac}
 | |
| 657 | \railterm{thintac}
 | |
| 658 | ||
| 659 | \railalias{subgoaltac}{subgoal\_tac}
 | |
| 660 | \railterm{subgoaltac}
 | |
| 661 | ||
| 9614 | 662 | \railalias{renametac}{rename\_tac}
 | 
| 663 | \railterm{renametac}
 | |
| 664 | ||
| 665 | \railalias{rotatetac}{rotate\_tac}
 | |
| 666 | \railterm{rotatetac}
 | |
| 667 | ||
| 9606 | 668 | \begin{rail}
 | 
| 669 | ( ruletac | eruletac | druletac | fruletac | cuttac | thintac ) goalspec? | |
| 670 | ( insts thmref | thmrefs ) | |
| 671 | ; | |
| 672 | subgoaltac goalspec? (prop +) | |
| 673 | ; | |
| 9614 | 674 | renametac goalspec? (name +) | 
| 675 | ; | |
| 676 | rotatetac goalspec? int? | |
| 677 | ; | |
| 9606 | 678 | 'tactic' text | 
| 679 | ; | |
| 680 | ||
| 681 | insts: ((name '=' term) + 'and') 'in' | |
| 682 | ; | |
| 683 | \end{rail}
 | |
| 684 | ||
| 685 | \begin{descr}
 | |
| 13041 | 686 | |
| 9606 | 687 | \item [$rule_tac$ etc.] do resolution of rules with explicit instantiation. | 
| 688 |   This works the same way as the ML tactics \texttt{res_inst_tac} etc. (see
 | |
| 689 |   \cite[\S3]{isabelle-ref}).
 | |
| 13041 | 690 | |
| 691 | Multiple rules may be only given if there is no instantiation; then | |
| 9606 | 692 |   $rule_tac$ is the same as \texttt{resolve_tac} in ML (see
 | 
| 693 |   \cite[\S3]{isabelle-ref}).
 | |
| 13041 | 694 | |
| 9606 | 695 | \item [$cut_tac$] inserts facts into the proof state as assumption of a | 
| 696 |   subgoal, see also \texttt{cut_facts_tac} in \cite[\S3]{isabelle-ref}.  Note
 | |
| 13027 | 697 | that the scope of schematic variables is spread over the main goal | 
| 698 | statement. Instantiations may be given as well, see also ML tactic | |
| 9606 | 699 |   \texttt{cut_inst_tac} in \cite[\S3]{isabelle-ref}.
 | 
| 13041 | 700 | |
| 9606 | 701 | \item [$thin_tac~\phi$] deletes the specified assumption from a subgoal; note | 
| 702 |   that $\phi$ may contain schematic variables.  See also \texttt{thin_tac} in
 | |
| 703 |   \cite[\S3]{isabelle-ref}.
 | |
| 13041 | 704 | |
| 9606 | 705 | \item [$subgoal_tac~\phi$] adds $\phi$ as an assumption to a subgoal. See | 
| 706 |   also \texttt{subgoal_tac} and \texttt{subgoals_tac} in
 | |
| 707 |   \cite[\S3]{isabelle-ref}.
 | |
| 13041 | 708 | |
| 9614 | 709 | \item [$rename_tac~\vec x$] renames parameters of a goal according to the list | 
| 710 |   $\vec x$, which refers to the \emph{suffix} of variables.
 | |
| 13041 | 711 | |
| 9614 | 712 | \item [$rotate_tac~n$] rotates the assumptions of a goal by $n$ positions: | 
| 713 | from right to left if $n$ is positive, and from left to right if $n$ is | |
| 714 |   negative; the default value is $1$.  See also \texttt{rotate_tac} in
 | |
| 715 |   \cite[\S3]{isabelle-ref}.
 | |
| 13041 | 716 | |
| 9606 | 717 | \item [$tactic~text$] produces a proof method from any ML text of type | 
| 718 |   \texttt{tactic}.  Apart from the usual ML environment and the current
 | |
| 719 | implicit theory context, the ML code may refer to the following locally | |
| 720 | bound values: | |
| 721 | ||
| 722 | {\footnotesize\begin{verbatim}
 | |
| 723 | val ctxt : Proof.context | |
| 724 | val facts : thm list | |
| 725 | val thm : string -> thm | |
| 726 | val thms : string -> thm list | |
| 727 | \end{verbatim}}
 | |
| 728 |   Here \texttt{ctxt} refers to the current proof context, \texttt{facts}
 | |
| 729 | indicates any current facts for forward-chaining, and | |
| 730 |   \texttt{thm}~/~\texttt{thms} retrieve named facts (including global
 | |
| 731 | theorems) from the context. | |
| 732 | \end{descr}
 | |
| 733 | ||
| 734 | ||
| 12621 | 735 | \subsection{The Simplifier}\label{sec:simplifier}
 | 
| 736 | ||
| 13048 | 737 | \subsubsection{Simplification methods}
 | 
| 12618 | 738 | |
| 8483 | 739 | \indexisarmeth{simp}\indexisarmeth{simp-all}
 | 
| 7315 | 740 | \begin{matharray}{rcl}
 | 
| 741 | simp & : & \isarmeth \\ | |
| 8483 | 742 | simp_all & : & \isarmeth \\ | 
| 7315 | 743 | \end{matharray}
 | 
| 744 | ||
| 8483 | 745 | \railalias{simpall}{simp\_all}
 | 
| 746 | \railterm{simpall}
 | |
| 747 | ||
| 8704 | 748 | \railalias{noasm}{no\_asm}
 | 
| 749 | \railterm{noasm}
 | |
| 750 | ||
| 751 | \railalias{noasmsimp}{no\_asm\_simp}
 | |
| 752 | \railterm{noasmsimp}
 | |
| 753 | ||
| 754 | \railalias{noasmuse}{no\_asm\_use}
 | |
| 755 | \railterm{noasmuse}
 | |
| 756 | ||
| 13617 | 757 | \railalias{asmlr}{asm\_lr}
 | 
| 758 | \railterm{asmlr}
 | |
| 759 | ||
| 11128 | 760 | \indexouternonterm{simpmod}
 | 
| 7315 | 761 | \begin{rail}
 | 
| 13027 | 762 |   ('simp' | simpall) ('!' ?) opt? (simpmod *)
 | 
| 7315 | 763 | ; | 
| 764 | ||
| 13617 | 765 |   opt: '(' (noasm | noasmsimp | noasmuse | asmlr) ')'
 | 
| 8704 | 766 | ; | 
| 9711 | 767 |   simpmod: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') |
 | 
| 9847 | 768 | 'split' (() | 'add' | 'del')) ':' thmrefs | 
| 7315 | 769 | ; | 
| 770 | \end{rail}
 | |
| 771 | ||
| 7321 | 772 | \begin{descr}
 | 
| 13015 | 773 | |
| 8547 | 774 | \item [$simp$] invokes Isabelle's simplifier, after declaring additional rules | 
| 8594 | 775 |   according to the arguments given.  Note that the \railtterm{only} modifier
 | 
| 8547 | 776 | first removes all other rewrite rules, congruences, and looper tactics | 
| 8594 | 777 |   (including splits), and then behaves like \railtterm{add}.
 | 
| 13041 | 778 | |
| 9711 | 779 |   \medskip The \railtterm{cong} modifiers add or delete Simplifier congruence
 | 
| 780 |   rules (see also \cite{isabelle-ref}), the default is to add.
 | |
| 13041 | 781 | |
| 9711 | 782 |   \medskip The \railtterm{split} modifiers add or delete rules for the
 | 
| 783 |   Splitter (see also \cite{isabelle-ref}), the default is to add.  This works
 | |
| 784 | only if the Simplifier method has been properly setup to include the | |
| 785 | Splitter (all major object logics such HOL, HOLCF, FOL, ZF do this already). | |
| 13041 | 786 | |
| 13015 | 787 | \item [$simp_all$] is similar to $simp$, but acts on all goals (backwards from | 
| 788 | the last to the first one). | |
| 789 | ||
| 7321 | 790 | \end{descr}
 | 
| 791 | ||
| 13015 | 792 | By default the Simplifier methods take local assumptions fully into account, | 
| 793 | using equational assumptions in the subsequent normalization process, or | |
| 13024 | 794 | simplifying assumptions themselves (cf.\ \texttt{asm_full_simp_tac} in
 | 
| 13015 | 795 | \cite[\S10]{isabelle-ref}).  In structured proofs this is usually quite well
 | 
| 796 | behaved in practice: just the local premises of the actual goal are involved, | |
| 13041 | 797 | additional facts may be inserted via explicit forward-chaining (using $\THEN$, | 
| 13015 | 798 | $\FROMNAME$ etc.). The full context of assumptions is only included if the | 
| 799 | ``$!$'' (bang) argument is given, which should be used with some care, though. | |
| 7321 | 800 | |
| 13015 | 801 | Additional Simplifier options may be specified to tune the behavior further | 
| 13041 | 802 | (mostly for unstructured scripts with many accidental local facts): | 
| 803 | ``$(no_asm)$'' means assumptions are ignored completely (cf.\ | |
| 804 | \texttt{simp_tac}), ``$(no_asm_simp)$'' means assumptions are used in the
 | |
| 805 | simplification of the conclusion but are not themselves simplified (cf.\ | |
| 806 | \texttt{asm_simp_tac}), and ``$(no_asm_use)$'' means assumptions are
 | |
| 807 | simplified but are not used in the simplification of each other or the | |
| 808 | conclusion (cf.\ \texttt{full_simp_tac}).
 | |
| 13617 | 809 | For compatibility reasons, there is also an option ``$(asm_lr)$'', | 
| 810 | which means that an assumption is only used for simplifying assumptions | |
| 811 | which are to the right of it (cf.\ \texttt{asm_lr_simp_tac}).
 | |
| 8704 | 812 | |
| 813 | \medskip | |
| 814 | ||
| 815 | The Splitter package is usually configured to work as part of the Simplifier. | |
| 9711 | 816 | The effect of repeatedly applying \texttt{split_tac} can be simulated by
 | 
| 13041 | 817 | ``$(simp~only\colon~split\colon~\vec a)$''. There is also a separate $split$ | 
| 818 | method available for single-step case splitting. | |
| 8483 | 819 | |
| 820 | ||
| 12621 | 821 | \subsubsection{Declaring rules}
 | 
| 8483 | 822 | |
| 8667 | 823 | \indexisarcmd{print-simpset}
 | 
| 8638 | 824 | \indexisaratt{simp}\indexisaratt{split}\indexisaratt{cong}
 | 
| 7321 | 825 | \begin{matharray}{rcl}
 | 
| 13024 | 826 |   \isarcmd{print_simpset}^* & : & \isarkeep{theory~|~proof} \\
 | 
| 7321 | 827 | simp & : & \isaratt \\ | 
| 9711 | 828 | cong & : & \isaratt \\ | 
| 8483 | 829 | split & : & \isaratt \\ | 
| 7321 | 830 | \end{matharray}
 | 
| 831 | ||
| 832 | \begin{rail}
 | |
| 9711 | 833 |   ('simp' | 'cong' | 'split') (() | 'add' | 'del')
 | 
| 7321 | 834 | ; | 
| 835 | \end{rail}
 | |
| 836 | ||
| 837 | \begin{descr}
 | |
| 13024 | 838 | |
| 839 | \item [$\isarcmd{print_simpset}$] prints the collection of rules declared to
 | |
| 840 | the Simplifier, which is also known as ``simpset'' internally | |
| 8667 | 841 |   \cite{isabelle-ref}.  This is a diagnostic command; $undo$ does not apply.
 | 
| 13024 | 842 | |
| 8547 | 843 | \item [$simp$] declares simplification rules. | 
| 13024 | 844 | |
| 8638 | 845 | \item [$cong$] declares congruence rules. | 
| 13024 | 846 | |
| 9711 | 847 | \item [$split$] declares case split rules. | 
| 13024 | 848 | |
| 7321 | 849 | \end{descr}
 | 
| 7319 | 850 | |
| 7315 | 851 | |
| 12621 | 852 | \subsubsection{Forward simplification}
 | 
| 853 | ||
| 9905 | 854 | \indexisaratt{simplified}
 | 
| 7315 | 855 | \begin{matharray}{rcl}
 | 
| 9905 | 856 | simplified & : & \isaratt \\ | 
| 7315 | 857 | \end{matharray}
 | 
| 858 | ||
| 9905 | 859 | \begin{rail}
 | 
| 13015 | 860 | 'simplified' opt? thmrefs? | 
| 9905 | 861 | ; | 
| 862 | ||
| 863 |   opt: '(' (noasm | noasmsimp | noasmuse) ')'
 | |
| 864 | ; | |
| 865 | \end{rail}
 | |
| 7905 | 866 | |
| 9905 | 867 | \begin{descr}
 | 
| 13048 | 868 | |
| 13015 | 869 | \item [$simplified~\vec a$] causes a theorem to be simplified, either by | 
| 870 | exactly the specified rules $\vec a$, or the implicit Simplifier context if | |
| 871 | no arguments are given. The result is fully simplified by default, | |
| 872 | including assumptions and conclusion; the options $no_asm$ etc.\ tune the | |
| 13048 | 873 | Simplifier in the same way as the for the $simp$ method. | 
| 13041 | 874 | |
| 13015 | 875 | Note that forward simplification restricts the simplifier to its most basic | 
| 876 |   operation of term rewriting; solver and looper tactics \cite{isabelle-ref}
 | |
| 877 |   are \emph{not} involved here.  The $simplified$ attribute should be only
 | |
| 878 | rarely required under normal circumstances. | |
| 879 | ||
| 9905 | 880 | \end{descr}
 | 
| 7315 | 881 | |
| 882 | ||
| 13048 | 883 | \subsubsection{Low-level equational reasoning}
 | 
| 9614 | 884 | |
| 12976 | 885 | \indexisarmeth{subst}\indexisarmeth{hypsubst}\indexisarmeth{split}
 | 
| 9614 | 886 | \begin{matharray}{rcl}
 | 
| 13015 | 887 | subst^* & : & \isarmeth \\ | 
| 9614 | 888 | hypsubst^* & : & \isarmeth \\ | 
| 13015 | 889 | split^* & : & \isarmeth \\ | 
| 9614 | 890 | \end{matharray}
 | 
| 891 | ||
| 892 | \begin{rail}
 | |
| 15995 | 893 |   'subst' ('(' 'asm' ')')? ('(' (nat+) ')')? thmref
 | 
| 9614 | 894 | ; | 
| 9799 | 895 |   'split' ('(' 'asm' ')')? thmrefs
 | 
| 9703 | 896 | ; | 
| 9614 | 897 | \end{rail}
 | 
| 898 | ||
| 13015 | 899 | These methods provide low-level facilities for equational reasoning that are | 
| 900 | intended for specialized applications only. Normally, single step | |
| 901 | calculations would be performed in a structured text (see also | |
| 902 | \S\ref{sec:calculation}), while the Simplifier methods provide the canonical
 | |
| 903 | way for automated normalization (see \S\ref{sec:simplifier}).
 | |
| 9614 | 904 | |
| 905 | \begin{descr}
 | |
| 13041 | 906 | |
| 15995 | 907 | \item [$subst~eq$] performs a single substitution step using rule $eq$, which | 
| 13041 | 908 | may be either a meta or object equality. | 
| 909 | ||
| 15995 | 910 | \item [$subst~(asm)~eq$] substitutes in an assumption. | 
| 911 | ||
| 912 | \item [$subst~(i \dots j)~eq$] performs several substitutions in the | |
| 913 | conclusion. The numbers $i$ to $j$ indicate the positions to substitute at. | |
| 914 | Positions are ordered from the top of the term tree moving down from left to | |
| 915 | right. For example, in $(a+b)+(c+d)$ there are three positions where | |
| 916 | commutativity of $+$ is applicable: 1 refers to the whole term, 2 to $a+b$ | |
| 917 | and 3 to $c+d$. If the positions in the list $(i \dots j)$ are | |
| 918 | non-overlapping (e.g. $(2~3)$ in $(a+b)+(c+d)$) you may assume all | |
| 919 | substitutions are performed simultaneously. Otherwise the behaviour of | |
| 920 | $subst$ is not specified. | |
| 921 | ||
| 922 | \item [$subst~(asm)~(i \dots j)~eq$] performs the substitutions in the | |
| 16010 | 923 | assumptions. Positions $1 \dots i@1$ refer | 
| 924 | to assumption 1, positions $i@1+1 \dots i@2$ to assumption 2, and so on. | |
| 15995 | 925 | |
| 13041 | 926 | \item [$hypsubst$] performs substitution using some assumption; this only | 
| 927 | works for equations of the form $x = t$ where $x$ is a free or bound | |
| 928 | variable. | |
| 929 | ||
| 930 | \item [$split~\vec a$] performs single-step case splitting using rules $thms$. | |
| 9799 | 931 | By default, splitting is performed in the conclusion of a goal; the $asm$ | 
| 932 | option indicates to operate on assumptions instead. | |
| 13048 | 933 | |
| 9703 | 934 | Note that the $simp$ method already involves repeated application of split | 
| 13048 | 935 | rules as declared in the current context. | 
| 9614 | 936 | \end{descr}
 | 
| 937 | ||
| 938 | ||
| 12621 | 939 | \subsection{The Classical Reasoner}\label{sec:classical}
 | 
| 7135 | 940 | |
| 13048 | 941 | \subsubsection{Basic methods}
 | 
| 7321 | 942 | |
| 13024 | 943 | \indexisarmeth{rule}\indexisarmeth{default}\indexisarmeth{contradiction}
 | 
| 944 | \indexisarmeth{intro}\indexisarmeth{elim}
 | |
| 7321 | 945 | \begin{matharray}{rcl}
 | 
| 946 | rule & : & \isarmeth \\ | |
| 13024 | 947 | contradiction & : & \isarmeth \\ | 
| 7321 | 948 | intro & : & \isarmeth \\ | 
| 949 | elim & : & \isarmeth \\ | |
| 950 | \end{matharray}
 | |
| 951 | ||
| 952 | \begin{rail}
 | |
| 8547 | 953 |   ('rule' | 'intro' | 'elim') thmrefs?
 | 
| 7321 | 954 | ; | 
| 955 | \end{rail}
 | |
| 956 | ||
| 957 | \begin{descr}
 | |
| 13041 | 958 | |
| 7466 | 959 | \item [$rule$] as offered by the classical reasoner is a refinement over the | 
| 13024 | 960 |   primitive one (see \S\ref{sec:pure-meth-att}).  Both versions essentially
 | 
| 961 | work the same, but the classical version observes the classical rule context | |
| 13041 | 962 | in addition to that of Isabelle/Pure. | 
| 963 | ||
| 964 | Common object logics (HOL, ZF, etc.) declare a rich collection of classical | |
| 965 | rules (even if these would qualify as intuitionistic ones), but only few | |
| 966 | declarations to the rule context of Isabelle/Pure | |
| 967 |   (\S\ref{sec:pure-meth-att}).
 | |
| 968 | ||
| 13024 | 969 | \item [$contradiction$] solves some goal by contradiction, deriving any result | 
| 13041 | 970 | from both $\neg A$ and $A$. Chained facts, which are guaranteed to | 
| 971 | participate, may appear in either order. | |
| 9614 | 972 | |
| 7466 | 973 | \item [$intro$ and $elim$] repeatedly refine some goal by intro- or | 
| 13041 | 974 | elim-resolution, after having inserted any chained facts. Exactly the rules | 
| 975 | given as arguments are taken into account; this allows fine-tuned | |
| 976 | decomposition of a proof problem, in contrast to common automated tools. | |
| 977 | ||
| 7321 | 978 | \end{descr}
 | 
| 979 | ||
| 980 | ||
| 13048 | 981 | \subsubsection{Automated methods}
 | 
| 7315 | 982 | |
| 9799 | 983 | \indexisarmeth{blast}\indexisarmeth{fast}\indexisarmeth{slow}
 | 
| 984 | \indexisarmeth{best}\indexisarmeth{safe}\indexisarmeth{clarify}
 | |
| 7321 | 985 | \begin{matharray}{rcl}
 | 
| 9780 | 986 | blast & : & \isarmeth \\ | 
| 987 | fast & : & \isarmeth \\ | |
| 9799 | 988 | slow & : & \isarmeth \\ | 
| 9780 | 989 | best & : & \isarmeth \\ | 
| 990 | safe & : & \isarmeth \\ | |
| 991 | clarify & : & \isarmeth \\ | |
| 7321 | 992 | \end{matharray}
 | 
| 993 | ||
| 11128 | 994 | \indexouternonterm{clamod}
 | 
| 7321 | 995 | \begin{rail}
 | 
| 13027 | 996 |   'blast' ('!' ?) nat? (clamod *)
 | 
| 7321 | 997 | ; | 
| 13027 | 998 |   ('fast' | 'slow' | 'best' | 'safe' | 'clarify') ('!' ?) (clamod *)
 | 
| 7321 | 999 | ; | 
| 1000 | ||
| 9408 | 1001 |   clamod: (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' thmrefs
 | 
| 7321 | 1002 | ; | 
| 1003 | \end{rail}
 | |
| 1004 | ||
| 1005 | \begin{descr}
 | |
| 1006 | \item [$blast$] refers to the classical tableau prover (see \texttt{blast_tac}
 | |
| 7335 | 1007 |   in \cite[\S11]{isabelle-ref}).  The optional argument specifies a
 | 
| 10858 | 1008 | user-supplied search bound (default 20). | 
| 9799 | 1009 | \item [$fast$, $slow$, $best$, $safe$, and $clarify$] refer to the generic | 
| 1010 |   classical reasoner.  See \texttt{fast_tac}, \texttt{slow_tac},
 | |
| 1011 |   \texttt{best_tac}, \texttt{safe_tac}, and \texttt{clarify_tac} in
 | |
| 1012 |   \cite[\S11]{isabelle-ref} for more information.
 | |
| 7321 | 1013 | \end{descr}
 | 
| 1014 | ||
| 13041 | 1015 | Any of the above methods support additional modifiers of the context of | 
| 1016 | classical rules. Their semantics is analogous to the attributes given before. | |
| 1017 | Facts provided by forward chaining are inserted into the goal before | |
| 1018 | commencing proof search. The ``!''~argument causes the full context of | |
| 1019 | assumptions to be included as well. | |
| 7321 | 1020 | |
| 7315 | 1021 | |
| 12621 | 1022 | \subsubsection{Combined automated methods}\label{sec:clasimp}
 | 
| 7315 | 1023 | |
| 9799 | 1024 | \indexisarmeth{auto}\indexisarmeth{force}\indexisarmeth{clarsimp}
 | 
| 1025 | \indexisarmeth{fastsimp}\indexisarmeth{slowsimp}\indexisarmeth{bestsimp}
 | |
| 7321 | 1026 | \begin{matharray}{rcl}
 | 
| 9606 | 1027 | auto & : & \isarmeth \\ | 
| 7321 | 1028 | force & : & \isarmeth \\ | 
| 9438 | 1029 | clarsimp & : & \isarmeth \\ | 
| 9606 | 1030 | fastsimp & : & \isarmeth \\ | 
| 9799 | 1031 | slowsimp & : & \isarmeth \\ | 
| 1032 | bestsimp & : & \isarmeth \\ | |
| 7321 | 1033 | \end{matharray}
 | 
| 1034 | ||
| 11128 | 1035 | \indexouternonterm{clasimpmod}
 | 
| 7321 | 1036 | \begin{rail}
 | 
| 13027 | 1037 | 'auto' '!'? (nat nat)? (clasimpmod *) | 
| 9780 | 1038 | ; | 
| 13027 | 1039 |   ('force' | 'clarsimp' | 'fastsimp' | 'slowsimp' | 'bestsimp') '!'? (clasimpmod *)
 | 
| 7321 | 1040 | ; | 
| 7315 | 1041 | |
| 9711 | 1042 |   clasimpmod: ('simp' (() | 'add' | 'del' | 'only') |
 | 
| 10031 | 1043 |     ('cong' | 'split') (() | 'add' | 'del') |
 | 
| 1044 | 'iff' (((() | 'add') '?'?) | 'del') | | |
| 9408 | 1045 |     (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' thmrefs
 | 
| 7321 | 1046 | \end{rail}
 | 
| 7315 | 1047 | |
| 7321 | 1048 | \begin{descr}
 | 
| 9799 | 1049 | \item [$auto$, $force$, $clarsimp$, $fastsimp$, $slowsimp$, and $bestsimp$] | 
| 1050 | provide access to Isabelle's combined simplification and classical reasoning | |
| 1051 |   tactics.  These correspond to \texttt{auto_tac}, \texttt{force_tac},
 | |
| 1052 |   \texttt{clarsimp_tac}, and Classical Reasoner tactics with the Simplifier
 | |
| 1053 |   added as wrapper, see \cite[\S11]{isabelle-ref} for more information.  The
 | |
| 13048 | 1054 |   modifier arguments correspond to those given in \S\ref{sec:simplifier} and
 | 
| 1055 |   \S\ref{sec:classical}.  Just note that the ones related to the Simplifier
 | |
| 1056 |   are prefixed by \railtterm{simp} here.
 | |
| 9614 | 1057 | |
| 7987 | 1058 | Facts provided by forward chaining are inserted into the goal before doing | 
| 1059 | the search. The ``!''~argument causes the full context of assumptions to be | |
| 1060 | included as well. | |
| 7321 | 1061 | \end{descr}
 | 
| 1062 | ||
| 7987 | 1063 | |
| 13048 | 1064 | \subsubsection{Declaring rules}
 | 
| 7135 | 1065 | |
| 8667 | 1066 | \indexisarcmd{print-claset}
 | 
| 7391 | 1067 | \indexisaratt{intro}\indexisaratt{elim}\indexisaratt{dest}
 | 
| 9936 | 1068 | \indexisaratt{iff}\indexisaratt{rule}
 | 
| 7321 | 1069 | \begin{matharray}{rcl}
 | 
| 13024 | 1070 |   \isarcmd{print_claset}^* & : & \isarkeep{theory~|~proof} \\
 | 
| 7321 | 1071 | intro & : & \isaratt \\ | 
| 1072 | elim & : & \isaratt \\ | |
| 1073 | dest & : & \isaratt \\ | |
| 9936 | 1074 | rule & : & \isaratt \\ | 
| 7391 | 1075 | iff & : & \isaratt \\ | 
| 7321 | 1076 | \end{matharray}
 | 
| 7135 | 1077 | |
| 7321 | 1078 | \begin{rail}
 | 
| 9408 | 1079 |   ('intro' | 'elim' | 'dest') ('!' | () | '?')
 | 
| 7321 | 1080 | ; | 
| 9936 | 1081 | 'rule' 'del' | 
| 1082 | ; | |
| 10031 | 1083 | 'iff' (((() | 'add') '?'?) | 'del') | 
| 9936 | 1084 | ; | 
| 7321 | 1085 | \end{rail}
 | 
| 7135 | 1086 | |
| 7321 | 1087 | \begin{descr}
 | 
| 13024 | 1088 | |
| 1089 | \item [$\isarcmd{print_claset}$] prints the collection of rules declared to
 | |
| 1090 | the Classical Reasoner, which is also known as ``simpset'' internally | |
| 8667 | 1091 |   \cite{isabelle-ref}.  This is a diagnostic command; $undo$ does not apply.
 | 
| 13024 | 1092 | |
| 8517 | 1093 | \item [$intro$, $elim$, and $dest$] declare introduction, elimination, and | 
| 11332 | 1094 | destruction rules, respectively. By default, rules are considered as | 
| 9408 | 1095 |   \emph{unsafe} (i.e.\ not applied blindly without backtracking), while a
 | 
| 13041 | 1096 |   single ``!'' classifies as \emph{safe}.  Rule declarations marked by ``?''
 | 
| 1097 |   coincide with those of Isabelle/Pure, cf.\ \S\ref{sec:pure-meth-att} (i.e.\
 | |
| 1098 | are only applied in single steps of the $rule$ method). | |
| 13024 | 1099 | |
| 11332 | 1100 | \item [$rule~del$] deletes introduction, elimination, or destruction rules from | 
| 9936 | 1101 | the context. | 
| 13041 | 1102 | |
| 1103 | \item [$iff$] declares logical equivalences to the Simplifier and the | |
| 13024 | 1104 | Classical reasoner at the same time. Non-conditional rules result in a | 
| 1105 | ``safe'' introduction and elimination pair; conditional ones are considered | |
| 1106 | ``unsafe''. Rules with negative conclusion are automatically inverted | |
| 13041 | 1107 | (using $\neg$ elimination internally). | 
| 1108 | ||
| 1109 | The ``?'' version of $iff$ declares rules to the Isabelle/Pure context only, | |
| 1110 | and omits the Simplifier declaration. | |
| 1111 | ||
| 7321 | 1112 | \end{descr}
 | 
| 7135 | 1113 | |
| 8203 
2fcc6017cb72
intro/elim/dest attributes: changed ! / !! flags to ? / ??;
 wenzelm parents: 
8195diff
changeset | 1114 | |
| 13048 | 1115 | \subsubsection{Classical operations}
 | 
| 13027 | 1116 | |
| 1117 | \indexisaratt{elim-format}\indexisaratt{swapped}
 | |
| 1118 | ||
| 1119 | \begin{matharray}{rcl}
 | |
| 1120 | elim_format & : & \isaratt \\ | |
| 1121 | swapped & : & \isaratt \\ | |
| 1122 | \end{matharray}
 | |
| 1123 | ||
| 1124 | \begin{descr}
 | |
| 13041 | 1125 | |
| 13027 | 1126 | \item [$elim_format$] turns a destruction rule into elimination rule format; | 
| 1127 | this operation is similar to the the intuitionistic version | |
| 1128 |   (\S\ref{sec:misc-meth-att}), but each premise of the resulting rule acquires
 | |
| 13041 | 1129 | an additional local fact of the negated main thesis; according to the | 
| 13027 | 1130 | classical principle $(\neg A \Imp A) \Imp A$. | 
| 13041 | 1131 | |
| 13027 | 1132 | \item [$swapped$] turns an introduction rule into an elimination, by resolving | 
| 1133 | with the classical swap principle $(\neg B \Imp A) \Imp (\neg A \Imp B)$. | |
| 1134 | ||
| 1135 | \end{descr}
 | |
| 1136 | ||
| 1137 | ||
| 12621 | 1138 | \subsection{Proof by cases and induction}\label{sec:cases-induct}
 | 
| 12618 | 1139 | |
| 13048 | 1140 | \subsubsection{Rule contexts}
 | 
| 12618 | 1141 | |
| 1142 | \indexisarcmd{case}\indexisarcmd{print-cases}
 | |
| 1143 | \indexisaratt{case-names}\indexisaratt{params}\indexisaratt{consumes}
 | |
| 1144 | \begin{matharray}{rcl}
 | |
| 1145 |   \isarcmd{case} & : & \isartrans{proof(state)}{proof(state)} \\
 | |
| 1146 |   \isarcmd{print_cases}^* & : & \isarkeep{proof} \\
 | |
| 1147 | case_names & : & \isaratt \\ | |
| 1148 | params & : & \isaratt \\ | |
| 1149 | consumes & : & \isaratt \\ | |
| 1150 | \end{matharray}
 | |
| 1151 | ||
| 1152 | Basically, Isar proof contexts are built up explicitly using commands like | |
| 1153 | $\FIXNAME$, $\ASSUMENAME$ etc.\ (see \S\ref{sec:proof-context}).  In typical
 | |
| 1154 | verification tasks this can become hard to manage, though. In particular, a | |
| 1155 | large number of local contexts may emerge from case analysis or induction over | |
| 1156 | inductive sets and types. | |
| 1157 | ||
| 1158 | \medskip | |
| 1159 | ||
| 1160 | The $\CASENAME$ command provides a shorthand to refer to certain parts of | |
| 1161 | logical context symbolically. Proof methods may provide an environment of | |
| 1162 | named ``cases'' of the form $c\colon \vec x, \vec \phi$. Then the effect of | |
| 13041 | 1163 | ``$\CASE{c}$'' is that of ``$\FIX{\vec x}~\ASSUME{c}{\vec\phi}$''.  Term
 | 
| 1164 | bindings may be covered as well, such as $\Var{case}$ for the intended
 | |
| 1165 | conclusion. | |
| 12618 | 1166 | |
| 13027 | 1167 | Normally the ``terminology'' of a case value (i.e.\ the parameters $\vec x$) | 
| 13041 | 1168 | are marked as hidden.  Using the explicit form ``$\CASE{(c~\vec x)}$'' enables
 | 
| 1169 | proof writers to choose their own names for the subsequent proof text. | |
| 12618 | 1170 | |
| 1171 | \medskip | |
| 1172 | ||
| 13027 | 1173 | It is important to note that $\CASENAME$ does \emph{not} provide direct means
 | 
| 1174 | to peek at the current goal state, which is generally considered | |
| 1175 | non-observable in Isar. The text of the cases basically emerge from standard | |
| 1176 | elimination or induction rules, which in turn are derived from previous theory | |
| 13041 | 1177 | specifications in a canonical way (say from $\isarkeyword{inductive}$
 | 
| 1178 | definitions). | |
| 13027 | 1179 | |
| 12618 | 1180 | Named cases may be exhibited in the current proof context only if both the | 
| 1181 | proof method and the rules involved support this. Case names and parameters | |
| 1182 | of basic rules may be declared by hand as well, by using appropriate | |
| 1183 | attributes. Thus variant versions of rules that have been derived manually | |
| 1184 | may be used in advanced case analysis later. | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1185 | |
| 12618 | 1186 | \railalias{casenames}{case\_names}
 | 
| 1187 | \railterm{casenames}
 | |
| 1188 | ||
| 1189 | \begin{rail}
 | |
| 13041 | 1190 |   'case' (caseref | '(' caseref ((name | underscore) +) ')')
 | 
| 12618 | 1191 | ; | 
| 13024 | 1192 | caseref: nameref attributes? | 
| 1193 | ; | |
| 1194 | ||
| 13027 | 1195 | casenames (name +) | 
| 12618 | 1196 | ; | 
| 13027 | 1197 | 'params' ((name *) + 'and') | 
| 12618 | 1198 | ; | 
| 1199 | 'consumes' nat? | |
| 1200 | ; | |
| 1201 | \end{rail}
 | |
| 1202 | ||
| 1203 | \begin{descr}
 | |
| 13041 | 1204 | |
| 1205 | \item [$\CASE{(c~\vec x)}$] invokes a named local context $c\colon \vec x,
 | |
| 1206 | \vec \phi$, as provided by an appropriate proof method (such as $cases$ and | |
| 1207 |   $induct$, see \S\ref{sec:cases-induct-meth}).  The command ``$\CASE{(c~\vec
 | |
| 1208 |     x)}$'' abbreviates ``$\FIX{\vec x}~\ASSUME{c}{\vec\phi}$''.
 | |
| 1209 | ||
| 12618 | 1210 | \item [$\isarkeyword{print_cases}$] prints all local contexts of the current
 | 
| 1211 | state, using Isar proof language notation. This is a diagnostic command; | |
| 1212 | $undo$ does not apply. | |
| 13041 | 1213 | |
| 12618 | 1214 | \item [$case_names~\vec c$] declares names for the local contexts of premises | 
| 1215 |   of some theorem; $\vec c$ refers to the \emph{suffix} of the list of
 | |
| 1216 | premises. | |
| 13041 | 1217 | |
| 12618 | 1218 | \item [$params~\vec p@1 \dots \vec p@n$] renames the innermost parameters of | 
| 1219 | premises $1, \dots, n$ of some theorem. An empty list of names may be given | |
| 1220 | to skip positions, leaving the present parameters unchanged. | |
| 13041 | 1221 | |
| 12618 | 1222 |   Note that the default usage of case rules does \emph{not} directly expose
 | 
| 1223 |   parameters to the proof context (see also \S\ref{sec:cases-induct-meth}).
 | |
| 13041 | 1224 | |
| 12618 | 1225 | \item [$consumes~n$] declares the number of ``major premises'' of a rule, | 
| 1226 | i.e.\ the number of facts to be consumed when it is applied by an | |
| 1227 |   appropriate proof method (cf.\ \S\ref{sec:cases-induct-meth}).  The default
 | |
| 1228 | value of $consumes$ is $n = 1$, which is appropriate for the usual kind of | |
| 13041 | 1229 | cases and induction rules for inductive sets (cf.\ | 
| 12618 | 1230 |   \S\ref{sec:hol-inductive}).  Rules without any $consumes$ declaration given
 | 
| 1231 | are treated as if $consumes~0$ had been specified. | |
| 13041 | 1232 | |
| 12618 | 1233 | Note that explicit $consumes$ declarations are only rarely needed; this is | 
| 1234 | already taken care of automatically by the higher-level $cases$ and $induct$ | |
| 1235 |   declarations, see also \S\ref{sec:cases-induct-att}.
 | |
| 13027 | 1236 | |
| 12618 | 1237 | \end{descr}
 | 
| 1238 | ||
| 1239 | ||
| 12621 | 1240 | \subsubsection{Proof methods}\label{sec:cases-induct-meth}
 | 
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1241 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1242 | \indexisarmeth{cases}\indexisarmeth{induct}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1243 | \begin{matharray}{rcl}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1244 | cases & : & \isarmeth \\ | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1245 | induct & : & \isarmeth \\ | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1246 | \end{matharray}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1247 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1248 | The $cases$ and $induct$ methods provide a uniform interface to case analysis | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1249 | and induction over datatypes, inductive sets, and recursive functions. The | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1250 | corresponding rules may be specified and instantiated in a casual manner. | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1251 | Furthermore, these methods provide named local contexts that may be invoked | 
| 13048 | 1252 | via the $\CASENAME$ proof command within the subsequent proof text. This | 
| 1253 | accommodates compact proof texts even when reasoning about large | |
| 1254 | specifications. | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1255 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1256 | \begin{rail}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1257 | 'cases' spec | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1258 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1259 | 'induct' spec | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1260 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1261 | |
| 13041 | 1262 | spec: open? args rule? | 
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1263 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1264 |   open: '(' 'open' ')'
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1265 | ; | 
| 13041 | 1266 | args: (insts * 'and') | 
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1267 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1268 |   rule: ('type' | 'set') ':' nameref | 'rule' ':' thmref
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1269 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1270 | \end{rail}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1271 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1272 | \begin{descr}
 | 
| 13041 | 1273 | |
| 1274 | \item [$cases~insts~R$] applies method $rule$ with an appropriate case | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1275 | distinction theorem, instantiated to the subjects $insts$. Symbolic case | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1276 | names are bound according to the rule's local contexts. | 
| 13041 | 1277 | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1278 | The rule is determined as follows, according to the facts and arguments | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1279 | passed to the $cases$ method: | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1280 |   \begin{matharray}{llll}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1281 |     \Text{facts}    &       & \Text{arguments} & \Text{rule} \\\hline
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1282 |                     & cases &           & \Text{classical case split} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1283 |                     & cases & t         & \Text{datatype exhaustion (type of $t$)} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1284 |     \edrv a \in A   & cases & \dots     & \Text{inductive set elimination (of $A$)} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1285 |     \dots           & cases & \dots ~ R & \Text{explicit rule $R$} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1286 |   \end{matharray}
 | 
| 13041 | 1287 | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1288 |   Several instantiations may be given, referring to the \emph{suffix} of
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1289 |   premises of the case rule; within each premise, the \emph{prefix} of
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1290 | variables is instantiated. In most situations, only a single term needs to | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1291 | be specified; this refers to the first variable of the last premise (it is | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1292 | usually the same for all cases). | 
| 13041 | 1293 | |
| 1294 | The ``$(open)$'' option causes the parameters of the new local contexts to | |
| 1295 | be exposed to the current proof context. Thus local variables stemming from | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1296 | distant parts of the theory development may be introduced in an implicit | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1297 | manner, which can be quite confusing to the reader. Furthermore, this | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1298 | option may cause unwanted hiding of existing local variables, resulting in | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1299 | less robust proof texts. | 
| 13041 | 1300 | |
| 1301 | \item [$induct~insts~R$] is analogous to the $cases$ method, but refers to | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1302 | induction rules, which are determined as follows: | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1303 |   \begin{matharray}{llll}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1304 |     \Text{facts}    &        & \Text{arguments} & \Text{rule} \\\hline
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1305 |                     & induct & P ~ x ~ \dots & \Text{datatype induction (type of $x$)} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1306 |     \edrv x \in A   & induct & \dots         & \Text{set induction (of $A$)} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1307 |     \dots           & induct & \dots ~ R     & \Text{explicit rule $R$} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1308 |   \end{matharray}
 | 
| 13041 | 1309 | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1310 | Several instantiations may be given, each referring to some part of a mutual | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1311 | inductive definition or datatype --- only related partial induction rules | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1312 | may be used together, though. Any of the lists of terms $P, x, \dots$ | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1313 |   refers to the \emph{suffix} of variables present in the induction rule.
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1314 | This enables the writer to specify only induction variables, or both | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1315 | predicates and variables, for example. | 
| 13041 | 1316 | |
| 1317 | The ``$(open)$'' option works the same way as for $cases$. | |
| 13027 | 1318 | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1319 | \end{descr}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1320 | |
| 13048 | 1321 | Above methods produce named local contexts, as determined by the instantiated | 
| 1322 | rule as specified in the text. Beyond that, the $induct$ method guesses | |
| 1323 | further instantiations from the goal specification itself. Any persisting | |
| 1324 | unresolved schematic variables of the resulting rule will render the the | |
| 1325 | corresponding case invalid.  The term binding $\Var{case}$\indexisarvar{case}
 | |
| 1326 | for the conclusion will be provided with each case, provided that term is | |
| 1327 | fully specified. | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1328 | |
| 13048 | 1329 | The $\isarkeyword{print_cases}$ command prints all named cases present in the
 | 
| 1330 | current proof state. | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1331 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1332 | \medskip | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1333 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1334 | It is important to note that there is a fundamental difference of the $cases$ | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1335 | and $induct$ methods in handling of non-atomic goal statements: $cases$ just | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1336 | applies a certain rule in backward fashion, splitting the result into new | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1337 | goals with the local contexts being augmented in a purely monotonic manner. | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1338 | |
| 13622 | 1339 | In contrast, $induct$ passes the full goal statement through the | 
| 1340 | ``recursive'' course involved in the induction. Thus the original statement | |
| 1341 | is basically replaced by separate copies, corresponding to the induction | |
| 1342 | hypotheses and conclusion; the original goal context is no longer available. | |
| 1343 | This behavior allows \emph{strengthened induction predicates} to be expressed
 | |
| 1344 | concisely as meta-level rule statements, i.e.\ $\All{\vec x} \vec\phi \Imp
 | |
| 1345 | \psi$ to indicate ``variable'' parameters $\vec x$ and ``recursive'' | |
| 1346 | assumptions $\vec\phi$. Note that ``$\isarcmd{case}~c$'' already performs
 | |
| 1347 | ``$\FIX{\vec x}$''.  Also note that local definitions may be expressed as
 | |
| 1348 | $\All{\vec x} n \equiv t[\vec x] \Imp \phi[n]$, with induction over $n$.
 | |
| 1349 | ||
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1350 | |
| 13425 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13411diff
changeset | 1351 | In induction proofs, local assumptions introduced by cases are split into two | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13411diff
changeset | 1352 | different kinds: $hyps$ stemming from the rule and $prems$ from the goal | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13411diff
changeset | 1353 | statement. This is reflected in the extracted cases accordingly, so invoking | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13411diff
changeset | 1354 | ``$\isarcmd{case}~c$'' will provide separate facts $c\mathord.hyps$ and
 | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13411diff
changeset | 1355 | $c\mathord.prems$, as well as fact $c$ to hold the all-inclusive list. | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13411diff
changeset | 1356 | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1357 | \medskip | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1358 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1359 | Facts presented to either method are consumed according to the number of | 
| 12618 | 1360 | ``major premises'' of the rule involved (see also \S\ref{sec:cases-induct}),
 | 
| 13041 | 1361 | which is usually $0$ for plain cases and induction rules of datatypes etc.\ | 
| 12618 | 1362 | and $1$ for rules of inductive sets and the like. The remaining facts are | 
| 1363 | inserted into the goal verbatim before the actual $cases$ or $induct$ rule is | |
| 1364 | applied (thus facts may be even passed through an induction). | |
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1365 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1366 | |
| 12621 | 1367 | \subsubsection{Declaring rules}\label{sec:cases-induct-att}
 | 
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1368 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1369 | \indexisarcmd{print-induct-rules}\indexisaratt{cases}\indexisaratt{induct}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1370 | \begin{matharray}{rcl}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1371 |   \isarcmd{print_induct_rules}^* & : & \isarkeep{theory~|~proof} \\
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1372 | cases & : & \isaratt \\ | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1373 | induct & : & \isaratt \\ | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1374 | \end{matharray}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1375 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1376 | \begin{rail}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1377 | 'cases' spec | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1378 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1379 | 'induct' spec | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1380 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1381 | |
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1382 |   spec: ('type' | 'set') ':' nameref
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1383 | ; | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1384 | \end{rail}
 | 
| 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1385 | |
| 13024 | 1386 | \begin{descr}
 | 
| 13041 | 1387 | |
| 13024 | 1388 | \item [$\isarkeyword{print_induct_rules}$] prints cases and induct rules for
 | 
| 1389 | sets and types of the current context. | |
| 13041 | 1390 | |
| 13024 | 1391 | \item [$cases$ and $induct$] (as attributes) augment the corresponding context | 
| 1392 | of rules for reasoning about inductive sets and types, using the | |
| 1393 | corresponding methods of the same name. Certain definitional packages of | |
| 1394 | object-logics usually declare emerging cases and induction rules as | |
| 1395 | expected, so users rarely need to intervene. | |
| 13048 | 1396 | |
| 13024 | 1397 | Manual rule declarations usually include the the $case_names$ and $ps$ | 
| 1398 | attributes to adjust names of cases and parameters of a rule (see | |
| 13048 | 1399 |   \S\ref{sec:cases-induct}); the $consumes$ declaration is taken care of
 | 
| 13024 | 1400 | automatically: $consumes~0$ is specified for ``type'' rules and $consumes~1$ | 
| 1401 | for ``set'' rules. | |
| 13041 | 1402 | |
| 13024 | 1403 | \end{descr}
 | 
| 11691 
fc9bd420162c
induct/cases made generic, removed simplified/stripped options;
 wenzelm parents: 
11469diff
changeset | 1404 | |
| 9614 | 1405 | %%% Local Variables: | 
| 7135 | 1406 | %%% mode: latex | 
| 1407 | %%% TeX-master: "isar-ref" | |
| 9614 | 1408 | %%% End: |