| author | wenzelm | 
| Fri, 15 Oct 2021 19:25:31 +0200 | |
| changeset 74525 | c960bfcb91db | 
| parent 73765 | ebaed09ce06e | 
| permissions | -rw-r--r-- | 
| 61656 | 1  | 
(*:maxLineLen=78:*)  | 
2  | 
||
| 29755 | 3  | 
theory Tactic  | 
4  | 
imports Base  | 
|
5  | 
begin  | 
|
| 18537 | 6  | 
|
| 58618 | 7  | 
chapter \<open>Tactical reasoning\<close>  | 
| 18537 | 8  | 
|
| 61854 | 9  | 
text \<open>  | 
10  | 
Tactical reasoning works by refining an initial claim in a backwards  | 
|
11  | 
fashion, until a solved form is reached. A \<open>goal\<close> consists of several  | 
|
12  | 
subgoals that need to be solved in order to achieve the main statement; zero  | 
|
13  | 
subgoals means that the proof may be finished. A \<open>tactic\<close> is a refinement  | 
|
14  | 
operation that maps a goal to a lazy sequence of potential successors. A  | 
|
15  | 
\<open>tactical\<close> is a combinator for composing tactics.  | 
|
16  | 
\<close>  | 
|
| 18537 | 17  | 
|
18  | 
||
| 58618 | 19  | 
section \<open>Goals \label{sec:tactical-goals}\<close>
 | 
