| author | wenzelm | 
| Tue, 13 Feb 2001 16:48:36 +0100 | |
| changeset 11110 | 306beb99e192 | 
| parent 11051 | 00b70f3196c2 | 
| child 11498 | 681aa3dfab4b | 
| permissions | -rw-r--r-- | 
| 7046 | 1 | |
| 7167 | 2 | \chapter{Isabelle/HOL Tools and Packages}\label{ch:hol-tools}
 | 
| 7135 | 3 | |
| 10240 | 4 | \section{Miscellaneous attributes}\label{sec:rule-format}
 | 
| 7990 | 5 | |
| 11039 | 6 | \indexisaratt{rule-format}\indexisaratt{split-format}
 | 
| 7990 | 7 | \begin{matharray}{rcl}
 | 
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 8 | rule_format & : & \isaratt \\ | 
| 11051 | 9 | split_format^* & : & \isaratt \\ | 
| 7990 | 10 | \end{matharray}
 | 
| 11 | ||
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 12 | \railalias{ruleformat}{rule\_format}
 | 
| 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 13 | \railterm{ruleformat}
 | 
| 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 14 | |
| 11039 | 15 | \railalias{splitformat}{split\_format}
 | 
| 16 | \railterm{splitformat}
 | |
| 17 | \railterm{complete}
 | |
| 18 | ||
| 9905 | 19 | \begin{rail}
 | 
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 20 |   ruleformat ('(' noasm ')')?
 | 
| 9905 | 21 | ; | 
| 11039 | 22 |   splitformat (((name * ) + 'and') | ('(' complete ')'))
 | 
| 23 | ; | |
| 9905 | 24 | \end{rail}
 | 
| 25 | ||
| 7990 | 26 | \begin{descr}
 | 
| 9848 | 27 | |
| 9941 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 28 | \item [$rule_format$] causes a theorem to be put into standard object-rule | 
| 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 29 | form, replacing implication and (bounded) universal quantification of HOL by | 
| 
fe05af7ec816
renamed atts: rulify to rule_format, elimify to elim_format;
 wenzelm parents: 
9935diff
changeset | 30 | the corresponding meta-logical connectives. By default, the result is fully | 
| 9905 | 31 | normalized, including assumptions and conclusions at any depth. The | 
| 32 | $no_asm$ option restricts the transformation to the conclusion of a rule. | |
| 11039 | 33 | |
| 34 | \item [$split_format~\vec p@1 \dots \vec p@n$] puts tuple objects into | |
| 35 | canonical form as specified by the arguments given; $\vec p@i$ refers to | |
| 11051 | 36 | occurrences in premise $i$ of the rule. The $split_format~(complete)$ form | 
| 37 |   causes \emph{all} arguments in function applications to be represented
 | |
| 38 | canonically according to their tuple type structure. | |
| 11039 | 39 | |
| 11051 | 40 | Note that these operations tend to invent funny names for new local | 
| 41 | parameters to be introduced. | |
| 11039 | 42 | |
| 7990 | 43 | \end{descr}
 | 
| 44 | ||
| 45 | ||
| 7135 | 46 | \section{Primitive types}
 | 
| 47 | ||
| 7141 | 48 | \indexisarcmd{typedecl}\indexisarcmd{typedef}
 | 
| 49 | \begin{matharray}{rcl}
 | |
| 50 |   \isarcmd{typedecl} & : & \isartrans{theory}{theory} \\
 | |
| 51 |   \isarcmd{typedef} & : & \isartrans{theory}{proof(prove)} \\
 | |
| 52 | \end{matharray}
 | |
| 53 | ||
| 54 | \begin{rail}
 | |
| 55 | 'typedecl' typespec infix? comment? | |
| 56 | ; | |
| 57 | 'typedef' parname? typespec infix? \\ '=' term comment? | |
| 58 | ; | |
| 59 | \end{rail}
 | |
| 60 | ||
| 7167 | 61 | \begin{descr}
 | 
| 7141 | 62 | \item [$\isarkeyword{typedecl}~(\vec\alpha)t$] is similar to the original
 | 
| 63 |   $\isarkeyword{typedecl}$ of Isabelle/Pure (see \S\ref{sec:types-pure}), but
 | |
