| 18537 |      1 | 
 | 
|  |      2 | (* $Id$ *)
 | 
|  |      3 | 
 | 
|  |      4 | theory tactic imports base begin
 | 
|  |      5 | 
 | 
| 20452 |      6 | chapter {* Tactical reasoning *}
 | 
| 18537 |      7 | 
 | 
| 20451 |      8 | text {*
 | 
| 20474 |      9 |   Tactical reasoning works by refining the initial claim in a
 | 
|  |     10 |   backwards fashion, until a solved form is reached.  A @{text "goal"}
 | 
|  |     11 |   consists of several subgoals that need to be solved in order to
 | 
|  |     12 |   achieve the main statement; zero subgoals means that the proof may
 | 
|  |     13 |   be finished.  A @{text "tactic"} is a refinement operation that maps
 | 
|  |     14 |   a goal to a lazy sequence of potential successors.  A @{text
 | 
|  |     15 |   "tactical"} is a combinator for composing tactics.
 | 
| 20451 |     16 | *}
 | 
| 18537 |     17 | 
 | 
|  |     18 | 
 | 
|  |     19 | section {* Goals \label{sec:tactical-goals} *}
 | 
|  |     20 | 
 | 
| 20451 |     21 | text {*
 | 
|  |     22 |   Isabelle/Pure represents a goal\glossary{Tactical goal}{A theorem of
 | 
|  |     23 |   \seeglossary{Horn Clause} form stating that a number of subgoals
 | 
|  |     24 |   imply the main conclusion, which is marked as a protected
 | 
|  |     25 |   proposition.} as a theorem stating that the subgoals imply the main
 | 
|  |     26 |   goal: @{text "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C"}.  The outermost goal
 | 
|  |     27 |   structure is that of a Horn Clause\glossary{Horn Clause}{An iterated
 | 
|  |     28 |   implication @{text "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C"}, without any
 | 
|  |     29 |   outermost quantifiers.  Strictly speaking, propositions @{text
 | 
|  |     30 |   "A\<^sub>i"} need to be atomic in Horn Clauses, but Isabelle admits
 | 
|  |     31 |   arbitrary substructure here (nested @{text "\<Longrightarrow>"} and @{text "\<And>"}
 | 
|  |     32 |   connectives).}: i.e.\ an iterated implication without any
 | 
|  |     33 |   quantifiers\footnote{Recall that outermost @{text "\<And>x. \<phi>[x]"} is
 | 
|  |     34 |   always represented via schematic variables in the body: @{text
 | 
|  |     35 |   "\<phi>[?x]"}.  These variables may get instantiated during the course of
 | 
|  |     36 |   reasoning.}.  For @{text "n = 0"} a goal is called ``solved''.
 | 
| 18537 |     37 | 
 | 
| 20451 |     38 |   The structure of each subgoal @{text "A\<^sub>i"} is that of a
 | 
|  |     39 |   general Hereditary Harrop Formula @{text "\<And>x\<^sub>1 \<dots>
 | 
| 20474 |     40 |   \<And>x\<^sub>k. H\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> H\<^sub>m \<Longrightarrow> B"} in normal form where.
 | 
|  |     41 |   Here @{text "x\<^sub>1, \<dots>, x\<^sub>k"} are goal parameters, i.e.\
 | 
| 20451 |     42 |   arbitrary-but-fixed entities of certain types, and @{text
 | 
|  |     43 |   "H\<^sub>1, \<dots>, H\<^sub>m"} are goal hypotheses, i.e.\ facts that may
 | 
|  |     44 |   be assumed locally.  Together, this forms the goal context of the
 | 
|  |     45 |   conclusion @{text B} to be established.  The goal hypotheses may be
 | 
| 20474 |     46 |   again arbitrary Hereditary Harrop Formulas, although the level of
 | 
|  |     47 |   nesting rarely exceeds 1--2 in practice.
 | 
| 18537 |     48 | 
 | 
| 20451 |     49 |   The main conclusion @{text C} is internally marked as a protected
 | 
|  |     50 |   proposition\glossary{Protected proposition}{An arbitrarily
 | 
|  |     51 |   structured proposition @{text "C"} which is forced to appear as
 | 
|  |     52 |   atomic by wrapping it into a propositional identity operator;
 | 
|  |     53 |   notation @{text "#C"}.  Protecting a proposition prevents basic
 | 
|  |     54 |   inferences from entering into that structure for the time being.},
 | 
| 20474 |     55 |   which is represented explicitly by the notation @{text "#C"}.  This
 | 
|  |     56 |   ensures that the decomposition into subgoals and main conclusion is
 | 
|  |     57 |   well-defined for arbitrarily structured claims.
 | 
| 18537 |     58 | 
 | 
| 20451 |     59 |   \medskip Basic goal management is performed via the following
 | 
|  |     60 |   Isabelle/Pure rules:
 | 
| 18537 |     61 | 
 | 
|  |     62 |   \[
 | 
|  |     63 |   \infer[@{text "(init)"}]{@{text "C \<Longrightarrow> #C"}}{} \qquad
 | 
| 20547 |     64 |   \infer[@{text "(finish)"}]{@{text "C"}}{@{text "#C"}}
 | 
| 18537 |     65 |   \]
 | 
|  |     66 | 
 | 
|  |     67 |   \medskip The following low-level variants admit general reasoning
 | 
|  |     68 |   with protected propositions:
 | 
|  |     69 | 
 | 
|  |     70 |   \[
 | 
|  |     71 |   \infer[@{text "(protect)"}]{@{text "#C"}}{@{text "C"}} \qquad
 | 
|  |     72 |   \infer[@{text "(conclude)"}]{@{text "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> C"}}{@{text "A\<^sub>1 \<Longrightarrow> \<dots> \<Longrightarrow> A\<^sub>n \<Longrightarrow> #C"}}
 | 
|  |     73 |   \]
 | 
|  |     74 | *}
 | 
|  |     75 | 
 | 
|  |     76 | text %mlref {*
 | 
|  |     77 |   \begin{mldecls}
 | 
|  |     78 |   @{index_ML Goal.init: "cterm -> thm"} \\
 | 
|  |     79 |   @{index_ML Goal.finish: "thm -> thm"} \\
 | 
|  |     80 |   @{index_ML Goal.protect: "thm -> thm"} \\
 | 
|  |     81 |   @{index_ML Goal.conclude: "thm -> thm"} \\
 | 
|  |     82 |   \end{mldecls}
 | 
|  |     83 | 
 | 
|  |     84 |   \begin{description}
 | 
|  |     85 | 
 | 
| 20474 |     86 |   \item @{ML "Goal.init"}~@{text C} initializes a tactical goal from
 | 
|  |     87 |   the well-formed proposition @{text C}.
 | 
| 18537 |     88 | 
 | 
| 20474 |     89 |   \item @{ML "Goal.finish"}~@{text "thm"} checks whether theorem
 | 
|  |     90 |   @{text "thm"} is a solved goal (no subgoals), and concludes the
 | 
|  |     91 |   result by removing the goal protection.
 | 
| 18537 |     92 | 
 | 
| 20474 |     93 |   \item @{ML "Goal.protect"}~@{text "thm"} protects the full statement
 | 
|  |     94 |   of theorem @{text "thm"}.
 | 
| 18537 |     95 | 
 | 
| 20474 |     96 |   \item @{ML "Goal.conclude"}~@{text "thm"} removes the goal
 | 
|  |     97 |   protection, even if there are pending subgoals.
 | 
| 18537 |     98 | 
 | 
|  |     99 |   \end{description}
 | 
|  |    100 | *}
 | 
|  |    101 | 
 | 
|  |    102 | 
 | 
|  |    103 | section {* Tactics *}
 | 
|  |    104 | 
 | 
|  |    105 | text {*
 | 
|  |    106 | 
 | 
|  |    107 | FIXME
 | 
|  |    108 | 
 | 
|  |    109 | \glossary{Tactic}{Any function that maps a \seeglossary{tactical goal}
 | 
|  |    110 | to a lazy sequence of possible sucessors.  This scheme subsumes
 | 
|  |    111 | failure (empty result sequence), as well as lazy enumeration of proof
 | 
|  |    112 | search results.  Error conditions are usually mapped to an empty
 | 
|  |    113 | result, which gives further tactics a chance to try in turn.
 | 
|  |    114 | Commonly, tactics either take an argument to address a particular
 | 
|  |    115 | subgoal, or operate on a certain range of subgoals in a uniform
 | 
|  |    116 | fashion.  In any case, the main conclusion is normally left untouched,
 | 
|  |    117 | apart from instantiating \seeglossary{schematic variables}.}
 | 
|  |    118 | 
 | 
|  |    119 | 
 | 
|  |    120 | *}
 | 
|  |    121 | 
 | 
|  |    122 | section {* Tacticals *}
 | 
|  |    123 | 
 | 
|  |    124 | text {*
 | 
|  |    125 | 
 | 
|  |    126 | FIXME
 | 
|  |    127 | 
 | 
|  |    128 | \glossary{Tactical}{A functional combinator for building up complex
 | 
|  |    129 | tactics from simpler ones.  Typical tactical perform sequential
 | 
|  |    130 | composition, disjunction (choice), iteration, or goal addressing.
 | 
|  |    131 | Various search strategies may be expressed via tacticals.}
 | 
|  |    132 | 
 | 
|  |    133 | *}
 | 
|  |    134 | 
 | 
|  |    135 | end
 | 
|  |    136 | 
 |