| 18537 | 20  | 
|
| 58618 | 21  | 
text \<open>  | 
| 61854 | 22  | 
Isabelle/Pure represents a goal as a theorem stating that the subgoals imply  | 
23  | 
the main goal: \<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C\<close>. The outermost goal structure is that of  | 
|
24  | 
a Horn Clause: i.e.\ an iterated implication without any quantifiers\<^footnote>\<open>Recall  | 
|
25  | 
that outermost \<open>\<And>x. \<phi>[x]\<close> is always represented via schematic variables in  | 
|
26  | 
the body: \<open>\<phi>[?x]\<close>. These variables may get instantiated during the course of  | 
|
27  | 
reasoning.\<close>. For \<open>n = 0\<close> a goal is called ``solved''.  | 
|
| 18537 | 28  | 
|
| 61854 | 29  | 
The structure of each subgoal \<open>A\<^sub>i\<close> is that of a general Hereditary Harrop  | 
30  | 
Formula \<open>\<And>x\<^sub>1 \<dots> \<And>x\<^sub>k. H\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> H\<^sub>m \<Longrightarrow> B\<close>. Here \<open>x\<^sub>1, \<dots>, x\<^sub>k\<close> are goal  | 
|
31  | 
parameters, i.e.\ arbitrary-but-fixed entities of certain types, and \<open>H\<^sub>1,  | 
|
32  | 
\<dots>, H\<^sub>m\<close> are goal hypotheses, i.e.\ facts that may be assumed locally.  | 
|
33  | 
Together, this forms the goal context of the conclusion \<open>B\<close> to be  | 
|
34  | 
established. The goal hypotheses may be again arbitrary Hereditary Harrop  | 
|
35  | 
Formulas, although the level of nesting rarely exceeds 1--2 in practice.  | 
|
| 18537 | 36  | 
|
| 61854 | 37  | 
The main conclusion \<open>C\<close> is internally marked as a protected proposition,  | 
38  | 
which is represented explicitly by the notation \<open>#C\<close> here. This ensures that  | 
|
39  | 
the decomposition into subgoals and main conclusion is well-defined for  | 
|
40  | 
arbitrarily structured claims.  | 
|
| 18537 | 41  | 
|
| 61416 | 42  | 
\<^medskip>  | 
| 61854 | 43  | 
Basic goal management is performed via the following Isabelle/Pure rules:  | 
| 18537 | 44  | 
|
45  | 
\[  | 
|
| 61493 | 46  | 
  \infer[\<open>(init)\<close>]{\<open>C \<Longrightarrow> #C\<close>}{} \qquad
 | 
47  | 
  \infer[\<open>(finish)\<close>]{\<open>C\<close>}{\<open>#C\<close>}
 | 
|
| 18537 | 48  | 
\]  | 
49  | 
||
| 61416 | 50  | 
\<^medskip>  | 
| 61854 | 51  | 
The following low-level variants admit general reasoning with protected  | 
52  | 
propositions:  | 
|
| 18537 | 53  | 
|
54  | 
\[  | 
|
| 61493 | 55  | 
  \infer[\<open>(protect n)\<close>]{\<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C\<close>}{\<open>A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C\<close>}
 | 
| 52456 | 56  | 
\]  | 
57  | 
\[  | 
|
| 61493 | 58  | 
  \infer[\<open>(conclude)\<close>]{\<open>A \<Longrightarrow> \<dots> \<Longrightarrow> C\<close>}{\<open>A \<Longrightarrow> \<dots> \<Longrightarrow> #C\<close>}
 | 
| 18537 | 59  | 
\]  | 
| 58618 | 60  | 
\<close>  | 
| 18537 | 61  | 
|
| 58618 | 62  | 
text %mlref \<open>  | 
| 18537 | 63  | 
  \begin{mldecls}
 | 
| 73764 | 64  | 
  @{define_ML Goal.init: "cterm -> thm"} \\
 | 
65  | 
  @{define_ML Goal.finish: "Proof.context -> thm -> thm"} \\
 | 
|
66  | 
  @{define_ML Goal.protect: "int -> thm -> thm"} \\
 | 
|
67  | 
  @{define_ML Goal.conclude: "thm -> thm"} \\
 | 
|
| 18537 | 68  | 
  \end{mldecls}
 | 
69  | 
||
| 69597 | 70  | 
\<^descr> \<^ML>\<open>Goal.init\<close>~\<open>C\<close> initializes a tactical goal from the well-formed  | 
| 61854 | 71  | 
proposition \<open>C\<close>.  | 
| 18537 | 72  | 
|
| 69597 | 73  | 
\<^descr> \<^ML>\<open>Goal.finish\<close>~\<open>ctxt thm\<close> checks whether theorem \<open>thm\<close> is a solved  | 
| 61854 | 74  | 
goal (no subgoals), and concludes the result by removing the goal  | 
75  | 
protection. The context is only required for printing error messages.  | 
|
| 18537 | 76  | 
|
| 69597 | 77  | 
\<^descr> \<^ML>\<open>Goal.protect\<close>~\<open>n thm\<close> protects the statement of theorem \<open>thm\<close>. The  | 
| 61854 | 78  | 
parameter \<open>n\<close> indicates the number of premises to be retained.  | 
| 18537 | 79  | 
|
| 69597 | 80  | 
\<^descr> \<^ML>\<open>Goal.conclude\<close>~\<open>thm\<close> removes the goal protection, even if there are  | 
| 61854 | 81  | 
pending subgoals.  | 
| 58618 | 82  | 
\<close>  | 
| 18537 | 83  | 
|
84  | 
||
| 58618 | 85  | 
section \<open>Tactics\label{sec:tactics}\<close>
 | 
| 18537 | 86  | 
|
| 61854 | 87  | 
text \<open>  | 
88  | 
A \<open>tactic\<close> is a function \<open>goal \<rightarrow> goal\<^sup>*\<^sup>*\<close> that maps a given goal state  | 
|
89  | 
  (represented as a theorem, cf.\ \secref{sec:tactical-goals}) to a lazy
 | 
|
90  | 
sequence of potential successor states. The underlying sequence  | 
|
91  | 
implementation is lazy both in head and tail, and is purely functional in  | 
|
92  | 
\<^emph>\<open>not\<close> supporting memoing.\<^footnote>\<open>The lack of memoing and the strict nature of ML  | 
|
93  | 
requires some care when working with low-level sequence operations, to avoid  | 
|
94  | 
duplicate or premature evaluation of results. It also means that modified  | 
|
95  | 
runtime behavior, such as timeout, is very hard to achieve for general  | 
|
96  | 
tactics.\<close>  | 
|
| 18537 | 97  | 
|
| 61854 | 98  | 
An \<^emph>\<open>empty result sequence\<close> means that the tactic has failed: in a compound  | 
99  | 
tactic expression other tactics might be tried instead, or the whole  | 
|
100  | 
refinement step might fail outright, producing a toplevel error message in  | 
|
101  | 
the end. When implementing tactics from scratch, one should take care to  | 
|
102  | 
observe the basic protocol of mapping regular error conditions to an empty  | 
|
103  | 
result; only serious faults should emerge as exceptions.  | 
|
| 28781 | 104  | 
|
| 61854 | 105  | 
By enumerating \<^emph>\<open>multiple results\<close>, a tactic can easily express the  | 
106  | 
potential outcome of an internal search process. There are also combinators  | 
|
107  | 
for building proof tools that involve search systematically, see also  | 
|
108  | 
  \secref{sec:tacticals}.
 | 
|
| 28781 | 109  | 
|
| 61416 | 110  | 
\<^medskip>  | 
| 61854 | 111  | 
As explained before, a goal state essentially consists of a list of subgoals  | 
112  | 
that imply the main goal (conclusion). Tactics may operate on all subgoals  | 
|
113  | 
or on a particularly specified subgoal, but must not change the main  | 
|
114  | 
conclusion (apart from instantiating schematic goal variables).  | 
|
| 18537 | 115  | 
|
| 61854 | 116  | 
Tactics with explicit \<^emph>\<open>subgoal addressing\<close> are of the form \<open>int \<rightarrow> tactic\<close>  | 
117  | 
and may be applied to a particular subgoal (counting from 1). If the subgoal  | 
|
118  | 
number is out of range, the tactic should fail with an empty result  | 
|
119  | 
sequence, but must not raise an exception!  | 
|
| 28781 | 120  | 
|
| 61854 | 121  | 
Operating on a particular subgoal means to replace it by an interval of zero  | 
122  | 
or more subgoals in the same place; other subgoals must not be affected,  | 
|
123  | 
apart from instantiating schematic variables ranging over the whole goal  | 
|
124  | 
state.  | 
|
| 28781 | 125  | 
|
| 61854 | 126  | 
A common pattern of composing tactics with subgoal addressing is to try the  | 
127  | 
first one, and then the second one only if the subgoal has not been solved  | 
|
128  | 
yet. Special care is required here to avoid bumping into unrelated subgoals  | 
|
129  | 
that happen to come after the original subgoal. Assuming that there is only  | 
|
130  | 
a single initial subgoal is a very common error when implementing tactics!  | 
|
| 28782 | 131  | 
|
| 61854 | 132  | 
Tactics with internal subgoal addressing should expose the subgoal index as  | 
133  | 
\<open>int\<close> argument in full generality; a hardwired subgoal 1 is not acceptable.  | 
|
| 28781 | 134  | 
|
| 61416 | 135  | 
\<^medskip>  | 
| 61854 | 136  | 
The main well-formedness conditions for proper tactics are summarized as  | 
137  | 
follows.  | 
|
| 28781 | 138  | 
|
| 61854 | 139  | 
\<^item> General tactic failure is indicated by an empty result, only serious  | 
140  | 
faults may produce an exception.  | 
|
| 28781 | 141  | 
|
| 61854 | 142  | 
\<^item> The main conclusion must not be changed, apart from instantiating  | 
143  | 
schematic variables.  | 
|
| 28781 | 144  | 
|
| 61854 | 145  | 
\<^item> A tactic operates either uniformly on all subgoals, or specifically on a  | 
146  | 
selected subgoal (without bumping into unrelated subgoals).  | 
|
| 28781 | 147  | 
|
| 61854 | 148  | 
\<^item> Range errors in subgoal addressing produce an empty result.  | 
| 28782 | 149  | 
|
| 61854 | 150  | 
Some of these conditions are checked by higher-level goal infrastructure  | 
151  | 
  (\secref{sec:struct-goals}); others are not checked explicitly, and
 | 
|
152  | 
violating them merely results in ill-behaved tactics experienced by the user  | 
|
153  | 
(e.g.\ tactics that insist in being applicable only to singleton goals, or  | 
|
| 69597 | 154  | 
prevent composition via standard tacticals such as \<^ML>\<open>REPEAT\<close>).  | 
| 58618 | 155  | 
\<close>  | 
| 28782 | 156  | 
|
| 58618 | 157  | 
text %mlref \<open>  | 
| 28782 | 158  | 
  \begin{mldecls}
 | 
| 73764 | 159  | 
  @{define_ML_type tactic = "thm -> thm Seq.seq"} \\
 | 
160  | 
  @{define_ML no_tac: tactic} \\
 | 
|
161  | 
  @{define_ML all_tac: tactic} \\
 | 
|
162  | 
  @{define_ML print_tac: "Proof.context -> string -> tactic"} \\[1ex]
 | 
|
163  | 
  @{define_ML PRIMITIVE: "(thm -> thm) -> tactic"} \\[1ex]
 | 
|
164  | 
  @{define_ML SUBGOAL: "(term * int -> tactic) -> int -> tactic"} \\
 | 
|
165  | 
  @{define_ML CSUBGOAL: "(cterm * int -> tactic) -> int -> tactic"} \\
 | 
|
166  | 
  @{define_ML SELECT_GOAL: "tactic -> int -> tactic"} \\
 | 
|
167  | 
  @{define_ML PREFER_GOAL: "tactic -> int -> tactic"} \\
 | 
|
| 28782 | 168  | 
  \end{mldecls}
 | 
169  | 
||
| 69597 | 170  | 
\<^descr> Type \<^ML_type>\<open>tactic\<close> represents tactics. The well-formedness conditions  | 
| 63680 | 171  | 
described above need to be observed. See also \<^file>\<open>~~/src/Pure/General/seq.ML\<close>  | 
172  | 
for the underlying implementation of lazy sequences.  | 
|
| 28782 | 173  | 
|
| 69597 | 174  | 
\<^descr> Type \<^ML_type>\<open>int -> tactic\<close> represents tactics with explicit subgoal  | 
| 61854 | 175  | 
addressing, with well-formedness conditions as described above.  | 
| 28782 | 176  | 
|
| 69597 | 177  | 
\<^descr> \<^ML>\<open>no_tac\<close> is a tactic that always fails, returning the empty sequence.  | 
| 28783 | 178  | 
|
| 69597 | 179  | 
\<^descr> \<^ML>\<open>all_tac\<close> is a tactic that always succeeds, returning a singleton  | 
| 61854 | 180  | 
sequence with unchanged goal state.  | 
| 28783 | 181  | 
|
| 69597 | 182  | 
\<^descr> \<^ML>\<open>print_tac\<close>~\<open>ctxt message\<close> is like \<^ML>\<open>all_tac\<close>, but prints a message  | 
| 61854 | 183  | 
together with the goal state on the tracing channel.  | 
| 28783 | 184  | 
|
| 69597 | 185  | 
\<^descr> \<^ML>\<open>PRIMITIVE\<close>~\<open>rule\<close> turns a primitive inference rule into a tactic with  | 
186  | 
unique result. Exception \<^ML>\<open>THM\<close> is considered a regular tactic failure  | 
|
| 61854 | 187  | 
and produces an empty result; other exceptions are passed through.  | 
| 28782 | 188  | 
|
| 69597 | 189  | 
\<^descr> \<^ML>\<open>SUBGOAL\<close>~\<open>(fn (subgoal, i) => tactic)\<close> is the most basic form to  | 
| 61854 | 190  | 
produce a tactic with subgoal addressing. The given abstraction over the  | 
191  | 
subgoal term and subgoal number allows to peek at the relevant information  | 
|
192  | 
of the full goal state. The subgoal range is checked as required above.  | 
|
| 28782 | 193  | 
|
| 69597 | 194  | 
\<^descr> \<^ML>\<open>CSUBGOAL\<close> is similar to \<^ML>\<open>SUBGOAL\<close>, but passes the subgoal as  | 
195  | 
\<^ML_type>\<open>cterm\<close> instead of raw \<^ML_type>\<open>term\<close>. This avoids expensive  | 
|
| 61854 | 196  | 
re-certification in situations where the subgoal is used directly for  | 
197  | 
primitive inferences.  | 
|
| 28782 | 198  | 
|
| 69597 | 199  | 
\<^descr> \<^ML>\<open>SELECT_GOAL\<close>~\<open>tac i\<close> confines a tactic to the specified subgoal \<open>i\<close>.  | 
| 61854 | 200  | 
This rearranges subgoals and the main goal protection  | 
201  | 
  (\secref{sec:tactical-goals}), while retaining the syntactic context of the
 | 
|
202  | 
overall goal state (concerning schematic variables etc.).  | 
|
| 52463 | 203  | 
|
| 69597 | 204  | 
\<^descr> \<^ML>\<open>PREFER_GOAL\<close>~\<open>tac i\<close> rearranges subgoals to put \<open>i\<close> in front. This is  | 
205  | 
similar to \<^ML>\<open>SELECT_GOAL\<close>, but without changing the main goal protection.  | 
|
| 58618 | 206  | 
\<close>  | 
| 18537 | 207  | 
|
208  | 
||
| 58618 | 209  | 
subsection \<open>Resolution and assumption tactics \label{sec:resolve-assume-tac}\<close>
 | 
| 28783 | 210  | 
|
| 61854 | 211  | 
text \<open>  | 
212  | 
\<^emph>\<open>Resolution\<close> is the most basic mechanism for refining a subgoal using a  | 
|
213  | 
theorem as object-level rule. \<^emph>\<open>Elim-resolution\<close> is particularly suited for  | 
|
214  | 
elimination rules: it resolves with a rule, proves its first premise by  | 
|
215  | 
assumption, and finally deletes that assumption from any new subgoals.  | 
|
216  | 
\<^emph>\<open>Destruct-resolution\<close> is like elim-resolution, but the given destruction  | 
|
217  | 
rules are first turned into canonical elimination format.  | 
|
218  | 
\<^emph>\<open>Forward-resolution\<close> is like destruct-resolution, but without deleting the  | 
|
219  | 
selected assumption. The \<open>r/e/d/f\<close> naming convention is maintained for  | 
|
220  | 
several different kinds of resolution rules and tactics.  | 
|
| 28783 | 221  | 
|
| 61854 | 222  | 
Assumption tactics close a subgoal by unifying some of its premises against  | 
223  | 
its conclusion.  | 
|
| 28783 | 224  | 
|
| 61416 | 225  | 
\<^medskip>  | 
| 61854 | 226  | 
All the tactics in this section operate on a subgoal designated by a  | 
227  | 
positive integer. Other subgoals might be affected indirectly, due to  | 
|
228  | 
instantiation of schematic variables.  | 
|
| 28783 | 229  | 
|
| 61854 | 230  | 
There are various sources of non-determinism, the tactic result sequence  | 
231  | 
enumerates all possibilities of the following choices (if applicable):  | 
|
| 28783 | 232  | 
|
| 61854 | 233  | 
\<^enum> selecting one of the rules given as argument to the tactic;  | 
234  | 
||
235  | 
\<^enum> selecting a subgoal premise to eliminate, unifying it against the first  | 
|
236  | 
premise of the rule;  | 
|
| 28783 | 237  | 
|
| 61854 | 238  | 
\<^enum> unifying the conclusion of the subgoal to the conclusion of the rule.  | 
| 28783 | 239  | 
|
| 61854 | 240  | 
Recall that higher-order unification may produce multiple results that are  | 
241  | 
enumerated here.  | 
|
| 58618 | 242  | 
\<close>  | 
| 28783 | 243  | 
|
| 58618 | 244  | 
text %mlref \<open>  | 
| 28783 | 245  | 
  \begin{mldecls}
 | 
| 73764 | 246  | 
  @{define_ML resolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | 
247  | 
  @{define_ML eresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | 
|
248  | 
  @{define_ML dresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | 
|
249  | 
  @{define_ML forward_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | 
|
250  | 
  @{define_ML biresolve_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\[1ex]
 | 
|
251  | 
  @{define_ML assume_tac: "Proof.context -> int -> tactic"} \\
 | 
|
252  | 
  @{define_ML eq_assume_tac: "int -> tactic"} \\[1ex]
 | 
|
253  | 
  @{define_ML match_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | 
|
254  | 
  @{define_ML ematch_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | 
|
255  | 
  @{define_ML dmatch_tac: "Proof.context -> thm list -> int -> tactic"} \\
 | 
|
256  | 
  @{define_ML bimatch_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\
 | 
|
| 28783 | 257  | 
  \end{mldecls}
 | 
258  | 
||
| 69597 | 259  | 
\<^descr> \<^ML>\<open>resolve_tac\<close>~\<open>ctxt thms i\<close> refines the goal state using the given  | 
| 61854 | 260  | 
theorems, which should normally be introduction rules. The tactic resolves a  | 
261  | 
rule's conclusion with subgoal \<open>i\<close>, replacing it by the corresponding  | 
|
262  | 
versions of the rule's premises.  | 
|
| 28783 | 263  | 
|
| 69597 | 264  | 
\<^descr> \<^ML>\<open>eresolve_tac\<close>~\<open>ctxt thms i\<close> performs elim-resolution with the given  | 
| 61854 | 265  | 
theorems, which are normally be elimination rules.  | 
| 46278 | 266  | 
|
| 69597 | 267  | 
Note that \<^ML_text>\<open>eresolve_tac ctxt [asm_rl]\<close> is equivalent to \<^ML_text>\<open>assume_tac ctxt\<close>, which facilitates mixing of assumption steps with  | 
| 46278 | 268  | 
genuine eliminations.  | 
| 28783 | 269  | 
|
| 69597 | 270  | 
\<^descr> \<^ML>\<open>dresolve_tac\<close>~\<open>ctxt thms i\<close> performs destruct-resolution with the  | 
| 61854 | 271  | 
given theorems, which should normally be destruction rules. This replaces an  | 
272  | 
assumption by the result of applying one of the rules.  | 
|
| 28783 | 273  | 
|
| 69597 | 274  | 
\<^descr> \<^ML>\<open>forward_tac\<close> is like \<^ML>\<open>dresolve_tac\<close> except that the selected  | 
| 61854 | 275  | 
assumption is not deleted. It applies a rule to an assumption, adding the  | 
276  | 
result as a new assumption.  | 
|
| 28783 | 277  | 
|
| 69597 | 278  | 
\<^descr> \<^ML>\<open>biresolve_tac\<close>~\<open>ctxt brls i\<close> refines the proof state by resolution or  | 
| 61854 | 279  | 
elim-resolution on each rule, as indicated by its flag. It affects subgoal  | 
280  | 
\<open>i\<close> of the proof state.  | 
|
| 50072 | 281  | 
|
| 61854 | 282  | 
For each pair \<open>(flag, rule)\<close>, it applies resolution if the flag is \<open>false\<close>  | 
283  | 
and elim-resolution if the flag is \<open>true\<close>. A single tactic call handles a  | 
|
284  | 
mixture of introduction and elimination rules, which is useful to organize  | 
|
285  | 
the search process systematically in proof tools.  | 
|
| 50072 | 286  | 
|
| 69597 | 287  | 
\<^descr> \<^ML>\<open>assume_tac\<close>~\<open>ctxt i\<close> attempts to solve subgoal \<open>i\<close> by assumption  | 
| 61854 | 288  | 
(modulo higher-order unification).  | 
| 28783 | 289  | 
|
| 69597 | 290  | 
\<^descr> \<^ML>\<open>eq_assume_tac\<close> is similar to \<^ML>\<open>assume_tac\<close>, but checks only for  | 
| 61854 | 291  | 
immediate \<open>\<alpha>\<close>-convertibility instead of using unification. It succeeds (with  | 
292  | 
a unique next state) if one of the assumptions is equal to the subgoal's  | 
|
293  | 
conclusion. Since it does not instantiate variables, it cannot make other  | 
|
294  | 
subgoals unprovable.  | 
|
| 28783 | 295  | 
|
| 69597 | 296  | 
\<^descr> \<^ML>\<open>match_tac\<close>, \<^ML>\<open>ematch_tac\<close>, \<^ML>\<open>dmatch_tac\<close>, and \<^ML>\<open>bimatch_tac\<close>  | 
297  | 
are similar to \<^ML>\<open>resolve_tac\<close>, \<^ML>\<open>eresolve_tac\<close>, \<^ML>\<open>dresolve_tac\<close>,  | 
|
298  | 
and \<^ML>\<open>biresolve_tac\<close>, respectively, but do not instantiate schematic  | 
|
| 61854 | 299  | 
variables in the goal state.\<^footnote>\<open>Strictly speaking, matching means to treat the  | 
300  | 
unknowns in the goal state as constants, but these tactics merely discard  | 
|
301  | 
unifiers that would update the goal state. In rare situations (where the  | 
|
302  | 
conclusion and goal state have flexible terms at the same position), the  | 
|
303  | 
tactic will fail even though an acceptable unifier exists.\<close> These tactics  | 
|
304  | 
were written for a specific application within the classical reasoner.  | 
|
| 28783 | 305  | 
|
306  | 
Flexible subgoals are not updated at will, but are left alone.  | 
|
| 58618 | 307  | 
\<close>  | 
| 28783 | 308  | 
|
309  | 
||
| 58618 | 310  | 
subsection \<open>Explicit instantiation within a subgoal context\<close>  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
311  | 
|
| 61854 | 312  | 
text \<open>  | 
313  | 
  The main resolution tactics (\secref{sec:resolve-assume-tac}) use
 | 
|
314  | 
higher-order unification, which works well in many practical situations  | 
|
315  | 
despite its daunting theoretical properties. Nonetheless, there are  | 
|
316  | 
important problem classes where unguided higher-order unification is not so  | 
|
317  | 
useful. This typically involves rules like universal elimination,  | 
|
318  | 
existential introduction, or equational substitution. Here the unification  | 
|
319  | 
problem involves fully flexible \<open>?P ?x\<close> schemes, which are hard to manage  | 
|
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
320  | 
without further hints.  | 
| 
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
321  | 
|
| 61854 | 322  | 
By providing a (small) rigid term for \<open>?x\<close> explicitly, the remaining  | 
323  | 
unification problem is to assign a (large) term to \<open>?P\<close>, according to the  | 
|
324  | 
shape of the given subgoal. This is sufficiently well-behaved in most  | 
|
325  | 
practical situations.  | 
|
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
326  | 
|
| 61416 | 327  | 
\<^medskip>  | 
| 61854 | 328  | 
Isabelle provides separate versions of the standard \<open>r/e/d/f\<close> resolution  | 
329  | 
tactics that allow to provide explicit instantiations of unknowns of the  | 
|
330  | 
given rule, wrt.\ terms that refer to the implicit context of the selected  | 
|
331  | 
subgoal.  | 
|
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
332  | 
|
| 61854 | 333  | 
An instantiation consists of a list of pairs of the form \<open>(?x, t)\<close>, where  | 
334  | 
\<open>?x\<close> is a schematic variable occurring in the given rule, and \<open>t\<close> is a term  | 
|
335  | 
from the current proof context, augmented by the local goal parameters of  | 
|
336  | 
the selected subgoal; cf.\ the \<open>focus\<close> operation described in  | 
|
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
337  | 
  \secref{sec:variables}.
 | 
| 
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
338  | 
|
| 61854 | 339  | 
Entering the syntactic context of a subgoal is a brittle operation, because  | 
340  | 
its exact form is somewhat accidental, and the choice of bound variable  | 
|
341  | 
names depends on the presence of other local and global names. Explicit  | 
|
342  | 
renaming of subgoal parameters prior to explicit instantiation might help to  | 
|
343  | 
achieve a bit more robustness.  | 
|
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
344  | 
|
| 61854 | 345  | 
Type instantiations may be given as well, via pairs like \<open>(?'a, \<tau>)\<close>. Type  | 
346  | 
instantiations are distinguished from term instantiations by the syntactic  | 
|
347  | 
form of the schematic variable. Types are instantiated before terms are.  | 
|
348  | 
Since term instantiation already performs simple type-inference, so explicit  | 
|
349  | 
type instantiations are seldom necessary.  | 
|
| 58618 | 350  | 
\<close>  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
351  | 
|
| 58618 | 352  | 
text %mlref \<open>  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
353  | 
  \begin{mldecls}
 | 
| 73764 | 354  | 
  @{define_ML Rule_Insts.res_inst_tac: "Proof.context ->
 | 
| 59780 | 355  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
356  | 
thm -> int -> tactic"} \\  | 
|
| 73764 | 357  | 
  @{define_ML Rule_Insts.eres_inst_tac: "Proof.context ->
 | 
| 59780 | 358  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
359  | 
thm -> int -> tactic"} \\  | 
|
| 73764 | 360  | 
  @{define_ML Rule_Insts.dres_inst_tac: "Proof.context ->
 | 
| 59780 | 361  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
362  | 
thm -> int -> tactic"} \\  | 
|
| 73764 | 363  | 
  @{define_ML Rule_Insts.forw_inst_tac: "Proof.context ->
 | 
| 59780 | 364  | 
((indexname * Position.T) * string) list -> (binding * string option * mixfix) list ->  | 
365  | 
thm -> int -> tactic"} \\  | 
|
| 73764 | 366  | 
  @{define_ML Rule_Insts.subgoal_tac: "Proof.context -> string ->
 | 
| 59780 | 367  | 
(binding * string option * mixfix) list -> int -> tactic"} \\  | 
| 73764 | 368  | 
  @{define_ML Rule_Insts.thin_tac: "Proof.context -> string ->
 | 
| 59780 | 369  | 
(binding * string option * mixfix) list -> int -> tactic"} \\  | 
| 73764 | 370  | 
  @{define_ML rename_tac: "string list -> int -> tactic"} \\
 | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
371  | 
  \end{mldecls}
 | 
| 
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
372  | 
|
| 69597 | 373  | 
\<^descr> \<^ML>\<open>Rule_Insts.res_inst_tac\<close>~\<open>ctxt insts thm i\<close> instantiates the rule  | 
| 61854 | 374  | 
\<open>thm\<close> with the instantiations \<open>insts\<close>, as described above, and then performs  | 
375  | 
resolution on subgoal \<open>i\<close>.  | 
|
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
376  | 
|
| 69597 | 377  | 
\<^descr> \<^ML>\<open>Rule_Insts.eres_inst_tac\<close> is like \<^ML>\<open>Rule_Insts.res_inst_tac\<close>, but  | 
| 61854 | 378  | 
performs elim-resolution.  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
379  | 
|
| 69597 | 380  | 
\<^descr> \<^ML>\<open>Rule_Insts.dres_inst_tac\<close> is like \<^ML>\<open>Rule_Insts.res_inst_tac\<close>, but  | 
| 61854 | 381  | 
performs destruct-resolution.  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
382  | 
|
| 69597 | 383  | 
\<^descr> \<^ML>\<open>Rule_Insts.forw_inst_tac\<close> is like \<^ML>\<open>Rule_Insts.dres_inst_tac\<close>  | 
| 59763 | 384  | 
except that the selected assumption is not deleted.  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
385  | 
|
| 69597 | 386  | 
\<^descr> \<^ML>\<open>Rule_Insts.subgoal_tac\<close>~\<open>ctxt \<phi> i\<close> adds the proposition \<open>\<phi>\<close> as local  | 
| 61854 | 387  | 
premise to subgoal \<open>i\<close>, and poses the same as a new subgoal \<open>i + 1\<close> (in the  | 
388  | 
original context).  | 
|
| 46271 | 389  | 
|
| 69597 | 390  | 
\<^descr> \<^ML>\<open>Rule_Insts.thin_tac\<close>~\<open>ctxt \<phi> i\<close> deletes the specified premise from  | 
| 61854 | 391  | 
subgoal \<open>i\<close>. Note that \<open>\<phi>\<close> may contain schematic variables, to abbreviate  | 
392  | 
the intended proposition; the first matching subgoal premise will be  | 
|
393  | 
deleted. Removing useless premises from a subgoal increases its readability  | 
|
394  | 
and can make search tactics run faster.  | 
|
| 46277 | 395  | 
|
| 69597 | 396  | 
\<^descr> \<^ML>\<open>rename_tac\<close>~\<open>names i\<close> renames the innermost parameters of subgoal \<open>i\<close>  | 
| 61854 | 397  | 
according to the provided \<open>names\<close> (which need to be distinct identifiers).  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
398  | 
|
| 34930 | 399  | 
|
| 61854 | 400  | 
For historical reasons, the above instantiation tactics take unparsed string  | 
401  | 
arguments, which makes them hard to use in general ML code. The slightly  | 
|
| 69597 | 402  | 
  more advanced \<^ML>\<open>Subgoal.FOCUS\<close> combinator of \secref{sec:struct-goals}
 | 
| 61854 | 403  | 
allows to refer to internal goal structure with explicit context management.  | 
| 58618 | 404  | 
\<close>  | 
| 
28785
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
405  | 
|
| 
 
64163cddf3e6
added section "Explicit instantiation within a subgoal context";
 
wenzelm 
parents: 
28783 
diff
changeset
 | 
406  | 
|
| 58618 | 407  | 
subsection \<open>Rearranging goal states\<close>  | 
| 46274 | 408  | 
|
| 61854 | 409  | 
text \<open>  | 
410  | 
In rare situations there is a need to rearrange goal states: either the  | 
|
411  | 
overall collection of subgoals, or the local structure of a subgoal. Various  | 
|
412  | 
administrative tactics allow to operate on the concrete presentation these  | 
|
413  | 
conceptual sets of formulae.  | 
|
414  | 
\<close>  | 
|
| 46274 | 415  | 
|
| 58618 | 416  | 
text %mlref \<open>  | 
| 46274 | 417  | 
  \begin{mldecls}
 | 
| 73764 | 418  | 
  @{define_ML rotate_tac: "int -> int -> tactic"} \\
 | 
419  | 
  @{define_ML distinct_subgoals_tac: tactic} \\
 | 
|
420  | 
  @{define_ML flexflex_tac: "Proof.context -> tactic"} \\
 | 
|
| 46274 | 421  | 
  \end{mldecls}
 | 
422  | 
||
| 69597 | 423  | 
\<^descr> \<^ML>\<open>rotate_tac\<close>~\<open>n i\<close> rotates the premises of subgoal \<open>i\<close> by \<open>n\<close>  | 
| 61854 | 424  | 
positions: from right to left if \<open>n\<close> is positive, and from left to right if  | 
425  | 
\<open>n\<close> is negative.  | 
|
| 46274 | 426  | 
|
| 69597 | 427  | 
\<^descr> \<^ML>\<open>distinct_subgoals_tac\<close> removes duplicate subgoals from a proof state.  | 
| 61854 | 428  | 
This is potentially inefficient.  | 
| 46276 | 429  | 
|
| 69597 | 430  | 
\<^descr> \<^ML>\<open>flexflex_tac\<close> removes all flex-flex pairs from the proof state by  | 
| 61854 | 431  | 
applying the trivial unifier. This drastic step loses information. It is  | 
432  | 
already part of the Isar infrastructure for facts resulting from goals, and  | 
|
433  | 
rarely needs to be invoked manually.  | 
|
| 46276 | 434  | 
|
435  | 
Flex-flex constraints arise from difficult cases of higher-order  | 
|
| 69597 | 436  | 
unification. To prevent this, use \<^ML>\<open>Rule_Insts.res_inst_tac\<close> to  | 
| 61854 | 437  | 
instantiate some variables in a rule. Normally flex-flex constraints can be  | 
438  | 
ignored; they often disappear as unknowns get instantiated.  | 
|
| 58618 | 439  | 
\<close>  | 
| 46274 | 440  | 
|
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
441  | 
|
| 58618 | 442  | 
subsection \<open>Raw composition: resolution without lifting\<close>  | 
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
443  | 
|
| 58618 | 444  | 
text \<open>  | 
| 61854 | 445  | 
Raw composition of two rules means resolving them without prior lifting or  | 
446  | 
renaming of unknowns. This low-level operation, which underlies the  | 
|
447  | 
resolution tactics, may occasionally be useful for special effects.  | 
|
448  | 
Schematic variables are not renamed by default, so beware of clashes!  | 
|
| 58618 | 449  | 
\<close>  | 
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
450  | 
|
| 58618 | 451  | 
text %mlref \<open>  | 
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
452  | 
  \begin{mldecls}
 | 
| 73764 | 453  | 
  @{define_ML compose_tac: "Proof.context -> (bool * thm * int) -> int -> tactic"} \\
 | 
454  | 
  @{define_ML Drule.compose: "thm * int * thm -> thm"} \\
 | 
|
455  | 
  @{define_ML_infix COMP: "thm * thm -> thm"} \\
 | 
|
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
456  | 
  \end{mldecls}
 | 
| 
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
457  | 
|
| 69597 | 458  | 
\<^descr> \<^ML>\<open>compose_tac\<close>~\<open>ctxt (flag, rule, m) i\<close> refines subgoal \<open>i\<close> using  | 
| 61854 | 459  | 
\<open>rule\<close>, without lifting. The \<open>rule\<close> is taken to have the form \<open>\<psi>\<^sub>1 \<Longrightarrow> \<dots> \<psi>\<^sub>m \<Longrightarrow>  | 
460  | 
\<psi>\<close>, where \<open>\<psi>\<close> need not be atomic; thus \<open>m\<close> determines the number of new  | 
|
461  | 
subgoals. If \<open>flag\<close> is \<open>true\<close> then it performs elim-resolution --- it solves  | 
|
462  | 
the first premise of \<open>rule\<close> by assumption and deletes that assumption.  | 
|
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
463  | 
|
| 69597 | 464  | 
\<^descr> \<^ML>\<open>Drule.compose\<close>~\<open>(thm\<^sub>1, i, thm\<^sub>2)\<close> uses \<open>thm\<^sub>1\<close>, regarded as an  | 
| 61854 | 465  | 
atomic formula, to solve premise \<open>i\<close> of \<open>thm\<^sub>2\<close>. Let \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> be  | 
466  | 
\<open>\<psi>\<close> and \<open>\<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>\<close>. The unique \<open>s\<close> that unifies \<open>\<psi>\<close> and \<open>\<phi>\<^sub>i\<close> yields  | 
|
467  | 
the theorem \<open>(\<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>i\<^sub>-\<^sub>1 \<Longrightarrow> \<phi>\<^sub>i\<^sub>+\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<phi>)s\<close>. Multiple results are  | 
|
| 69597 | 468  | 
considered as error (exception \<^ML>\<open>THM\<close>).  | 
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
469  | 
|
| 61854 | 470  | 
\<^descr> \<open>thm\<^sub>1 COMP thm\<^sub>2\<close> is the same as \<open>Drule.compose (thm\<^sub>1, 1, thm\<^sub>2)\<close>.  | 
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
471  | 
|
| 
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
472  | 
|
| 
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
473  | 
  \begin{warn}
 | 
| 61854 | 474  | 
These low-level operations are stepping outside the structure imposed by  | 
475  | 
regular rule resolution. Used without understanding of the consequences,  | 
|
476  | 
they may produce results that cause problems with standard rules and tactics  | 
|
477  | 
later on.  | 
|
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
478  | 
  \end{warn}
 | 
| 58618 | 479  | 
\<close>  | 
| 
50074
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
480  | 
|
| 
 
0b02aaf7c7c5
some coverage of "resolution without lifting", which should be normally avoided;
 
wenzelm 
parents: 
50072 
diff
changeset
 | 
481  | 
|
| 58618 | 482  | 
section \<open>Tacticals \label{sec:tacticals}\<close>
 | 
| 18537 | 483  | 
|
| 61854 | 484  | 
text \<open>  | 
485  | 
A \<^emph>\<open>tactical\<close> is a functional combinator for building up complex tactics  | 
|
486  | 
from simpler ones. Common tacticals perform sequential composition,  | 
|
487  | 
disjunctive choice, iteration, or goal addressing. Various search strategies  | 
|
488  | 
may be expressed via tacticals.  | 
|
| 58618 | 489  | 
\<close>  | 
| 46258 | 490  | 
|
491  | 
||
| 58618 | 492  | 
subsection \<open>Combining tactics\<close>  | 
| 46258 | 493  | 
|
| 61854 | 494  | 
text \<open>  | 
495  | 
Sequential composition and alternative choices are the most basic ways to  | 
|
496  | 
combine tactics, similarly to ``\<^verbatim>\<open>,\<close>'' and ``\<^verbatim>\<open>|\<close>'' in Isar method notation.  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
497  | 
This corresponds to \<^ML_infix>\<open>THEN\<close> and \<^ML_infix>\<open>ORELSE\<close> in ML, but there  | 
| 61854 | 498  | 
are further possibilities for fine-tuning alternation of tactics such as  | 
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
499  | 
\<^ML_infix>\<open>APPEND\<close>. Further details become visible in ML due to explicit  | 
| 46262 | 500  | 
subgoal addressing.  | 
| 58618 | 501  | 
\<close>  | 
| 46258 | 502  | 
|
| 58618 | 503  | 
text %mlref \<open>  | 
| 46258 | 504  | 
  \begin{mldecls}
 | 
| 73764 | 505  | 
  @{define_ML_infix "THEN": "tactic * tactic -> tactic"} \\
 | 
506  | 
  @{define_ML_infix "ORELSE": "tactic * tactic -> tactic"} \\
 | 
|
507  | 
  @{define_ML_infix "APPEND": "tactic * tactic -> tactic"} \\
 | 
|
508  | 
  @{define_ML "EVERY": "tactic list -> tactic"} \\
 | 
|
509  | 
  @{define_ML "FIRST": "tactic list -> tactic"} \\[0.5ex]
 | 
|
| 46258 | 510  | 
|
| 73764 | 511  | 
  @{define_ML_infix "THEN'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
 | 
512  | 
  @{define_ML_infix "ORELSE'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
 | 
|
513  | 
  @{define_ML_infix "APPEND'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\
 | 
|
514  | 
  @{define_ML "EVERY'": "('a -> tactic) list -> 'a -> tactic"} \\
 | 
|
515  | 
  @{define_ML "FIRST'": "('a -> tactic) list -> 'a -> tactic"} \\
 | 
|
| 46258 | 516  | 
  \end{mldecls}
 | 
517  | 
||
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
518  | 
\<^descr> \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>2\<close> is the sequential composition of \<open>tac\<^sub>1\<close> and  | 
| 61854 | 519  | 
\<open>tac\<^sub>2\<close>. Applied to a goal state, it returns all states reachable in two  | 
520  | 
steps by applying \<open>tac\<^sub>1\<close> followed by \<open>tac\<^sub>2\<close>. First, it applies \<open>tac\<^sub>1\<close> to  | 
|
521  | 
the goal state, getting a sequence of possible next states; then, it applies  | 
|
522  | 
\<open>tac\<^sub>2\<close> to each of these and concatenates the results to produce again one  | 
|
523  | 
flat sequence of states.  | 
|
| 46258 | 524  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
525  | 
\<^descr> \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac\<^sub>2\<close> makes a choice between \<open>tac\<^sub>1\<close> and  | 
| 61854 | 526  | 
\<open>tac\<^sub>2\<close>. Applied to a state, it tries \<open>tac\<^sub>1\<close> and returns the result if  | 
527  | 
non-empty; if \<open>tac\<^sub>1\<close> fails then it uses \<open>tac\<^sub>2\<close>. This is a deterministic  | 
|
528  | 
choice: if \<open>tac\<^sub>1\<close> succeeds then \<open>tac\<^sub>2\<close> is excluded from the result.  | 
|
| 46258 | 529  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
530  | 
\<^descr> \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>APPEND\<close>~\<open>tac\<^sub>2\<close> concatenates the possible results of  | 
| 
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
531  | 
\<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close>. Unlike \<^ML_infix>\<open>ORELSE\<close> there is \<^emph>\<open>no commitment\<close> to  | 
| 
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
532  | 
either tactic, so \<^ML_infix>\<open>APPEND\<close> helps to avoid incompleteness during  | 
| 61854 | 533  | 
search, at the cost of potential inefficiencies.  | 
| 39852 | 534  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
535  | 
\<^descr> \<^ML>\<open>EVERY\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>n\<close>. Note that \<^ML>\<open>EVERY []\<close> is the same as  | 
| 69597 | 536  | 
\<^ML>\<open>all_tac\<close>: it always succeeds.  | 
| 46258 | 537  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
538  | 
\<^descr> \<^ML>\<open>FIRST\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>n]\<close> abbreviates \<open>tac\<^sub>1\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac\<^sub>n\<close>. Note that \<^ML>\<open>FIRST []\<close> is the  | 
| 69597 | 539  | 
same as \<^ML>\<open>no_tac\<close>: it always fails.  | 
| 46258 | 540  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
541  | 
\<^descr> \<^ML_infix>\<open>THEN'\<close> is the lifted version of \<^ML_infix>\<open>THEN\<close>, for tactics  | 
| 
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
542  | 
with explicit subgoal addressing. So \<open>(tac\<^sub>1\<close>~\<^ML_infix>\<open>THEN'\<close>~\<open>tac\<^sub>2) i\<close> is  | 
| 
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
543  | 
the same as \<open>(tac\<^sub>1 i\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>2 i)\<close>.  | 
| 46258 | 544  | 
|
| 46264 | 545  | 
The other primed tacticals work analogously.  | 
| 58618 | 546  | 
\<close>  | 
| 30272 | 547  | 
|
| 46259 | 548  | 
|
| 58618 | 549  | 
subsection \<open>Repetition tacticals\<close>  | 
| 46259 | 550  | 
|
| 61854 | 551  | 
text \<open>  | 
552  | 
These tacticals provide further control over repetition of tactics, beyond  | 
|
553  | 
the stylized forms of ``\<^verbatim>\<open>?\<close>'' and ``\<^verbatim>\<open>+\<close>'' in Isar method expressions.  | 
|
554  | 
\<close>  | 
|
| 46259 | 555  | 
|
| 58618 | 556  | 
text %mlref \<open>  | 
| 46259 | 557  | 
  \begin{mldecls}
 | 
| 73764 | 558  | 
  @{define_ML "TRY": "tactic -> tactic"} \\
 | 
559  | 
  @{define_ML "REPEAT": "tactic -> tactic"} \\
 | 
|
560  | 
  @{define_ML "REPEAT1": "tactic -> tactic"} \\
 | 
|
561  | 
  @{define_ML "REPEAT_DETERM": "tactic -> tactic"} \\
 | 
|
562  | 
  @{define_ML "REPEAT_DETERM_N": "int -> tactic -> tactic"} \\
 | 
|
| 46259 | 563  | 
  \end{mldecls}
 | 
564  | 
||
| 69597 | 565  | 
\<^descr> \<^ML>\<open>TRY\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and returns the resulting  | 
| 61854 | 566  | 
sequence, if non-empty; otherwise it returns the original state. Thus, it  | 
567  | 
applies \<open>tac\<close> at most once.  | 
|
| 46259 | 568  | 
|
| 61854 | 569  | 
Note that for tactics with subgoal addressing, the combinator can be applied  | 
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
570  | 
via functional composition: \<^ML>\<open>TRY\<close>~\<^ML_infix>\<open>o\<close>~\<open>tac\<close>. There is no need  | 
| 61854 | 571  | 
for \<^verbatim>\<open>TRY'\<close>.  | 
| 46259 | 572  | 
|
| 69597 | 573  | 
\<^descr> \<^ML>\<open>REPEAT\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and, recursively, to  | 
| 61854 | 574  | 
each element of the resulting sequence. The resulting sequence consists of  | 
575  | 
those states that make \<open>tac\<close> fail. Thus, it applies \<open>tac\<close> as many times as  | 
|
| 46259 | 576  | 
possible (including zero times), and allows backtracking over each  | 
| 69597 | 577  | 
invocation of \<open>tac\<close>. \<^ML>\<open>REPEAT\<close> is more general than \<^ML>\<open>REPEAT_DETERM\<close>,  | 
| 61854 | 578  | 
but requires more space.  | 
| 46259 | 579  | 
|
| 69597 | 580  | 
\<^descr> \<^ML>\<open>REPEAT1\<close>~\<open>tac\<close> is like \<^ML>\<open>REPEAT\<close>~\<open>tac\<close> but it always applies \<open>tac\<close>  | 
| 61854 | 581  | 
at least once, failing if this is impossible.  | 
| 46259 | 582  | 
|
| 69597 | 583  | 
\<^descr> \<^ML>\<open>REPEAT_DETERM\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and,  | 
| 61854 | 584  | 
recursively, to the head of the resulting sequence. It returns the first  | 
585  | 
state to make \<open>tac\<close> fail. It is deterministic, discarding alternative  | 
|
586  | 
outcomes.  | 
|
| 46266 | 587  | 
|
| 69597 | 588  | 
\<^descr> \<^ML>\<open>REPEAT_DETERM_N\<close>~\<open>n tac\<close> is like \<^ML>\<open>REPEAT_DETERM\<close>~\<open>tac\<close> but the  | 
589  | 
number of repetitions is bound by \<open>n\<close> (where \<^ML>\<open>~1\<close> means \<open>\<infinity>\<close>).  | 
|
| 58618 | 590  | 
\<close>  | 
| 46259 | 591  | 
|
| 61854 | 592  | 
text %mlex \<open>  | 
593  | 
The basic tactics and tacticals considered above follow some algebraic laws:  | 
|
| 46259 | 594  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
595  | 
\<^item> \<^ML>\<open>all_tac\<close> is the identity element of the tactical \<^ML_infix>\<open>THEN\<close>.  | 
| 46259 | 596  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
597  | 
\<^item> \<^ML>\<open>no_tac\<close> is the identity element of \<^ML_infix>\<open>ORELSE\<close> and \<^ML_infix>\<open>APPEND\<close>. Also, it is a zero element for \<^ML_infix>\<open>THEN\<close>, which means that  | 
| 
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
598  | 
\<open>tac\<close>~\<^ML_infix>\<open>THEN\<close>~\<^ML>\<open>no_tac\<close> is equivalent to \<^ML>\<open>no_tac\<close>.  | 
| 46259 | 599  | 
|
| 69597 | 600  | 
\<^item> \<^ML>\<open>TRY\<close> and \<^ML>\<open>REPEAT\<close> can be expressed as (recursive) functions over  | 
| 61854 | 601  | 
more basic combinators (ignoring some internal implementation tricks):  | 
| 58618 | 602  | 
\<close>  | 
| 46259 | 603  | 
|
| 58618 | 604  | 
ML \<open>  | 
| 46259 | 605  | 
fun TRY tac = tac ORELSE all_tac;  | 
606  | 
fun REPEAT tac st = ((tac THEN REPEAT tac) ORELSE all_tac) st;  | 
|
| 58618 | 607  | 
\<close>  | 
| 46259 | 608  | 
|
| 61854 | 609  | 
text \<open>  | 
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
610  | 
If \<open>tac\<close> can return multiple outcomes then so can \<^ML>\<open>REPEAT\<close>~\<open>tac\<close>. \<^ML>\<open>REPEAT\<close> uses \<^ML_infix>\<open>ORELSE\<close> and not \<^ML_infix>\<open>APPEND\<close>, it applies \<open>tac\<close>  | 
| 61854 | 611  | 
as many times as possible in each outcome.  | 
| 46259 | 612  | 
|
613  | 
  \begin{warn}
 | 
|
| 61854 | 614  | 
Note the explicit abstraction over the goal state in the ML definition of  | 
| 69597 | 615  | 
\<^ML>\<open>REPEAT\<close>. Recursive tacticals must be coded in this awkward fashion to  | 
| 61854 | 616  | 
avoid infinite recursion of eager functional evaluation in Standard ML. The  | 
| 69597 | 617  | 
following attempt would make \<^ML>\<open>REPEAT\<close>~\<open>tac\<close> loop:  | 
| 46259 | 618  | 
  \end{warn}
 | 
| 58618 | 619  | 
\<close>  | 
| 46259 | 620  | 
|
| 59902 | 621  | 
ML_val \<open>  | 
| 46260 | 622  | 
(*BAD -- does not terminate!*)  | 
623  | 
fun REPEAT tac = (tac THEN REPEAT tac) ORELSE all_tac;  | 
|
| 58618 | 624  | 
\<close>  | 
| 46259 | 625  | 
|
| 46263 | 626  | 
|
| 58618 | 627  | 
subsection \<open>Applying tactics to subgoal ranges\<close>  | 
| 46263 | 628  | 
|
| 61854 | 629  | 
text \<open>  | 
| 69597 | 630  | 
Tactics with explicit subgoal addressing \<^ML_type>\<open>int -> tactic\<close> can be  | 
| 61854 | 631  | 
used together with tacticals that act like ``subgoal quantifiers'': guided  | 
632  | 
by success of the body tactic a certain range of subgoals is covered. Thus  | 
|
633  | 
the body tactic is applied to \<^emph>\<open>all\<close> subgoals, \<^emph>\<open>some\<close> subgoal etc.  | 
|
| 46263 | 634  | 
|
| 61854 | 635  | 
Suppose that the goal state has \<open>n \<ge> 0\<close> subgoals. Many of these tacticals  | 
636  | 
address subgoal ranges counting downwards from \<open>n\<close> towards \<open>1\<close>. This has the  | 
|
637  | 
fortunate effect that newly emerging subgoals are concatenated in the  | 
|
638  | 
result, without interfering each other. Nonetheless, there might be  | 
|
639  | 
situations where a different order is desired.  | 
|
640  | 
\<close>  | 
|
| 46263 | 641  | 
|
| 58618 | 642  | 
text %mlref \<open>  | 
| 46263 | 643  | 
  \begin{mldecls}
 | 
| 73764 | 644  | 
  @{define_ML ALLGOALS: "(int -> tactic) -> tactic"} \\
 | 
645  | 
  @{define_ML SOMEGOAL: "(int -> tactic) -> tactic"} \\
 | 
|
646  | 
  @{define_ML FIRSTGOAL: "(int -> tactic) -> tactic"} \\
 | 
|
647  | 
  @{define_ML HEADGOAL: "(int -> tactic) -> tactic"} \\
 | 
|
648  | 
  @{define_ML REPEAT_SOME: "(int -> tactic) -> tactic"} \\
 | 
|
649  | 
  @{define_ML REPEAT_FIRST: "(int -> tactic) -> tactic"} \\
 | 
|
650  | 
  @{define_ML RANGE: "(int -> tactic) list -> int -> tactic"} \\
 | 
|
| 46263 | 651  | 
  \end{mldecls}
 | 
652  | 
||
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
653  | 
\<^descr> \<^ML>\<open>ALLGOALS\<close>~\<open>tac\<close> is equivalent to \<open>tac n\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac 1\<close>. It applies the \<open>tac\<close> to all the subgoals, counting downwards.  | 
| 69597 | 654  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
655  | 
\<^descr> \<^ML>\<open>SOMEGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac n\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac 1\<close>. It applies \<open>tac\<close> to one subgoal, counting downwards.  | 
| 46263 | 656  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
657  | 
\<^descr> \<^ML>\<open>FIRSTGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac 1\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>ORELSE\<close>~\<open>tac n\<close>. It applies \<open>tac\<close> to one subgoal, counting upwards.  | 
| 46263 | 658  | 
|
| 69597 | 659  | 
\<^descr> \<^ML>\<open>HEADGOAL\<close>~\<open>tac\<close> is equivalent to \<open>tac 1\<close>. It applies \<open>tac\<close>  | 
| 61854 | 660  | 
unconditionally to the first subgoal.  | 
| 46267 | 661  | 
|
| 69597 | 662  | 
\<^descr> \<^ML>\<open>REPEAT_SOME\<close>~\<open>tac\<close> applies \<open>tac\<close> once or more to a subgoal, counting  | 
| 61854 | 663  | 
downwards.  | 
| 46263 | 664  | 
|
| 69597 | 665  | 
\<^descr> \<^ML>\<open>REPEAT_FIRST\<close>~\<open>tac\<close> applies \<open>tac\<close> once or more to a subgoal, counting  | 
| 61854 | 666  | 
upwards.  | 
| 46263 | 667  | 
|
| 69597 | 668  | 
\<^descr> \<^ML>\<open>RANGE\<close>~\<open>[tac\<^sub>1, \<dots>, tac\<^sub>k] i\<close> is equivalent to \<open>tac\<^sub>k (i + k -  | 
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
669  | 
1)\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>\<dots>\<close>~\<^ML_infix>\<open>THEN\<close>~\<open>tac\<^sub>1 i\<close>. It applies the given list of  | 
| 61854 | 670  | 
tactics to the corresponding range of subgoals, counting downwards.  | 
| 58618 | 671  | 
\<close>  | 
| 46263 | 672  | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
673  | 
|
| 58618 | 674  | 
subsection \<open>Control and search tacticals\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
675  | 
|
| 61854 | 676  | 
text \<open>  | 
| 69597 | 677  | 
A predicate on theorems \<^ML_type>\<open>thm -> bool\<close> can test whether a goal  | 
| 61854 | 678  | 
state enjoys some desirable property --- such as having no subgoals. Tactics  | 
679  | 
that search for satisfactory goal states are easy to express. The main  | 
|
680  | 
search procedures, depth-first, breadth-first and best-first, are provided  | 
|
681  | 
as tacticals. They generate the search tree by repeatedly applying a given  | 
|
682  | 
tactic.  | 
|
683  | 
\<close>  | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
684  | 
|
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
685  | 
|
| 46270 | 686  | 
text %mlref ""  | 
687  | 
||
| 58618 | 688  | 
subsubsection \<open>Filtering a tactic's results\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
689  | 
|
| 58618 | 690  | 
text \<open>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
691  | 
  \begin{mldecls}
 | 
| 73764 | 692  | 
  @{define_ML FILTER: "(thm -> bool) -> tactic -> tactic"} \\
 | 
693  | 
  @{define_ML CHANGED: "tactic -> tactic"} \\
 | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
694  | 
  \end{mldecls}
 | 
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
695  | 
|
| 69597 | 696  | 
\<^descr> \<^ML>\<open>FILTER\<close>~\<open>sat tac\<close> applies \<open>tac\<close> to the goal state and returns a  | 
| 61854 | 697  | 
sequence consisting of those result goal states that are satisfactory in the  | 
698  | 
sense of \<open>sat\<close>.  | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
699  | 
|
| 69597 | 700  | 
\<^descr> \<^ML>\<open>CHANGED\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and returns precisely  | 
701  | 
those states that differ from the original state (according to \<^ML>\<open>Thm.eq_thm\<close>). Thus \<^ML>\<open>CHANGED\<close>~\<open>tac\<close> always has some effect on the state.  | 
|
| 58618 | 702  | 
\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
703  | 
|
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
704  | 
|
| 58618 | 705  | 
subsubsection \<open>Depth-first search\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
706  | 
|
| 58618 | 707  | 
text \<open>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
708  | 
  \begin{mldecls}
 | 
| 73764 | 709  | 
  @{define_ML DEPTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
 | 
710  | 
  @{define_ML DEPTH_SOLVE: "tactic -> tactic"} \\
 | 
|
711  | 
  @{define_ML DEPTH_SOLVE_1: "tactic -> tactic"} \\
 | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
712  | 
  \end{mldecls}
 | 
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
713  | 
|
| 69597 | 714  | 
\<^descr> \<^ML>\<open>DEPTH_FIRST\<close>~\<open>sat tac\<close> returns the goal state if \<open>sat\<close> returns true.  | 
| 61854 | 715  | 
Otherwise it applies \<open>tac\<close>, then recursively searches from each element of  | 
716  | 
the resulting sequence. The code uses a stack for efficiency, in effect  | 
|
| 
73765
 
ebaed09ce06e
more uniform document antiquotations for ML: consolidate former setup for manuals;
 
wenzelm 
parents: 
73764 
diff
changeset
 | 
717  | 
applying \<open>tac\<close>~\<^ML_infix>\<open>THEN\<close>~\<^ML>\<open>DEPTH_FIRST\<close>~\<open>sat tac\<close> to the state.  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
718  | 
|
| 69597 | 719  | 
\<^descr> \<^ML>\<open>DEPTH_SOLVE\<close>\<open>tac\<close> uses \<^ML>\<open>DEPTH_FIRST\<close> to search for states having  | 
| 61854 | 720  | 
no subgoals.  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
721  | 
|
| 69597 | 722  | 
\<^descr> \<^ML>\<open>DEPTH_SOLVE_1\<close>~\<open>tac\<close> uses \<^ML>\<open>DEPTH_FIRST\<close> to search for states  | 
| 61854 | 723  | 
having fewer subgoals than the given state. Thus, it insists upon solving at  | 
724  | 
least one subgoal.  | 
|
| 58618 | 725  | 
\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
726  | 
|
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
727  | 
|
| 58618 | 728  | 
subsubsection \<open>Other search strategies\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
729  | 
|
| 58618 | 730  | 
text \<open>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
731  | 
  \begin{mldecls}
 | 
| 73764 | 732  | 
  @{define_ML BREADTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\
 | 
733  | 
  @{define_ML BEST_FIRST: "(thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
 | 
|
734  | 
  @{define_ML THEN_BEST_FIRST: "tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic"} \\
 | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
735  | 
  \end{mldecls}
 | 
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
736  | 
|
| 61854 | 737  | 
These search strategies will find a solution if one exists. However, they do  | 
738  | 
not enumerate all solutions; they terminate after the first satisfactory  | 
|
739  | 
result from \<open>tac\<close>.  | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
740  | 
|
| 69597 | 741  | 
\<^descr> \<^ML>\<open>BREADTH_FIRST\<close>~\<open>sat tac\<close> uses breadth-first search to find states for  | 
| 61854 | 742  | 
which \<open>sat\<close> is true. For most applications, it is too slow.  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
743  | 
|
| 69597 | 744  | 
\<^descr> \<^ML>\<open>BEST_FIRST\<close>~\<open>(sat, dist) tac\<close> does a heuristic search, using \<open>dist\<close>  | 
| 61854 | 745  | 
to estimate the distance from a satisfactory state (in the sense of \<open>sat\<close>).  | 
746  | 
It maintains a list of states ordered by distance. It applies \<open>tac\<close> to the  | 
|
747  | 
head of this list; if the result contains any satisfactory states, then it  | 
|
| 69597 | 748  | 
returns them. Otherwise, \<^ML>\<open>BEST_FIRST\<close> adds the new states to the list,  | 
| 61854 | 749  | 
and continues.  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
750  | 
|
| 69597 | 751  | 
The distance function is typically \<^ML>\<open>size_of_thm\<close>, which computes the  | 
| 61854 | 752  | 
size of the state. The smaller the state, the fewer and simpler subgoals it  | 
753  | 
has.  | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
754  | 
|
| 69597 | 755  | 
\<^descr> \<^ML>\<open>THEN_BEST_FIRST\<close>~\<open>tac\<^sub>0 (sat, dist) tac\<close> is like \<^ML>\<open>BEST_FIRST\<close>,  | 
| 61854 | 756  | 
except that the priority queue initially contains the result of applying  | 
757  | 
\<open>tac\<^sub>0\<close> to the goal state. This tactical permits separate tactics for  | 
|
758  | 
starting the search and continuing the search.  | 
|
| 58618 | 759  | 
\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
760  | 
|
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
761  | 
|
| 58618 | 762  | 
subsubsection \<open>Auxiliary tacticals for searching\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
763  | 
|
| 58618 | 764  | 
text \<open>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
765  | 
  \begin{mldecls}
 | 
| 73764 | 766  | 
  @{define_ML COND: "(thm -> bool) -> tactic -> tactic -> tactic"} \\
 | 
767  | 
  @{define_ML IF_UNSOLVED: "tactic -> tactic"} \\
 | 
|
768  | 
  @{define_ML SOLVE: "tactic -> tactic"} \\
 | 
|
769  | 
  @{define_ML DETERM: "tactic -> tactic"} \\
 | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
770  | 
  \end{mldecls}
 | 
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
771  | 
|
| 69597 | 772  | 
\<^descr> \<^ML>\<open>COND\<close>~\<open>sat tac\<^sub>1 tac\<^sub>2\<close> applies \<open>tac\<^sub>1\<close> to the goal state if it  | 
| 61854 | 773  | 
satisfies predicate \<open>sat\<close>, and applies \<open>tac\<^sub>2\<close>. It is a conditional tactical  | 
774  | 
in that only one of \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close> is applied to a goal state. However,  | 
|
775  | 
both \<open>tac\<^sub>1\<close> and \<open>tac\<^sub>2\<close> are evaluated because ML uses eager evaluation.  | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
776  | 
|
| 69597 | 777  | 
\<^descr> \<^ML>\<open>IF_UNSOLVED\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state if it has any  | 
| 61854 | 778  | 
subgoals, and simply returns the goal state otherwise. Many common tactics,  | 
| 69597 | 779  | 
such as \<^ML>\<open>resolve_tac\<close>, fail if applied to a goal state that has no  | 
| 61854 | 780  | 
subgoals.  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
781  | 
|
| 69597 | 782  | 
\<^descr> \<^ML>\<open>SOLVE\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and then fails iff there  | 
| 61854 | 783  | 
are subgoals left.  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
784  | 
|
| 69597 | 785  | 
\<^descr> \<^ML>\<open>DETERM\<close>~\<open>tac\<close> applies \<open>tac\<close> to the goal state and returns the head of  | 
786  | 
the resulting sequence. \<^ML>\<open>DETERM\<close> limits the search space by making its  | 
|
| 61854 | 787  | 
argument deterministic.  | 
| 58618 | 788  | 
\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
789  | 
|
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
790  | 
|
| 58618 | 791  | 
subsubsection \<open>Predicates and functions useful for searching\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
792  | 
|
| 58618 | 793  | 
text \<open>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
794  | 
  \begin{mldecls}
 | 
| 73764 | 795  | 
  @{define_ML has_fewer_prems: "int -> thm -> bool"} \\
 | 
796  | 
  @{define_ML Thm.eq_thm: "thm * thm -> bool"} \\
 | 
|
797  | 
  @{define_ML Thm.eq_thm_prop: "thm * thm -> bool"} \\
 | 
|
798  | 
  @{define_ML size_of_thm: "thm -> int"} \\
 | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
799  | 
  \end{mldecls}
 | 
| 
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
800  | 
|
| 69597 | 801  | 
\<^descr> \<^ML>\<open>has_fewer_prems\<close>~\<open>n thm\<close> reports whether \<open>thm\<close> has fewer than \<open>n\<close>  | 
| 61854 | 802  | 
premises.  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
803  | 
|
| 69597 | 804  | 
\<^descr> \<^ML>\<open>Thm.eq_thm\<close>~\<open>(thm\<^sub>1, thm\<^sub>2)\<close> reports whether \<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> are  | 
| 61854 | 805  | 
equal. Both theorems must have the same conclusions, the same set of  | 
806  | 
hypotheses, and the same set of sort hypotheses. Names of bound variables  | 
|
807  | 
are ignored as usual.  | 
|
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
808  | 
|
| 69597 | 809  | 
\<^descr> \<^ML>\<open>Thm.eq_thm_prop\<close>~\<open>(thm\<^sub>1, thm\<^sub>2)\<close> reports whether the propositions of  | 
| 61854 | 810  | 
\<open>thm\<^sub>1\<close> and \<open>thm\<^sub>2\<close> are equal. Names of bound variables are ignored.  | 
811  | 
||
| 69597 | 812  | 
\<^descr> \<^ML>\<open>size_of_thm\<close>~\<open>thm\<close> computes the size of \<open>thm\<close>, namely the number of  | 
| 61854 | 813  | 
variables, constants and abstractions in its conclusion. It may serve as a  | 
| 69597 | 814  | 
distance function for \<^ML>\<open>BEST_FIRST\<close>.  | 
| 58618 | 815  | 
\<close>  | 
| 
46269
 
e75181672150
updated "Control and search tacticals" (moved from ref to implementation);
 
wenzelm 
parents: 
46267 
diff
changeset
 | 
816  | 
|
| 18537 | 817  | 
end  |