| 64 | also declares type arity $t :: (term, \dots, term) term$, making $t$ an | |
| 65 | actual HOL type constructor. | |
| 66 | \item [$\isarkeyword{typedef}~(\vec\alpha)t = A$] sets up a goal stating
 | |
| 67 | non-emptiness of the set $A$. After finishing the proof, the theory will be | |
| 7175 | 68 |   augmented by a Gordon/HOL-style type definition.  See \cite{isabelle-HOL}
 | 
| 7335 | 69 | for more information. Note that user-level theories usually do not directly | 
| 70 |   refer to the HOL $\isarkeyword{typedef}$ primitive, but use more advanced
 | |
| 71 |   packages such as $\isarkeyword{record}$ (see \S\ref{sec:record}) and
 | |
| 7175 | 72 |   $\isarkeyword{datatype}$ (see \S\ref{sec:datatype}).
 | 
| 7167 | 73 | \end{descr}
 | 
| 7141 | 74 | |
| 75 | ||
| 76 | \section{Records}\label{sec:record}
 | |
| 77 | ||
| 78 | \indexisarcmd{record}
 | |
| 79 | \begin{matharray}{rcl}
 | |
| 80 |   \isarcmd{record} & : & \isartrans{theory}{theory} \\
 | |
| 81 | \end{matharray}
 | |
| 82 | ||
| 83 | \begin{rail}
 | |
| 84 | 'record' typespec '=' (type '+')? (field +) | |
| 85 | ; | |
| 7135 | 86 | |
| 7141 | 87 | field: name '::' type comment? | 
| 88 | ; | |
| 89 | \end{rail}
 | |
| 90 | ||
| 7167 | 91 | \begin{descr}
 | 
| 7141 | 92 | \item [$\isarkeyword{record}~(\vec\alpha)t = \tau + \vec c :: \vec\sigma$]
 | 
| 93 | defines extensible record type $(\vec\alpha)t$, derived from the optional | |
| 94 | parent record $\tau$ by adding new field components $\vec c :: \vec\sigma$. | |
| 7335 | 95 |   See \cite{isabelle-HOL,NaraschewskiW-TPHOLs98} for more information only
 | 
| 96 | simply-typed extensible records. | |
| 7167 | 97 | \end{descr}
 | 
| 7141 | 98 | |
| 99 | ||
| 100 | \section{Datatypes}\label{sec:datatype}
 | |
| 101 | ||
| 7167 | 102 | \indexisarcmd{datatype}\indexisarcmd{rep-datatype}
 | 
| 7141 | 103 | \begin{matharray}{rcl}
 | 
| 104 |   \isarcmd{datatype} & : & \isartrans{theory}{theory} \\
 | |
| 105 |   \isarcmd{rep_datatype} & : & \isartrans{theory}{theory} \\
 | |
| 106 | \end{matharray}
 | |
| 107 | ||
| 108 | \railalias{repdatatype}{rep\_datatype}
 | |
| 109 | \railterm{repdatatype}
 | |
| 110 | ||
| 111 | \begin{rail}
 | |
| 9848 | 112 | 'datatype' (dtspec + 'and') | 
| 7141 | 113 | ; | 
| 9848 | 114 | repdatatype (name * ) dtrules | 
| 7141 | 115 | ; | 
| 116 | ||
| 9848 | 117 | dtspec: parname? typespec infix? '=' (cons + '|') | 
| 7141 | 118 | ; | 
| 9848 | 119 | cons: name (type * ) mixfix? comment? | 
| 120 | ; | |
| 121 | dtrules: 'distinct' thmrefs 'inject' thmrefs 'induction' thmrefs | |
| 7141 | 122 | \end{rail}
 | 
| 123 | ||
| 7167 | 124 | \begin{descr}
 | 
| 7319 | 125 | \item [$\isarkeyword{datatype}$] defines inductive datatypes in HOL.
 | 
| 126 | \item [$\isarkeyword{rep_datatype}$] represents existing types as inductive
 | |
| 127 | ones, generating the standard infrastructure of derived concepts (primitive | |
| 128 | recursion etc.). | |
| 7167 | 129 | \end{descr}
 | 
| 7141 | 130 | |
| 8449 | 131 | The induction and exhaustion theorems generated provide case names according | 
| 132 | to the constructors involved, while parameters are named after the types (see | |
| 133 | also \S\ref{sec:induct-method}).
 | |
