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