doc-src/IsarRef/Thy/ML_Tactic.thy
author wenzelm
Mon, 16 Jun 2008 22:13:39 +0200
changeset 27239 f2f42f9fa09d
parent 27208 5fe899199f85
child 30168 9a20be5be90b
permissions -rw-r--r--
pervasive RuleInsts;
     1 (* $Id$ *)
     2 
     3 theory ML_Tactic
     4 imports Main
     5 begin
     6 
     7 chapter {* ML tactic expressions *}
     8 
     9 text {*
    10   Isar Proof methods closely resemble traditional tactics, when used
    11   in unstructured sequences of @{command "apply"} commands.
    12   Isabelle/Isar provides emulations for all major ML tactics of
    13   classic Isabelle --- mostly for the sake of easy porting of existing
    14   developments, as actual Isar proof texts would demand much less
    15   diversity of proof methods.
    16 
    17   Unlike tactic expressions in ML, Isar proof methods provide proper
    18   concrete syntax for additional arguments, options, modifiers etc.
    19   Thus a typical method text is usually more concise than the
    20   corresponding ML tactic.  Furthermore, the Isar versions of classic
    21   Isabelle tactics often cover several variant forms by a single
    22   method with separate options to tune the behavior.  For example,
    23   method @{method simp} replaces all of @{ML simp_tac}~/ @{ML
    24   asm_simp_tac}~/ @{ML full_simp_tac}~/ @{ML asm_full_simp_tac}, there
    25   is also concrete syntax for augmenting the Simplifier context (the
    26   current ``simpset'') in a convenient way.
    27 *}
    28 
    29 
    30 section {* Resolution tactics *}
    31 
    32 text {*
    33   Classic Isabelle provides several variant forms of tactics for
    34   single-step rule applications (based on higher-order resolution).
    35   The space of resolution tactics has the following main dimensions.
    36 
    37   \begin{enumerate}
    38 
    39   \item The ``mode'' of resolution: intro, elim, destruct, or forward
    40   (e.g.\ @{ML resolve_tac}, @{ML eresolve_tac}, @{ML dresolve_tac},
    41   @{ML forward_tac}).
    42 
    43   \item Optional explicit instantiation (e.g.\ @{ML resolve_tac} vs.\
    44   @{ML res_inst_tac}).
    45 
    46   \item Abbreviations for singleton arguments (e.g.\ @{ML resolve_tac}
    47   vs.\ @{ML rtac}).
    48 
    49   \end{enumerate}
    50 
    51   Basically, the set of Isar tactic emulations @{method rule_tac},
    52   @{method erule_tac}, @{method drule_tac}, @{method frule_tac} (see
    53   \secref{sec:tactics}) would be sufficient to cover the four modes,
    54   either with or without instantiation, and either with single or
    55   multiple arguments.  Although it is more convenient in most cases to
    56   use the plain @{method rule} method (see
    57   \secref{sec:pure-meth-att}), or any of its ``improper'' variants
    58   @{method erule}, @{method drule}, @{method frule} (see
    59   \secref{sec:misc-meth-att}).  Note that explicit goal addressing is
    60   only supported by the actual @{method rule_tac} version.
    61 
    62   With this in mind, plain resolution tactics correspond to Isar
    63   methods as follows.
    64 
    65   \medskip
    66   \begin{tabular}{lll}
    67     @{ML rtac}~@{text "a 1"} & & @{text "rule a"} \\
    68     @{ML resolve_tac}~@{text "[a\<^sub>1, \<dots>] 1"} & & @{text "rule a\<^sub>1 \<dots>"} \\
    69     @{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \<dots>] a 1"} & &
    70     @{text "rule_tac x\<^sub>1 = t\<^sub>1 \<AND> \<dots> \<IN> a"} \\[0.5ex]
    71     @{ML rtac}~@{text "a i"} & & @{text "rule_tac [i] a"} \\
    72     @{ML resolve_tac}~@{text "[a\<^sub>1, \<dots>] i"} & & @{text "rule_tac [i] a\<^sub>1 \<dots>"} \\
    73     @{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \<dots>] a i"} & &
    74     @{text "rule_tac [i] x\<^sub>1 = t\<^sub>1 \<AND> \<dots> \<IN> a"} \\
    75   \end{tabular}
    76   \medskip
    77 
    78   Note that explicit goal addressing may be usually avoided by
    79   changing the order of subgoals with @{command "defer"} or @{command
    80   "prefer"} (see \secref{sec:tactic-commands}).
    81 *}
    82 
    83 
    84 section {* Simplifier tactics *}
    85 
    86 text {*
    87   The main Simplifier tactics @{ML simp_tac} and variants (cf.\
    88   \cite{isabelle-ref}) are all covered by the @{method simp} and
    89   @{method simp_all} methods (see \secref{sec:simplifier}).  Note that
    90   there is no individual goal addressing available, simplification
    91   acts either on the first goal (@{method simp}) or all goals
    92   (@{method simp_all}).
    93 
    94   \medskip
    95   \begin{tabular}{lll}
    96     @{ML asm_full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp} \\
    97     @{ML ALLGOALS}~(@{ML asm_full_simp_tac}~@{text "@{simpset}"}) & & @{method simp_all} \\[0.5ex]
    98     @{ML simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm)"} \\
    99     @{ML asm_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_simp)"} \\
   100     @{ML full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_use)"} \\
   101     @{ML asm_lr_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(asm_lr)"} \\
   102   \end{tabular}
   103   \medskip
   104 *}
   105 
   106 
   107 section {* Classical Reasoner tactics *}
   108 
   109 text {*
   110   The Classical Reasoner provides a rather large number of variations
   111   of automated tactics, such as @{ML blast_tac}, @{ML fast_tac}, @{ML
   112   clarify_tac} etc.\ (see \cite{isabelle-ref}).  The corresponding
   113   Isar methods usually share the same base name, such as @{method
   114   blast}, @{method fast}, @{method clarify} etc.\ (see
   115   \secref{sec:classical}).
   116 *}
   117 
   118 
   119 section {* Miscellaneous tactics *}
   120 
   121 text {*
   122   There are a few additional tactics defined in various theories of
   123   Isabelle/HOL, some of these also in Isabelle/FOL or Isabelle/ZF.
   124   The most common ones of these may be ported to Isar as follows.
   125 
   126   \medskip
   127   \begin{tabular}{lll}
   128     @{ML stac}~@{text "a 1"} & & @{text "subst a"} \\
   129     @{ML hyp_subst_tac}~@{text 1} & & @{text hypsubst} \\
   130     @{ML strip_tac}~@{text 1} & @{text "\<approx>"} & @{text "intro strip"} \\
   131     @{ML split_all_tac}~@{text 1} & & @{text "simp (no_asm_simp) only: split_tupled_all"} \\
   132       & @{text "\<approx>"} & @{text "simp only: split_tupled_all"} \\
   133       & @{text "\<lless>"} & @{text "clarify"} \\
   134   \end{tabular}
   135 *}
   136 
   137 
   138 section {* Tacticals *}
   139 
   140 text {*
   141   Classic Isabelle provides a huge amount of tacticals for combination
   142   and modification of existing tactics.  This has been greatly reduced
   143   in Isar, providing the bare minimum of combinators only: ``@{text
   144   ","}'' (sequential composition), ``@{text "|"}'' (alternative
   145   choices), ``@{text "?"}'' (try), ``@{text "+"}'' (repeat at least
   146   once).  These are usually sufficient in practice; if all fails,
   147   arbitrary ML tactic code may be invoked via the @{method tactic}
   148   method (see \secref{sec:tactics}).
   149 
   150   \medskip Common ML tacticals may be expressed directly in Isar as
   151   follows:
   152 
   153   \medskip
   154   \begin{tabular}{lll}
   155     @{text "tac\<^sub>1"}~@{ML_text THEN}~@{text "tac\<^sub>2"} & & @{text "meth\<^sub>1, meth\<^sub>2"} \\
   156     @{text "tac\<^sub>1"}~@{ML_text ORELSE}~@{text "tac\<^sub>2"} & & @{text "meth\<^sub>1 | meth\<^sub>2"} \\
   157     @{ML TRY}~@{text tac} & & @{text "meth?"} \\
   158     @{ML REPEAT1}~@{text tac} & & @{text "meth+"} \\
   159     @{ML REPEAT}~@{text tac} & & @{text "(meth+)?"} \\
   160     @{ML EVERY}~@{text "[tac\<^sub>1, \<dots>]"} & & @{text "meth\<^sub>1, \<dots>"} \\
   161     @{ML FIRST}~@{text "[tac\<^sub>1, \<dots>]"} & & @{text "meth\<^sub>1 | \<dots>"} \\
   162   \end{tabular}
   163   \medskip
   164 
   165   \medskip @{ML CHANGED} (see \cite{isabelle-ref}) is usually not
   166   required in Isar, since most basic proof methods already fail unless
   167   there is an actual change in the goal state.  Nevertheless, ``@{text
   168   "?"}''  (try) may be used to accept \emph{unchanged} results as
   169   well.
   170 
   171   \medskip @{ML ALLGOALS}, @{ML SOMEGOAL} etc.\ (see
   172   \cite{isabelle-ref}) are not available in Isar, since there is no
   173   direct goal addressing.  Nevertheless, some basic methods address
   174   all goals internally, notably @{method simp_all} (see
   175   \secref{sec:simplifier}).  Also note that @{ML ALLGOALS} can be
   176   often replaced by ``@{text "+"}'' (repeat at least once), although
   177   this usually has a different operational behavior, such as solving
   178   goals in a different order.
   179 
   180   \medskip Iterated resolution, such as @{ML_text "REPEAT (FIRSTGOAL
   181   (resolve_tac \<dots>))"}, is usually better expressed using the @{method
   182   intro} and @{method elim} methods of Isar (see
   183   \secref{sec:classical}).
   184 *}
   185 
   186 end