| 134 | ||
| 7319 | 135 | See \cite{isabelle-HOL} for more details on datatypes.  Note that the theory
 | 
| 7335 | 136 | syntax above has been slightly simplified over the old version, usually | 
| 8531 | 137 | requiring more quotes and less parentheses. Apart from proper proof methods | 
| 138 | for case-analysis and induction, there are also emulations of ML tactics | |
| 8945 | 139 | \texttt{case_tac} and \texttt{induct_tac} available, see
 | 
| 8665 | 140 | \S\ref{sec:induct_tac}.
 | 
| 7319 | 141 | |
| 7135 | 142 | |
| 143 | \section{Recursive functions}
 | |
| 144 | ||
| 10771 | 145 | \indexisarcmd{primrec}\indexisarcmd{recdef}\indexisarcmd{recdef-tc}
 | 
| 7141 | 146 | \begin{matharray}{rcl}
 | 
| 147 |   \isarcmd{primrec} & : & \isartrans{theory}{theory} \\
 | |
| 148 |   \isarcmd{recdef} & : & \isartrans{theory}{theory} \\
 | |
| 10771 | 149 |   \isarcmd{recdef_tc}^* & : & \isartrans{theory}{proof(prove)} \\
 | 
| 7141 | 150 | %FIXME | 
| 151 | %  \isarcmd{defer_recdef} & : & \isartrans{theory}{theory} \\
 | |
| 152 | \end{matharray}
 | |
| 153 | ||
| 9949 | 154 | \railalias{recdefsimp}{recdef\_simp}
 | 
| 155 | \railterm{recdefsimp}
 | |
| 156 | ||
| 157 | \railalias{recdefcong}{recdef\_cong}
 | |
| 158 | \railterm{recdefcong}
 | |
| 159 | ||
| 160 | \railalias{recdefwf}{recdef\_wf}
 | |
| 161 | \railterm{recdefwf}
 | |
| 162 | ||
| 10771 | 163 | \railalias{recdeftc}{recdef\_tc}
 | 
| 164 | \railterm{recdeftc}
 | |
| 165 | ||
| 7141 | 166 | \begin{rail}
 | 
| 8657 | 167 | 'primrec' parname? (equation + ) | 
| 168 | ; | |
| 9848 | 169 | 'recdef' name term (eqn + ) hints? | 
| 170 | ; | |
| 10771 | 171 | recdeftc thmdecl? tc comment? | 
| 172 | ; | |
| 8657 | 173 | |
| 9848 | 174 | equation: thmdecl? eqn | 
| 175 | ; | |
| 176 | eqn: prop comment? | |
| 8657 | 177 | ; | 
| 9848 | 178 |   hints: '(' 'hints' (recdefmod * ) ')'
 | 
| 179 | ; | |
| 9949 | 180 | recdefmod: ((recdefsimp | recdefcong | recdefwf) (() | 'add' | 'del') ':' thmrefs) | clasimpmod | 
| 7141 | 181 | ; | 
| 10771 | 182 |   tc: nameref ('(' nat ')')?
 | 
| 183 | ; | |
| 7141 | 184 | \end{rail}
 | 
| 185 | ||
| 7167 | 186 | \begin{descr}
 | 
| 7319 | 187 | \item [$\isarkeyword{primrec}$] defines primitive recursive functions over
 | 
| 9848 | 188 |   datatypes, see also \cite{isabelle-HOL}.
 | 
| 7319 | 189 | \item [$\isarkeyword{recdef}$] defines general well-founded recursive
 | 
| 9848 | 190 |   functions (using the TFL package), see also \cite{isabelle-HOL}.  The
 | 
| 9949 | 191 | $recdef_simp$, $recdef_cong$, and $recdef_wf$ hints refer to auxiliary rules | 
| 192 | to be used in the internal automated proof process of TFL. Additional | |
| 193 |   $clasimpmod$ declarations (cf.\ \S\ref{sec:clasimp}) may be given to tune
 | |
| 194 |   the context of the Simplifier (cf.\ \S\ref{sec:simplifier}) and Classical
 | |
| 195 |   reasoner (cf.\ \S\ref{sec:classical}).
 | |
| 10771 | 196 | \item [$\isarkeyword{recdef_tc}~c~(i)$] recommences the proof for leftover
 | 
| 197 | termination condition number $i$ (default $1$) as generated by a | |
| 198 |   $\isarkeyword{recdef}$ definition of constant $c$.
 | |
