| author | clasohm | 
| Fri, 27 Jan 1995 12:28:05 +0100 | |
| changeset 881 | d7dcba167ed8 | 
| parent 866 | 2d3d020eef11 | 
| child 1222 | d99d13a0213f | 
| permissions | -rw-r--r-- | 
| 104 | 1 | %% $Id$ | 
| 2 | \chapter{Proof Management: The Subgoal Module}
 | |
| 3 | \index{proofs|(}
 | |
| 4 | \index{subgoal module|(}
 | |
| 5 | \index{reading!goals|see{proofs, starting}}
 | |
| 321 | 6 | |
| 7 | The subgoal module stores the current proof state\index{proof state} and
 | |
| 8 | many previous states; commands can produce new states or return to previous | |
| 9 | ones.  The {\em state list\/} at level $n$ is a list of pairs
 | |
| 104 | 10 | \[ [(\psi@n,\Psi@n),\; (\psi@{n-1},\Psi@{n-1}),\; \ldots,\; (\psi@0,[])] \]
 | 
| 11 | where $\psi@n$ is the current proof state, $\psi@{n-1}$ is the previous
 | |
| 12 | one, \ldots, and $\psi@0$ is the initial proof state. The $\Psi@i$ are | |
| 13 | sequences (lazy lists) of proof states, storing branch points where a | |
| 14 | tactic returned a list longer than one. The state lists permit various | |
| 15 | forms of backtracking. | |
| 16 | ||
| 17 | Chopping elements from the state list reverts to previous proof states. | |
| 18 | Besides this, the \ttindex{undo} command keeps a list of state lists.  The
 | |
| 19 | module actually maintains a stack of state lists, to support several | |
| 20 | proofs at the same time. | |
| 21 | ||
| 22 | The subgoal module always contains some proof state. At the start of the | |
| 23 | Isabelle session, this state consists of a dummy formula. | |
| 24 | ||
| 25 | ||
| 26 | \section{Basic commands}
 | |
| 27 | Most proofs begin with {\tt goal} or {\tt goalw} and require no other
 | |
| 28 | commands than {\tt by}, {\tt chop} and {\tt undo}.  They typically end with
 | |
| 29 | a call to {\tt result}.
 | |
| 30 | \subsection{Starting a backward proof}
 | |
| 321 | 31 | \index{proofs!starting}
 | 
| 104 | 32 | \begin{ttbox} 
 | 
| 33 | goal : theory -> string -> thm list | |
| 34 | goalw : theory -> thm list -> string -> thm list | |
| 35 | goalw_cterm : thm list -> Sign.cterm -> thm list | |
| 36 | premises : unit -> thm list | |
| 37 | \end{ttbox}
 | |
| 38 | The {\tt goal} commands start a new proof by setting the goal.  They
 | |
| 39 | replace the current state list by a new one consisting of the initial proof | |
| 40 | state.  They also empty the \ttindex{undo} list; this command cannot be
 | |
| 41 | undone! | |
| 42 | ||
| 43 | They all return a list of meta-hypotheses taken from the main goal. If | |
| 44 | this list is non-empty, bind its value to an \ML{} identifier by typing
 | |
| 45 | something like | |
| 46 | \begin{ttbox} 
 | |
| 332 | 47 | val prems = goal{\it theory\/ formula};
 | 
| 321 | 48 | \end{ttbox}\index{assumptions!of main goal}
 | 
| 104 | 49 | These assumptions serve as the premises when you are deriving a rule. They | 
| 50 | are also stored internally and can be retrieved later by the function {\tt
 | |
| 321 | 51 |   premises}.  When the proof is finished, {\tt result} compares the
 | 
| 104 | 52 | stored assumptions with the actual assumptions in the proof state. | 
| 53 | ||
| 321 | 54 | \index{definitions!unfolding}
 | 
| 104 | 55 | Some of the commands unfold definitions using meta-rewrite rules. This | 
| 332 | 56 | expansion affects both the initial subgoal and the premises, which would | 
| 321 | 57 | otherwise require use of {\tt rewrite_goals_tac} and
 | 
| 58 | {\tt rewrite_rule}.
 | |
| 104 | 59 | |
| 321 | 60 | \index{*"!"! symbol!in main goal}
 | 
| 61 | If the main goal has the form {\tt"!!{\it vars}.\ \ldots"}, with an
 | |
| 62 | outermost quantifier, then the list of premises will be empty. Subgoal~1 | |
| 63 | will contain the meta-quantified {\it vars\/} as parameters and the goal's
 | |
| 64 | premises as assumptions. This avoids having to call | |
| 65 | \ttindex{cut_facts_tac} with the list of premises (\S\ref{cut_facts_tac}).
 | |
| 104 | 66 | |
| 321 | 67 | \begin{ttdescription}
 | 
| 68 | \item[\ttindexbold{goal} {\it theory} {\it formula};] 
 | |
| 104 | 69 | begins a new proof, where {\it theory} is usually an \ML\ identifier
 | 
| 70 | and the {\it formula\/} is written as an \ML\ string.
 | |
| 71 | ||
| 321 | 72 | \item[\ttindexbold{goalw} {\it theory} {\it defs} {\it formula};] 
 | 
| 104 | 73 | is like {\tt goal} but also applies the list of {\it defs\/} as
 | 
| 74 | meta-rewrite rules to the first subgoal and the premises. | |
| 321 | 75 | \index{meta-rewriting}
 | 
| 104 | 76 | |
| 321 | 77 | \item[\ttindexbold{goalw_cterm} {\it theory} {\it defs} {\it ct};] 
 | 
| 78 | is a version of {\tt goalw} for programming applications.  The main goal is
 | |
| 104 | 79 | supplied as a cterm, not as a string. Typically, the cterm is created from | 
| 80 | a term~$t$ by \hbox{\tt Sign.cterm_of (sign_of thy) $t$}.
 | |
| 81 | \index{*Sign.cterm_of}\index{*sign_of}
 | |
| 82 | ||
| 83 | \item[\ttindexbold{premises}()] 
 | |
| 321 | 84 | returns the list of meta-hypotheses associated with the current proof (in | 
| 85 | case you forgot to bind them to an \ML{} identifier).
 | |
| 86 | \end{ttdescription}
 | |
| 87 | ||
| 104 | 88 | |
| 89 | \subsection{Applying a tactic}
 | |
| 321 | 90 | \index{tactics!commands for applying}
 | 
| 104 | 91 | \begin{ttbox} 
 | 
| 92 | by : tactic -> unit | |
| 93 | byev : tactic list -> unit | |
| 94 | \end{ttbox}
 | |
| 95 | These commands extend the state list. They apply a tactic to the current | |
| 96 | proof state. If the tactic succeeds, it returns a non-empty sequence of | |
| 97 | next states. The head of the sequence becomes the next state, while the | |
| 98 | tail is retained for backtracking (see~{\tt back}).
 | |
| 321 | 99 | \begin{ttdescription} \item[\ttindexbold{by} {\it tactic};] 
 | 
| 104 | 100 | applies the {\it tactic\/} to the proof state.
 | 
| 101 | ||
| 321 | 102 | \item[\ttindexbold{byev} {\it tactics};] 
 | 
| 104 | 103 | applies the list of {\it tactics}, one at a time.  It is useful for testing
 | 
| 104 | calls to {\tt prove_goal}, and abbreviates \hbox{\tt by (EVERY {\it
 | |
| 105 | tactics})}. | |
| 321 | 106 | \end{ttdescription}
 | 
| 104 | 107 | |
| 286 | 108 | \noindent{\it Error indications:}\nobreak
 | 
| 104 | 109 | \begin{itemize}
 | 
| 286 | 110 | \item {\footnotesize\tt "by:\ tactic failed"} means that the tactic
 | 
| 111 | returned an empty sequence when applied to the current proof state. | |
| 112 | \item {\footnotesize\tt "Warning:\ same as previous level"} means that the
 | |
| 113 | new proof state is identical to the previous state. | |
| 114 | \item{\footnotesize\tt "Warning:\ signature of proof state has changed"}
 | |
| 115 | means that some rule was applied whose theory is outside the theory of | |
| 116 | the initial proof state. This could signify a mistake such as expressing | |
| 117 | the goal in intuitionistic logic and proving it using classical logic. | |
| 104 | 118 | \end{itemize}
 | 
| 119 | ||
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 120 | \subsection{Extracting and storing the proved theorem}
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 121 | \label{ExtractingAndStoringTheProvedTheorem}
 | 
| 321 | 122 | \index{theorems!from subgoal module}
 | 
| 104 | 123 | \begin{ttbox} 
 | 
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 124 | result : unit -> thm | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 125 | uresult : unit -> thm | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 126 | bind_thm : string * thm -> unit | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 127 | qed : string -> unit | 
| 104 | 128 | \end{ttbox}
 | 
| 321 | 129 | \begin{ttdescription}
 | 
| 130 | \item[\ttindexbold{result}()]\index{assumptions!of main goal}
 | |
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 131 | returns the final theorem, after converting the free variables to | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 132 | schematics. It discharges the assumptions supplied to the matching | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 133 |   {\tt goal} command.  
 | 
| 104 | 134 | |
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 135 | It raises an exception unless the proof state passes certain checks. There | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 136 |   must be no assumptions other than those supplied to {\tt goal}.  There
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 137 | must be no subgoals. The theorem proved must be a (first-order) instance | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 138 |   of the original goal, as stated in the {\tt goal} command.  This allows
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 139 |   {\bf answer extraction} --- instantiation of variables --- but no other
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 140 | changes to the main goal. The theorem proved must have the same signature | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 141 | as the initial proof state. | 
| 104 | 142 | |
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 143 | These checks are needed because an Isabelle tactic can return any proof | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 144 | state at all. | 
| 321 | 145 | |
| 146 | \item[\ttindexbold{uresult}()] is like {\tt result()} but omits the checks.
 | |
| 147 | It is needed when the initial goal contains function unknowns, when | |
| 148 | definitions are unfolded in the main goal (by calling | |
| 149 |   \ttindex{rewrite_tac}),\index{definitions!unfolding} or when
 | |
| 150 |   \ttindex{assume_ax} has been used.
 | |
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 151 | |
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 152 | \item[\ttindexbold{bind_thm}($name$, $thm$)]\index{theorems!storing of}
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 153 |   stores {\tt standard($thm$)} (see \S\ref{MiscellaneousForwardRules})
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 154 | in Isabelle's theorem database and in the ML variable $name$. The | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 155 |   theorem can be retrieved from Isabelle's database by {\tt get_thm}
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 156 |   (see \S\ref{BasicOperationsOnTheories}).
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 157 | |
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 158 | \item[\ttindexbold{qed} $name$]
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 159 |   combines {\tt result} and {\tt bind_thm} in that it first uses {\tt
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 160 |   result()} to get the theorem and then stores it like {\tt bind_thm}.
 | 
| 321 | 161 | \end{ttdescription}
 | 
| 104 | 162 | |
| 163 | ||
| 164 | \subsection{Undoing and backtracking}
 | |
| 165 | \begin{ttbox} 
 | |
| 166 | chop : unit -> unit | |
| 167 | choplev : int -> unit | |
| 168 | back : unit -> unit | |
| 169 | undo : unit -> unit | |
| 170 | \end{ttbox}
 | |
| 321 | 171 | \begin{ttdescription}
 | 
| 104 | 172 | \item[\ttindexbold{chop}();] 
 | 
| 173 | deletes the top level of the state list, cancelling the last \ttindex{by}
 | |
| 174 | command.  It provides a limited undo facility, and the {\tt undo} command
 | |
| 175 | can cancel it. | |
| 176 | ||
| 177 | \item[\ttindexbold{choplev} {\it n};] 
 | |
| 178 | truncates the state list to level~{\it n}. 
 | |
| 179 | ||
| 180 | \item[\ttindexbold{back}();]
 | |
| 181 | searches the state list for a non-empty branch point, starting from the top | |
| 182 | level. The first one found becomes the current proof state --- the most | |
| 286 | 183 | recent alternative branch is taken. This is a form of interactive | 
| 184 | backtracking. | |
| 104 | 185 | |
| 186 | \item[\ttindexbold{undo}();] 
 | |
| 187 | cancels the most recent change to the proof state by the commands \ttindex{by},
 | |
| 188 | {\tt chop}, {\tt choplev}, and~{\tt back}.  It {\bf cannot}
 | |
| 189 | cancel {\tt goal} or {\tt undo} itself.  It can be repeated to
 | |
| 190 | cancel a series of commands. | |
| 321 | 191 | \end{ttdescription}
 | 
| 286 | 192 | |
| 193 | \goodbreak | |
| 194 | \noindent{\it Error indications for {\tt back}:}\par\nobreak
 | |
| 104 | 195 | \begin{itemize}
 | 
| 286 | 196 | \item{\footnotesize\tt"Warning:\ same as previous choice at this level"}
 | 
| 197 |   means {\tt back} found a non-empty branch point, but that it contained
 | |
| 198 | the same proof state as the current one. | |
| 199 | \item{\footnotesize\tt "Warning:\ signature of proof state has changed"}
 | |
| 200 | means the signature of the alternative proof state differs from that of | |
| 201 | the current state. | |
| 202 | \item {\footnotesize\tt "back:\ no alternatives"} means {\tt back} could
 | |
| 203 | find no alternative proof state. | |
| 104 | 204 | \end{itemize}
 | 
| 205 | ||
| 507 | 206 | \subsection{Printing the proof state}\label{sec:goals-printing}
 | 
| 321 | 207 | \index{proof state!printing of}
 | 
| 104 | 208 | \begin{ttbox} 
 | 
| 209 | pr : unit -> unit | |
| 210 | prlev : int -> unit | |
| 211 | goals_limit: int ref \hfill{\bf initially 10}
 | |
| 212 | \end{ttbox}
 | |
| 507 | 213 | See also the printing control options described in | 
| 214 | \S\ref{sec:printing-control}. 
 | |
| 321 | 215 | \begin{ttdescription}
 | 
| 104 | 216 | \item[\ttindexbold{pr}();] 
 | 
| 217 | prints the current proof state. | |
| 218 | ||
| 219 | \item[\ttindexbold{prlev} {\it n};] 
 | |
| 220 | prints the proof state at level {\it n}.  This allows you to review the
 | |
| 221 | previous steps of the proof. | |
| 222 | ||
| 321 | 223 | \item[\ttindexbold{goals_limit} := {\it k};] 
 | 
| 104 | 224 | specifies~$k$ as the maximum number of subgoals to print. | 
| 321 | 225 | \end{ttdescription}
 | 
| 104 | 226 | |
| 227 | ||
| 228 | \subsection{Timing}
 | |
| 321 | 229 | \index{timing statistics}\index{proofs!timing}
 | 
| 104 | 230 | \begin{ttbox} 
 | 
| 231 | proof_timing: bool ref \hfill{\bf initially false}
 | |
| 232 | \end{ttbox}
 | |
| 233 | ||
| 321 | 234 | \begin{ttdescription}
 | 
| 235 | \item[\ttindexbold{proof_timing} := true;] 
 | |
| 104 | 236 | makes the \ttindex{by} and \ttindex{prove_goal} commands display how much
 | 
| 237 | processor time was spent. This information is compiler-dependent. | |
| 321 | 238 | \end{ttdescription}
 | 
| 104 | 239 | |
| 240 | ||
| 241 | \section{Shortcuts for applying tactics}
 | |
| 321 | 242 | \index{shortcuts!for {\tt by} commands}
 | 
| 104 | 243 | These commands call \ttindex{by} with common tactics.  Their chief purpose
 | 
| 244 | is to minimise typing, although the scanning shortcuts are useful in their | |
| 245 | own right.  Chapter~\ref{tactics} explains the tactics themselves.
 | |
| 246 | ||
| 247 | \subsection{Refining a given subgoal}
 | |
| 248 | \begin{ttbox} 
 | |
| 321 | 249 | ba : int -> unit | 
| 250 | br : thm -> int -> unit | |
| 251 | be : thm -> int -> unit | |
| 252 | bd : thm -> int -> unit | |
| 253 | brs : thm list -> int -> unit | |
| 254 | bes : thm list -> int -> unit | |
| 255 | bds : thm list -> int -> unit | |
| 104 | 256 | \end{ttbox}
 | 
| 257 | ||
| 321 | 258 | \begin{ttdescription}
 | 
| 104 | 259 | \item[\ttindexbold{ba} {\it i};] 
 | 
| 321 | 260 | performs \hbox{\tt by (assume_tac {\it i});}
 | 
| 104 | 261 | |
| 262 | \item[\ttindexbold{br} {\it thm} {\it i};] 
 | |
| 321 | 263 | performs \hbox{\tt by (resolve_tac [{\it thm}] {\it i});}
 | 
| 104 | 264 | |
| 265 | \item[\ttindexbold{be} {\it thm} {\it i};] 
 | |
| 321 | 266 | performs \hbox{\tt by (eresolve_tac [{\it thm}] {\it i});}
 | 
| 104 | 267 | |
| 268 | \item[\ttindexbold{bd} {\it thm} {\it i};] 
 | |
| 321 | 269 | performs \hbox{\tt by (dresolve_tac [{\it thm}] {\it i});}
 | 
| 104 | 270 | |
| 271 | \item[\ttindexbold{brs} {\it thms} {\it i};] 
 | |
| 321 | 272 | performs \hbox{\tt by (resolve_tac {\it thms} {\it i});}
 | 
| 104 | 273 | |
| 274 | \item[\ttindexbold{bes} {\it thms} {\it i};] 
 | |
| 321 | 275 | performs \hbox{\tt by (eresolve_tac {\it thms} {\it i});}
 | 
| 104 | 276 | |
| 277 | \item[\ttindexbold{bds} {\it thms} {\it i};] 
 | |
| 321 | 278 | performs \hbox{\tt by (dresolve_tac {\it thms} {\it i});}
 | 
| 279 | \end{ttdescription}
 | |
| 104 | 280 | |
| 281 | ||
| 282 | \subsection{Scanning shortcuts}
 | |
| 283 | These shortcuts scan for a suitable subgoal (starting from subgoal~1). | |
| 284 | They refine the first subgoal for which the tactic succeeds. Thus, they | |
| 285 | require less typing than {\tt br}, etc.  They display the selected
 | |
| 286 | subgoal's number; please watch this, for it may not be what you expect! | |
| 287 | ||
| 288 | \begin{ttbox} 
 | |
| 321 | 289 | fa : unit -> unit | 
| 290 | fr : thm -> unit | |
| 291 | fe : thm -> unit | |
| 292 | fd : thm -> unit | |
| 293 | frs : thm list -> unit | |
| 294 | fes : thm list -> unit | |
| 295 | fds : thm list -> unit | |
| 104 | 296 | \end{ttbox}
 | 
| 297 | ||
| 321 | 298 | \begin{ttdescription}
 | 
| 104 | 299 | \item[\ttindexbold{fa}();] 
 | 
| 321 | 300 | solves some subgoal by assumption. | 
| 104 | 301 | |
| 302 | \item[\ttindexbold{fr} {\it thm};] 
 | |
| 303 | refines some subgoal using \hbox{\tt resolve_tac [{\it thm}]}
 | |
| 304 | ||
| 305 | \item[\ttindexbold{fe} {\it thm};] 
 | |
| 306 | refines some subgoal using \hbox{\tt eresolve_tac [{\it thm}]}
 | |
| 307 | ||
| 308 | \item[\ttindexbold{fd} {\it thm};] 
 | |
| 309 | refines some subgoal using \hbox{\tt dresolve_tac [{\it thm}]}
 | |
| 310 | ||
| 311 | \item[\ttindexbold{frs} {\it thms};] 
 | |
| 312 | refines some subgoal using \hbox{\tt resolve_tac {\it thms}}
 | |
| 313 | ||
| 314 | \item[\ttindexbold{fes} {\it thms};] 
 | |
| 315 | refines some subgoal using \hbox{\tt eresolve_tac {\it thms}} 
 | |
| 316 | ||
| 317 | \item[\ttindexbold{fds} {\it thms};] 
 | |
| 318 | refines some subgoal using \hbox{\tt dresolve_tac {\it thms}} 
 | |
| 321 | 319 | \end{ttdescription}
 | 
| 104 | 320 | |
| 321 | \subsection{Other shortcuts}
 | |
| 322 | \begin{ttbox} 
 | |
| 323 | bw : thm -> unit | |
| 324 | bws : thm list -> unit | |
| 325 | ren : string -> int -> unit | |
| 326 | \end{ttbox}
 | |
| 321 | 327 | \begin{ttdescription}
 | 
| 104 | 328 | \item[\ttindexbold{bw} {\it def};] 
 | 
| 329 | performs \hbox{\tt by (rewrite_goals_tac [{\it def}]);}
 | |
| 330 | It unfolds definitions in the subgoals (but not the main goal), by | |
| 321 | 331 | meta-rewriting with the given definition. | 
| 332 | \index{meta-rewriting}
 | |
| 104 | 333 | |
| 334 | \item[\ttindexbold{bws}] 
 | |
| 335 | is like {\tt bw} but takes a list of definitions.
 | |
| 336 | ||
| 337 | \item[\ttindexbold{ren} {\it names} {\it i};] 
 | |
| 338 | performs \hbox{\tt by (rename_tac {\it names} {\it i});} it renames
 | |
| 332 | 339 | parameters in subgoal~$i$.  (Ignore the message {\footnotesize\tt Warning:\
 | 
| 340 | same as previous level}.) | |
| 321 | 341 | \index{parameters!renaming}
 | 
| 342 | \end{ttdescription}
 | |
| 104 | 343 | |
| 344 | ||
| 321 | 345 | \section{Executing batch proofs}
 | 
| 346 | \index{batch execution}
 | |
| 286 | 347 | \begin{ttbox}
 | 
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 348 | prove_goal : theory-> string->(thm list->tactic list)->thm | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 349 | qed_goal : string->theory-> string->(thm list->tactic list)->unit | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 350 | prove_goalw: theory->thm list->string->(thm list->tactic list)->thm | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 351 | qed_goalw : string->theory->thm list->string->(thm list->tactic list)->unit | 
| 104 | 352 | prove_goalw_cterm: thm list->Sign.cterm->(thm list->tactic list)->thm | 
| 353 | \end{ttbox}
 | |
| 321 | 354 | These batch functions create an initial proof state, then apply a tactic to | 
| 355 | it, yielding a sequence of final proof states. The head of the sequence is | |
| 104 | 356 | returned, provided it is an instance of the theorem originally proposed. | 
| 357 | The forms {\tt prove_goal}, {\tt prove_goalw} and {\tt prove_goalw_cterm}
 | |
| 321 | 358 | are analogous to {\tt goal}, {\tt goalw} and {\tt goalw_cterm}.
 | 
| 104 | 359 | |
| 360 | The tactic is specified by a function from theorem lists to tactic lists. | |
| 332 | 361 | The function is applied to the list of meta-assumptions taken from | 
| 362 | the main goal.  The resulting tactics are applied in sequence (using {\tt
 | |
| 363 | EVERY}). For example, a proof consisting of the commands | |
| 104 | 364 | \begin{ttbox} 
 | 
| 365 | val prems = goal {\it theory} {\it formula};
 | |
| 366 | by \(tac@1\); \ldots by \(tac@n\); | |
| 367 | val my_thm = result(); | |
| 368 | \end{ttbox}
 | |
| 369 | can be transformed to an expression as follows: | |
| 370 | \begin{ttbox} 
 | |
| 371 | val my_thm = prove_goal {\it theory} {\it formula}
 | |
| 372 | (fn prems=> [ \(tac@1\), \ldots, \(tac@n\) ]); | |
| 373 | \end{ttbox}
 | |
| 374 | The methods perform identical processing of the initial {\it formula} and
 | |
| 332 | 375 | the final proof state.  But {\tt prove_goal} executes the tactic as a
 | 
| 376 | atomic operation, bypassing the subgoal module; the current interactive | |
| 377 | proof is unaffected. | |
| 378 | % | |
| 321 | 379 | \begin{ttdescription}
 | 
| 380 | \item[\ttindexbold{prove_goal} {\it theory} {\it formula} {\it tacsf};] 
 | |
| 104 | 381 | executes a proof of the {\it formula\/} in the given {\it theory}, using
 | 
| 382 | the given tactic function. | |
| 383 | ||
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 384 | \item[\ttindexbold{qed_goal} $name$ $theory$ $formula$ $tacsf$;]
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 385 | acts like {\tt prove_goal} but also stores the resulting theorem in
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 386 | Isabelle's theorem database and in the ML variable $name$ (see | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 387 | \S\ref{ExtractingAndStoringTheProvedTheorem}).
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 388 | |
| 104 | 389 | \item[\ttindexbold{prove_goalw} {\it theory} {\it defs} {\it formula} 
 | 
| 321 | 390 |       {\it tacsf};]\index{meta-rewriting}
 | 
| 104 | 391 | is like {\tt prove_goal} but also applies the list of {\it defs\/} as
 | 
| 392 | meta-rewrite rules to the first subgoal and the premises. | |
| 393 | ||
| 866 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 394 | \item[\ttindexbold{qed_goalw} $name$ $theory$ $defs$ $formula$ $tacsf$;]
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 395 | is analogous to {\tt qed_goal}.
 | 
| 
2d3d020eef11
added documentation of bind_thm, qed, qed_goal, get_thm, thms_of
 clasohm parents: 
507diff
changeset | 396 | |
| 104 | 397 | \item[\ttindexbold{prove_goalw_cterm} {\it theory} {\it defs} {\it ct}
 | 
| 321 | 398 |       {\it tacsf};] 
 | 
| 399 | is a version of {\tt prove_goalw} for programming applications.  The main
 | |
| 104 | 400 | goal is supplied as a cterm, not as a string. Typically, the cterm is | 
| 286 | 401 | created from a term~$t$ as follows: | 
| 402 | \begin{ttbox}
 | |
| 403 | Sign.cterm_of (sign_of thy) \(t\) | |
| 404 | \end{ttbox}
 | |
| 104 | 405 | \index{*Sign.cterm_of}\index{*sign_of}
 | 
| 321 | 406 | \end{ttdescription}
 | 
| 104 | 407 | |
| 408 | ||
| 321 | 409 | \section{Managing multiple proofs}
 | 
| 410 | \index{proofs!managing multiple}
 | |
| 104 | 411 | You may save the current state of the subgoal module and resume work on it | 
| 412 | later. This serves two purposes. | |
| 413 | \begin{enumerate}
 | |
| 414 | \item At some point, you may be uncertain of the next step, and | |
| 415 | wish to experiment. | |
| 416 | ||
| 417 | \item During a proof, you may see that a lemma should be proved first. | |
| 418 | \end{enumerate}
 | |
| 419 | Each saved proof state consists of a list of levels; \ttindex{chop} behaves
 | |
| 420 | independently for each of the saved proofs. In addition, each saved state | |
| 421 | carries a separate \ttindex{undo} list.
 | |
| 422 | ||
| 321 | 423 | \subsection{The stack of proof states}
 | 
| 424 | \index{proofs!stacking}
 | |
| 104 | 425 | \begin{ttbox} 
 | 
| 426 | push_proof : unit -> unit | |
| 427 | pop_proof : unit -> thm list | |
| 428 | rotate_proof : unit -> thm list | |
| 429 | \end{ttbox}
 | |
| 430 | The subgoal module maintains a stack of proof states. Most subgoal | |
| 321 | 431 | commands affect only the top of the stack.  The \ttindex{goal} command {\em
 | 
| 432 | replaces\/} the top of the stack; the only command that pushes a proof on the | |
| 104 | 433 | stack is {\tt push_proof}.
 | 
| 434 | ||
| 435 | To save some point of the proof, call {\tt push_proof}.  You may now
 | |
| 321 | 436 | state a lemma using {\tt goal}, or simply continue to apply tactics.
 | 
| 104 | 437 | Later, you can return to the saved point by calling {\tt pop_proof} or 
 | 
| 438 | {\tt rotate_proof}. 
 | |
| 439 | ||
| 440 | To view the entire stack, call {\tt rotate_proof} repeatedly; as it rotates
 | |
| 441 | the stack, it prints the new top element. | |
| 442 | ||
| 321 | 443 | \begin{ttdescription}
 | 
| 104 | 444 | \item[\ttindexbold{push_proof}();]  
 | 
| 445 | duplicates the top element of the stack, pushing a copy of the current | |
| 446 | proof state on to the stack. | |
| 447 | ||
| 448 | \item[\ttindexbold{pop_proof}();]  
 | |
| 449 | discards the top element of the stack. It returns the list of | |
| 332 | 450 | assumptions associated with the new proof; you should bind these to an | 
| 104 | 451 | \ML\ identifier.  They can also be obtained by calling \ttindex{premises}.
 | 
| 452 | ||
| 321 | 453 | \item[\ttindexbold{rotate_proof}();]
 | 
| 454 | \index{assumptions!of main goal}
 | |
| 104 | 455 | rotates the stack, moving the top element to the bottom. It returns the | 
| 456 | list of assumptions associated with the new proof. | |
| 321 | 457 | \end{ttdescription}
 | 
| 104 | 458 | |
| 459 | ||
| 321 | 460 | \subsection{Saving and restoring proof states}
 | 
| 461 | \index{proofs!saving and restoring}
 | |
| 104 | 462 | \begin{ttbox} 
 | 
| 463 | save_proof : unit -> proof | |
| 464 | restore_proof : proof -> thm list | |
| 465 | \end{ttbox}
 | |
| 466 | States of the subgoal module may be saved as \ML\ values of | |
| 321 | 467 | type~\mltydx{proof}, and later restored.
 | 
| 104 | 468 | |
| 321 | 469 | \begin{ttdescription}
 | 
| 104 | 470 | \item[\ttindexbold{save_proof}();]  
 | 
| 471 | returns the current state, which is on top of the stack. | |
| 472 | ||
| 321 | 473 | \item[\ttindexbold{restore_proof} {\it prf};]\index{assumptions!of main goal}
 | 
| 474 |   replaces the top of the stack by~{\it prf}.  It returns the list of
 | |
| 475 | assumptions associated with the new proof. | |
| 476 | \end{ttdescription}
 | |
| 104 | 477 | |
| 478 | ||
| 321 | 479 | \section{Debugging and inspecting}
 | 
| 480 | \index{tactics!debugging}
 | |
| 104 | 481 | These specialized operations support the debugging of tactics. They refer | 
| 321 | 482 | to the current proof state of the subgoal module. | 
| 104 | 483 | |
| 321 | 484 | \subsection{Reading and printing terms}
 | 
| 485 | \index{terms!reading of}\index{terms!printing of}\index{types!printing of}
 | |
| 104 | 486 | \begin{ttbox} 
 | 
| 487 | read : string -> term | |
| 488 | prin : term -> unit | |
| 489 | printyp : typ -> unit | |
| 490 | \end{ttbox}
 | |
| 491 | These read and print terms (or types) using the syntax associated with the | |
| 492 | proof state. | |
| 493 | ||
| 321 | 494 | \begin{ttdescription}
 | 
| 104 | 495 | \item[\ttindexbold{read} {\it string}]  
 | 
| 496 | reads the {\it string} as a term, without type checking.
 | |
| 497 | ||
| 498 | \item[\ttindexbold{prin} {\it t};]  
 | |
| 499 | prints the term~$t$ at the terminal. | |
| 500 | ||
| 501 | \item[\ttindexbold{printyp} {\it T};]  
 | |
| 502 | prints the type~$T$ at the terminal. | |
| 321 | 503 | \end{ttdescription}
 | 
| 104 | 504 | |
| 321 | 505 | \subsection{Inspecting the proof state}
 | 
| 506 | \index{proofs!inspecting the state}
 | |
| 104 | 507 | \begin{ttbox} 
 | 
| 508 | topthm : unit -> thm | |
| 509 | getgoal : int -> term | |
| 510 | gethyps : int -> thm list | |
| 511 | \end{ttbox}
 | |
| 512 | ||
| 321 | 513 | \begin{ttdescription}
 | 
| 104 | 514 | \item[\ttindexbold{topthm}()]  
 | 
| 515 | returns the proof state as an Isabelle theorem.  This is what \ttindex{by}
 | |
| 516 | would supply to a tactic at this point. It omits the post-processing of | |
| 517 | \ttindex{result} and \ttindex{uresult}.
 | |
| 518 | ||
| 519 | \item[\ttindexbold{getgoal} {\it i}]  
 | |
| 520 | returns subgoal~$i$ of the proof state, as a term. You may print | |
| 521 | this using {\tt prin}, though you may have to examine the internal
 | |
| 522 | data structure in order to locate the problem! | |
| 523 | ||
| 321 | 524 | \item[\ttindexbold{gethyps} {\it i}]
 | 
| 525 | returns the hypotheses of subgoal~$i$ as meta-level assumptions. In | |
| 526 | these theorems, the subgoal's parameters become free variables. This | |
| 527 |   command is supplied for debugging uses of \ttindex{METAHYPS}.
 | |
| 528 | \end{ttdescription}
 | |
| 104 | 529 | |
| 321 | 530 | \subsection{Filtering lists of rules}
 | 
| 104 | 531 | \begin{ttbox} 
 | 
| 532 | filter_goal: (term*term->bool) -> thm list -> int -> thm list | |
| 533 | \end{ttbox}
 | |
| 534 | ||
| 321 | 535 | \begin{ttdescription}
 | 
| 104 | 536 | \item[\ttindexbold{filter_goal} {\it could} {\it ths} {\it i}] 
 | 
| 537 | applies \hbox{\tt filter_thms {\it could}} to subgoal~$i$ of the proof
 | |
| 538 | state and returns the list of theorems that survive the filtering. | |
| 321 | 539 | \end{ttdescription}
 | 
| 104 | 540 | |
| 541 | \index{subgoal module|)}
 | |
| 542 | \index{proofs|)}
 |