| 199 | ||
| 200 |   Note that in most cases, $\isarkeyword{recdef}$ is able to finish its
 | |
| 201 | internal proofs without manual intervention. | |
| 7167 | 202 | \end{descr}
 | 
| 7141 | 203 | |
| 9848 | 204 | Both kinds of recursive definitions accommodate reasoning by induction (cf.\ | 
| 8449 | 205 | \S\ref{sec:induct-method}): rule $c\mathord{.}induct$ (where $c$ is the name
 | 
| 206 | of the function definition) refers to a specific induction rule, with | |
| 207 | parameters named according to the user-specified equations. Case names of | |
| 208 | $\isarkeyword{primrec}$ are that of the datatypes involved, while those of
 | |
| 209 | $\isarkeyword{recdef}$ are numbered (starting from $1$).
 | |
| 210 | ||
| 8657 | 211 | The equations provided by these packages may be referred later as theorem list | 
| 212 | $f\mathord.simps$, where $f$ is the (collective) name of the functions | |
| 213 | defined. Individual equations may be named explicitly as well; note that for | |
| 214 | $\isarkeyword{recdef}$ each specification given by the user may result in
 | |
| 215 | several theorems. | |
| 216 | ||
| 9935 | 217 | \medskip Hints for $\isarkeyword{recdef}$ may be also declared globally, using
 | 
| 218 | the following attributes. | |
| 219 | ||
| 220 | \indexisaratt{recdef-simp}\indexisaratt{recdef-cong}\indexisaratt{recdef-wf}
 | |
| 221 | \begin{matharray}{rcl}
 | |
| 222 | recdef_simp & : & \isaratt \\ | |
| 223 | recdef_cong & : & \isaratt \\ | |
| 224 | recdef_wf & : & \isaratt \\ | |
| 225 | \end{matharray}
 | |
| 226 | ||
| 227 | \railalias{recdefsimp}{recdef\_simp}
 | |
| 228 | \railterm{recdefsimp}
 | |
| 229 | ||
| 230 | \railalias{recdefcong}{recdef\_cong}
 | |
| 231 | \railterm{recdefcong}
 | |
| 232 | ||
| 233 | \railalias{recdefwf}{recdef\_wf}
 | |
| 234 | \railterm{recdefwf}
 | |
| 235 | ||
| 236 | \begin{rail}
 | |
| 237 | (recdefsimp | recdefcong | recdefwf) (() | 'add' | 'del') | |
| 238 | ; | |
| 239 | \end{rail}
 | |
| 240 | ||
| 7141 | 241 | |
| 10549 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 242 | \section{(Co)Inductive sets}\label{sec:inductive}
 | 
| 7135 | 243 | |
| 9602 | 244 | \indexisarcmd{inductive}\indexisarcmd{coinductive}\indexisaratt{mono}
 | 
| 7141 | 245 | \begin{matharray}{rcl}
 | 
| 246 |   \isarcmd{inductive} & : & \isartrans{theory}{theory} \\
 | |
| 9848 | 247 |   \isarcmd{coinductive} & : & \isartrans{theory}{theory} \\
 | 
| 7990 | 248 | mono & : & \isaratt \\ | 
| 7141 | 249 | \end{matharray}
 | 
| 250 | ||
| 251 | \railalias{condefs}{con\_defs}
 | |
| 9602 | 252 | \railterm{condefs}
 | 
| 7141 | 253 | |
| 254 | \begin{rail}
 | |
| 9848 | 255 |   ('inductive' | 'coinductive') sets intros monos?
 | 
| 7141 | 256 | ; | 
| 7990 | 257 | 'mono' (() | 'add' | 'del') | 
| 258 | ; | |
| 9848 | 259 | |
| 260 | sets: (term comment? +) | |
| 261 | ; | |
| 262 | intros: 'intros' attributes? (thmdecl? prop comment? +) | |
| 263 | ; | |
| 264 | monos: 'monos' thmrefs comment? | |
| 265 | ; | |
| 7141 | 266 | \end{rail}
 | 
| 267 | ||
| 7167 | 268 | \begin{descr}
 | 
| 7319 | 269 | \item [$\isarkeyword{inductive}$ and $\isarkeyword{coinductive}$] define
 | 
| 270 | (co)inductive sets from the given introduction rules. | |
| 8547 | 271 | \item [$mono$] declares monotonicity rules. These rule are involved in the | 
| 272 |   automated monotonicity proof of $\isarkeyword{inductive}$.
 | |
| 7167 | 273 | \end{descr}
 | 
| 7141 | 274 | |
| 8449 | 275 | See \cite{isabelle-HOL} for further information on inductive definitions in
 | 
| 276 | HOL. | |
| 7319 | 277 | |
| 7141 | 278 | |
| 8449 | 279 | \section{Proof by cases and induction}\label{sec:induct-method}
 | 
| 280 | ||
| 8666 | 281 | \subsection{Proof methods}\label{sec:induct-method-proper}
 | 
| 7141 | 282 | |
| 8449 | 283 | \indexisarmeth{cases}\indexisarmeth{induct}
 | 
| 7319 | 284 | \begin{matharray}{rcl}
 | 
| 8449 | 285 | cases & : & \isarmeth \\ | 
| 7319 | 286 | induct & : & \isarmeth \\ | 
| 287 | \end{matharray}
 | |
| 288 | ||
| 8449 | 289 | The $cases$ and $induct$ methods provide a uniform interface to case analysis | 
| 290 | and induction over datatypes, inductive sets, and recursive functions. The | |
| 291 | corresponding rules may be specified and instantiated in a casual manner. | |
| 292 | Furthermore, these methods provide named local contexts that may be invoked | |
| 293 | via the $\CASENAME$ proof command within the subsequent proof text (cf.\ | |
| 8484 | 294 | \S\ref{sec:cases}).  This accommodates compact proof texts even when reasoning
 | 
| 295 | about large specifications. | |
| 7319 | 296 | |
| 297 | \begin{rail}
 | |
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 298 |   'cases' ('(' 'simplified' ')')? spec
 | 
| 9848 | 299 | ; | 
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 300 |   'induct' ('(' 'stripped' ')')? spec
 | 
| 7319 | 301 | ; | 
| 302 | ||
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 303 | spec: open? args rule? params? | 
| 9848 | 304 | ; | 
| 305 |   open: '(' 'open' ')'
 | |
| 306 | ; | |
| 307 | args: (insts * 'and') | |
| 308 | ; | |
| 8449 | 309 |   rule: ('type' | 'set') ':' nameref | 'rule' ':' thmref
 | 
| 7319 | 310 | ; | 
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 311 | params: 'of' ':' insts | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 312 | ; | 
| 7319 | 313 | \end{rail}
 | 
| 314 | ||
| 315 | \begin{descr}
 | |
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 316 | \item [$cases~insts~R~ps$] applies method $rule$ with an appropriate case | 
| 9602 | 317 | distinction theorem, instantiated to the subjects $insts$. Symbolic case | 
| 318 | names are bound according to the rule's local contexts. | |
| 8449 | 319 | |
| 320 | The rule is determined as follows, according to the facts and arguments | |
| 321 | passed to the $cases$ method: | |
| 322 |   \begin{matharray}{llll}
 | |
| 9695 | 323 |     \Text{facts}    &       & \Text{arguments} & \Text{rule} \\\hline
 | 
| 324 |                     & cases &           & \Text{classical case split} \\
 | |
| 325 |                     & cases & t         & \Text{datatype exhaustion (type of $t$)} \\
 | |
| 326 |     \edrv a \in A   & cases & \dots     & \Text{inductive set elimination (of $A$)} \\
 | |
| 327 |     \dots           & cases & \dots ~ R & \Text{explicit rule $R$} \\
 | |
| 8449 | 328 |   \end{matharray}
 | 
| 9602 | 329 | |
| 330 |   Several instantiations may be given, referring to the \emph{suffix} of
 | |
| 331 |   premises of the case rule; within each premise, the \emph{prefix} of
 | |
| 332 | variables is instantiated. In most situations, only a single term needs to | |
| 333 | be specified; this refers to the first variable of the last premise (it is | |
| 334 | usually the same for all cases). | |
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 335 | |
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 336 | Additional parameters may be specified as $ps$; these are applied after the | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 337 | primary instantiation in the same manner as by the $of$ attribute (cf.\ | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 338 |   \S\ref{sec:pure-meth-att}).  This feature is rarely needed in practice; a
 | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 339 | typical application would be to specify additional arguments for rules | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 340 | stemming from parameterized inductive definitions (see also | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 341 |   \S\ref{sec:inductive}).
 | 
| 8449 | 342 | |
| 343 | The $simplified$ option causes ``obvious cases'' of the rule to be solved | |
| 344 | beforehand, while the others are left unscathed. | |
| 345 | ||
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 346 | The $open$ option causes the parameters of the new local contexts to be | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 347 | exposed to the current proof context. Thus local variables stemming from | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 348 | distant parts of the theory development may be introduced in an implicit | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 349 | manner, which can be quite confusing to the reader. Furthermore, this | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 350 | option may cause unwanted hiding of existing local variables, resulting in | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 351 | less robust proof texts. | 
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 352 | |
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 353 | \item [$induct~insts~R~ps$] is analogous to the $cases$ method, but refers to | 
| 8449 | 354 | induction rules, which are determined as follows: | 
| 355 |   \begin{matharray}{llll}
 | |
| 9695 | 356 |     \Text{facts}    &        & \Text{arguments} & \Text{rule} \\\hline
 | 
| 357 |                     & induct & P ~ x ~ \dots & \Text{datatype induction (type of $x$)} \\
 | |
| 358 |     \edrv x \in A   & induct & \dots         & \Text{set induction (of $A$)} \\
 | |
| 359 |     \dots           & induct & \dots ~ R     & \Text{explicit rule $R$} \\
 | |
| 8449 | 360 |   \end{matharray}
 | 
| 361 | ||
| 362 | Several instantiations may be given, each referring to some part of a mutual | |
| 363 | inductive definition or datatype --- only related partial induction rules | |
| 364 | may be used together, though. Any of the lists of terms $P, x, \dots$ | |
| 365 |   refers to the \emph{suffix} of variables present in the induction rule.
 | |
| 366 | This enables the writer to specify only induction variables, or both | |
| 367 | predicates and variables, for example. | |
| 7507 | 368 | |
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 369 | Additional parameters may be given in the same way as for $cases$. | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 370 | |
| 8449 | 371 | The $stripped$ option causes implications and (bounded) universal | 
| 372 | quantifiers to be removed from each new subgoal emerging from the | |
| 10456 | 373 | application of the induction rule. This accommodates special applications | 
| 374 | of ``strengthened induction predicates''. This option is rarely needed, the | |
| 375 | $induct$ method already handles proper rules appropriately by default. | |
| 9307 | 376 | |
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 377 | The $open$ option has the same effect as for the $cases$ method, see above. | 
| 7319 | 378 | \end{descr}
 | 
| 7141 | 379 | |
| 8484 | 380 | Above methods produce named local contexts (cf.\ \S\ref{sec:cases}), as
 | 
| 381 | determined by the instantiated rule \emph{before} it has been applied to the
 | |
| 382 | internal proof state.\footnote{As a general principle, Isar proof text may
 | |
| 8449 | 383 | never refer to parts of proof states directly.} Thus proper use of symbolic | 
| 384 | cases usually require the rule to be instantiated fully, as far as the | |
| 385 | emerging local contexts and subgoals are concerned. In particular, for | |
| 386 | induction both the predicates and variables have to be specified. Otherwise | |
| 8547 | 387 | the $\CASENAME$ command would refuse to invoke cases containing schematic | 
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 388 | variables. Furthermore the resulting local goal statement is bound to the | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 389 | term variable $\Var{case}$\indexisarvar{case} --- for each case where it is
 | 
| 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 390 | fully specified. | 
| 8449 | 391 | |
| 9602 | 392 | The $\isarkeyword{print_cases}$ command (\S\ref{sec:cases}) prints all named
 | 
| 8547 | 393 | cases present in the current proof state. | 
| 8449 | 394 | |
| 10456 | 395 | \medskip | 
| 396 | ||
| 397 | It is important to note that there is a fundamental difference of the $cases$ | |
| 398 | and $induct$ methods in handling of non-atomic goal statements: $cases$ just | |
| 399 | applies a certain rule in backward fashion, splitting the result into new | |
| 400 | goals with the local contexts being augmented in a purely monotonic manner. | |
| 401 | ||
| 402 | In contrast, $induct$ passes the full goal statement through the ``recursive'' | |
| 403 | course involved in the induction. Thus the original statement is basically | |
| 404 | replaced by separate copies, corresponding to the induction hypotheses and | |
| 405 | conclusion; the original goal context is no longer available. This behavior | |
| 406 | allows \emph{strengthened induction predicates} to be expressed concisely as
 | |
| 407 | meta-level rule statements, i.e.\ $\All{\vec x} \vec\phi \Imp \psi$ to
 | |
| 408 | indicate ``variable'' parameters $\vec x$ and ``recursive'' assumptions | |
| 409 | $\vec\phi$.  Also note that local definitions may be expressed as $\All{\vec
 | |
| 410 | x} n \equiv t[\vec x] \Imp \phi[n]$, with induction over $n$. | |
| 411 | ||
| 10549 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 412 | \medskip | 
| 8449 | 413 | |
| 10549 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 414 | Facts presented to either method are consumed according to the number of | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 415 | ``major premises'' of the rule involved (see also \S\ref{sec:induct-att} and
 | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 416 | \S\ref{sec:cases}), which is usually $0$ for plain cases and induction rules
 | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 417 | of datatypes etc.\ and $1$ for rules of inductive sets and the like. The | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 418 | remaining facts are inserted into the goal verbatim before the actual $cases$ | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 419 | or $induct$ rule is applied (thus facts may be even passed through an | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 420 | induction). | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 421 | |
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 422 | Note that whenever facts are present, the default rule selection scheme would | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 423 | provide a ``set'' rule only, with the first fact consumed and the rest | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 424 | inserted into the goal. In order to pass all facts into a ``type'' rule | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 425 | instead, one would have to specify this explicitly, e.g.\ by appending | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 426 | ``$type: name$'' to the method argument. | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 427 | |
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 428 | |
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 429 | \subsection{Declaring rules}\label{sec:induct-att}
 | 
| 8449 | 430 | |
| 431 | \indexisaratt{cases}\indexisaratt{induct}
 | |
| 432 | \begin{matharray}{rcl}
 | |
| 433 | cases & : & \isaratt \\ | |
| 434 | induct & : & \isaratt \\ | |
| 435 | \end{matharray}
 | |
| 436 | ||
| 437 | \begin{rail}
 | |
| 438 | 'cases' spec | |
| 439 | ; | |
| 440 | 'induct' spec | |
| 441 | ; | |
| 442 | ||
| 443 |   spec: ('type' | 'set') ':' nameref
 | |
| 444 | ; | |
| 445 | \end{rail}
 | |
| 446 | ||
| 447 | The $cases$ and $induct$ attributes augment the corresponding context of rules | |
| 448 | for reasoning about inductive sets and types. The standard rules are already | |
| 449 | declared by HOL definitional packages. For special applications, these may be | |
| 450 | replaced manually by variant versions. | |
| 451 | ||
| 10802 
7fa042e28c43
'cases' / 'induct' method: ?case binding, 'of:' spec;
 wenzelm parents: 
10771diff
changeset | 452 | Refer to the $case_names$ and $ps$ attributes (see \S\ref{sec:cases}) to
 | 
| 8484 | 453 | adjust names of cases and parameters of a rule. | 
| 454 | ||
| 10549 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 455 | The $consumes$ declaration (cf.\ \S\ref{sec:cases}) is taken care of
 | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 456 | automatically (if none had been given already): $consumes~0$ is specified for | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 457 | ``type'' rules and $consumes~1$ for ``set'' rules. | 
| 
5e19ae8d9582
cases/induct: tuned handling of facts ('consumes');
 wenzelm parents: 
10456diff
changeset | 458 | |
| 7046 | 459 | |
| 8665 | 460 | \subsection{Emulating tactic scripts}\label{sec:induct_tac}
 | 
| 461 | ||
| 462 | \indexisarmeth{case-tac}\indexisarmeth{induct-tac}
 | |
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 463 | \indexisarmeth{ind-cases}\indexisarcmd{inductive-cases}
 | 
| 8665 | 464 | \begin{matharray}{rcl}
 | 
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 465 | case_tac^* & : & \isarmeth \\ | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 466 | induct_tac^* & : & \isarmeth \\ | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 467 | ind_cases^* & : & \isarmeth \\ | 
| 9602 | 468 |   \isarcmd{inductive_cases} & : & \isartrans{theory}{theory} \\
 | 
| 8665 | 469 | \end{matharray}
 | 
| 470 | ||
| 471 | \railalias{casetac}{case\_tac}
 | |
| 472 | \railterm{casetac}
 | |
| 9602 | 473 | |
| 8665 | 474 | \railalias{inducttac}{induct\_tac}
 | 
| 475 | \railterm{inducttac}
 | |
| 476 | ||
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 477 | \railalias{indcases}{ind\_cases}
 | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 478 | \railterm{indcases}
 | 
| 9602 | 479 | |
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 480 | \railalias{inductivecases}{inductive\_cases}
 | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 481 | \railterm{inductivecases}
 | 
| 9602 | 482 | |
| 8665 | 483 | \begin{rail}
 | 
| 8666 | 484 | casetac goalspec? term rule? | 
| 8665 | 485 | ; | 
| 8692 | 486 | inducttac goalspec? (insts * 'and') rule? | 
| 8666 | 487 | ; | 
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 488 | indcases (prop +) | 
| 9602 | 489 | ; | 
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 490 | inductivecases thmdecl? (prop +) comment? | 
| 9602 | 491 | ; | 
| 8666 | 492 | |
| 493 |   rule: ('rule' ':' thmref)
 | |
| 8665 | 494 | ; | 
| 495 | \end{rail}
 | |
| 496 | ||
| 9602 | 497 | \begin{descr}
 | 
| 498 | \item [$case_tac$ and $induct_tac$] admit to reason about inductive datatypes | |
| 499 | only (unless an alternative rule is given explicitly). Furthermore, | |
| 500 | $case_tac$ does a classical case split on booleans; $induct_tac$ allows only | |
| 501 | variables to be given as instantiation. These tactic emulations feature | |
| 502 | both goal addressing and dynamic instantiation. Note that named local | |
| 503 |   contexts (see \S\ref{sec:cases}) are \emph{not} provided as would be by the
 | |
| 504 | proper $induct$ and $cases$ proof methods (see | |
| 505 |   \S\ref{sec:induct-method-proper}).
 | |
| 506 | ||
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 507 | \item [$ind_cases$ and $\isarkeyword{inductive_cases}$] provide an interface
 | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 508 |   to the \texttt{mk_cases} operation.  Rules are simplified in an unrestricted
 | 
| 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 509 | forward manner, unlike the proper $cases$ method (see | 
| 9602 | 510 |   \S\ref{sec:induct-method-proper}) which requires simplified cases to be
 | 
| 511 | solved completely. | |
| 512 | ||
| 9616 
b80ea2b32f8e
cases/induct method: 'opaque' by default; added 'open' option;
 wenzelm parents: 
9602diff
changeset | 513 | While $ind_cases$ is a proof method to apply the result immediately as | 
| 9602 | 514 |   elimination rules, $\isarkeyword{inductive_cases}$ provides case split
 | 
| 515 | theorems at the theory level for later use, | |
| 516 | \end{descr}
 | |
| 8665 | 517 | |
| 518 | ||
| 7390 | 519 | \section{Arithmetic}
 | 
| 520 | ||
| 9642 | 521 | \indexisarmeth{arith}\indexisaratt{arith-split}
 | 
| 7390 | 522 | \begin{matharray}{rcl}
 | 
| 523 | arith & : & \isarmeth \\ | |
| 9602 | 524 | arith_split & : & \isaratt \\ | 
| 7390 | 525 | \end{matharray}
 | 
| 526 | ||
| 8506 | 527 | \begin{rail}
 | 
| 528 | 'arith' '!'? | |
| 529 | ; | |
| 530 | \end{rail}
 | |
| 531 | ||
| 7390 | 532 | The $arith$ method decides linear arithmetic problems (on types $nat$, $int$, | 
| 8506 | 533 | $real$). Any current facts are inserted into the goal before running the | 
| 534 | procedure. The ``!''~argument causes the full context of assumptions to be | |
| 9602 | 535 | included. The $arith_split$ attribute declares case split rules to be | 
| 536 | expanded before the arithmetic procedure is invoked. | |
| 8506 | 537 | |
| 538 | Note that a simpler (but faster) version of arithmetic reasoning is already | |
| 539 | performed by the Simplifier. | |
| 7390 | 540 | |
| 541 | ||
| 7046 | 542 | %%% Local Variables: | 
| 543 | %%% mode: latex | |
| 544 | %%% TeX-master: "isar-ref" | |
| 545 | %%% End: |