src/Doc/IsarRef/Proof.thy
changeset 56420 b266e7a86485
parent 56419 f47de9e82b0f
child 56431 4eb88149c7b2
equal deleted inserted replaced
56419:f47de9e82b0f 56420:b266e7a86485
     1 theory Proof
       
     2 imports Base Main
       
     3 begin
       
     4 
       
     5 chapter {* Proofs \label{ch:proofs} *}
       
     6 
       
     7 text {*
       
     8   Proof commands perform transitions of Isar/VM machine
       
     9   configurations, which are block-structured, consisting of a stack of
       
    10   nodes with three main components: logical proof context, current
       
    11   facts, and open goals.  Isar/VM transitions are typed according to
       
    12   the following three different modes of operation:
       
    13 
       
    14   \begin{description}
       
    15 
       
    16   \item @{text "proof(prove)"} means that a new goal has just been
       
    17   stated that is now to be \emph{proven}; the next command may refine
       
    18   it by some proof method, and enter a sub-proof to establish the
       
    19   actual result.
       
    20 
       
    21   \item @{text "proof(state)"} is like a nested theory mode: the
       
    22   context may be augmented by \emph{stating} additional assumptions,
       
    23   intermediate results etc.
       
    24 
       
    25   \item @{text "proof(chain)"} is intermediate between @{text
       
    26   "proof(state)"} and @{text "proof(prove)"}: existing facts (i.e.\
       
    27   the contents of the special ``@{fact_ref this}'' register) have been
       
    28   just picked up in order to be used when refining the goal claimed
       
    29   next.
       
    30 
       
    31   \end{description}
       
    32 
       
    33   The proof mode indicator may be understood as an instruction to the
       
    34   writer, telling what kind of operation may be performed next.  The
       
    35   corresponding typings of proof commands restricts the shape of
       
    36   well-formed proof texts to particular command sequences.  So dynamic
       
    37   arrangements of commands eventually turn out as static texts of a
       
    38   certain structure.
       
    39 
       
    40   \Appref{ap:refcard} gives a simplified grammar of the (extensible)
       
    41   language emerging that way from the different types of proof
       
    42   commands.  The main ideas of the overall Isar framework are
       
    43   explained in \chref{ch:isar-framework}.
       
    44 *}
       
    45 
       
    46 
       
    47 section {* Proof structure *}
       
    48 
       
    49 subsection {* Formal notepad *}
       
    50 
       
    51 text {*
       
    52   \begin{matharray}{rcl}
       
    53     @{command_def "notepad"} & : & @{text "local_theory \<rightarrow> proof(state)"} \\
       
    54   \end{matharray}
       
    55 
       
    56   @{rail \<open>
       
    57     @@{command notepad} @'begin'
       
    58     ;
       
    59     @@{command end}
       
    60   \<close>}
       
    61 
       
    62   \begin{description}
       
    63 
       
    64   \item @{command "notepad"}~@{keyword "begin"} opens a proof state
       
    65   without any goal statement.  This allows to experiment with Isar,
       
    66   without producing any persistent result.
       
    67 
       
    68   The notepad can be closed by @{command "end"} or discontinued by
       
    69   @{command "oops"}.
       
    70 
       
    71   \end{description}
       
    72 *}
       
    73 
       
    74 
       
    75 subsection {* Blocks *}
       
    76 
       
    77 text {*
       
    78   \begin{matharray}{rcl}
       
    79     @{command_def "next"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
    80     @{command_def "{"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
    81     @{command_def "}"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
    82   \end{matharray}
       
    83 
       
    84   While Isar is inherently block-structured, opening and closing
       
    85   blocks is mostly handled rather casually, with little explicit
       
    86   user-intervention.  Any local goal statement automatically opens
       
    87   \emph{two} internal blocks, which are closed again when concluding
       
    88   the sub-proof (by @{command "qed"} etc.).  Sections of different
       
    89   context within a sub-proof may be switched via @{command "next"},
       
    90   which is just a single block-close followed by block-open again.
       
    91   The effect of @{command "next"} is to reset the local proof context;
       
    92   there is no goal focus involved here!
       
    93 
       
    94   For slightly more advanced applications, there are explicit block
       
    95   parentheses as well.  These typically achieve a stronger forward
       
    96   style of reasoning.
       
    97 
       
    98   \begin{description}
       
    99 
       
   100   \item @{command "next"} switches to a fresh block within a
       
   101   sub-proof, resetting the local context to the initial one.
       
   102 
       
   103   \item @{command "{"} and @{command "}"} explicitly open and close
       
   104   blocks.  Any current facts pass through ``@{command "{"}''
       
   105   unchanged, while ``@{command "}"}'' causes any result to be
       
   106   \emph{exported} into the enclosing context.  Thus fixed variables
       
   107   are generalized, assumptions discharged, and local definitions
       
   108   unfolded (cf.\ \secref{sec:proof-context}).  There is no difference
       
   109   of @{command "assume"} and @{command "presume"} in this mode of
       
   110   forward reasoning --- in contrast to plain backward reasoning with
       
   111   the result exported at @{command "show"} time.
       
   112 
       
   113   \end{description}
       
   114 *}
       
   115 
       
   116 
       
   117 subsection {* Omitting proofs *}
       
   118 
       
   119 text {*
       
   120   \begin{matharray}{rcl}
       
   121     @{command_def "oops"} & : & @{text "proof \<rightarrow> local_theory | theory"} \\
       
   122   \end{matharray}
       
   123 
       
   124   The @{command "oops"} command discontinues the current proof
       
   125   attempt, while considering the partial proof text as properly
       
   126   processed.  This is conceptually quite different from ``faking''
       
   127   actual proofs via @{command_ref "sorry"} (see
       
   128   \secref{sec:proof-steps}): @{command "oops"} does not observe the
       
   129   proof structure at all, but goes back right to the theory level.
       
   130   Furthermore, @{command "oops"} does not produce any result theorem
       
   131   --- there is no intended claim to be able to complete the proof
       
   132   in any way.
       
   133 
       
   134   A typical application of @{command "oops"} is to explain Isar proofs
       
   135   \emph{within} the system itself, in conjunction with the document
       
   136   preparation tools of Isabelle described in \chref{ch:document-prep}.
       
   137   Thus partial or even wrong proof attempts can be discussed in a
       
   138   logically sound manner.  Note that the Isabelle {\LaTeX} macros can
       
   139   be easily adapted to print something like ``@{text "\<dots>"}'' instead of
       
   140   the keyword ``@{command "oops"}''.
       
   141 *}
       
   142 
       
   143 
       
   144 section {* Statements *}
       
   145 
       
   146 subsection {* Context elements \label{sec:proof-context} *}
       
   147 
       
   148 text {*
       
   149   \begin{matharray}{rcl}
       
   150     @{command_def "fix"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
   151     @{command_def "assume"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
   152     @{command_def "presume"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
   153     @{command_def "def"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
   154   \end{matharray}
       
   155 
       
   156   The logical proof context consists of fixed variables and
       
   157   assumptions.  The former closely correspond to Skolem constants, or
       
   158   meta-level universal quantification as provided by the Isabelle/Pure
       
   159   logical framework.  Introducing some \emph{arbitrary, but fixed}
       
   160   variable via ``@{command "fix"}~@{text x}'' results in a local value
       
   161   that may be used in the subsequent proof as any other variable or
       
   162   constant.  Furthermore, any result @{text "\<turnstile> \<phi>[x]"} exported from
       
   163   the context will be universally closed wrt.\ @{text x} at the
       
   164   outermost level: @{text "\<turnstile> \<And>x. \<phi>[x]"} (this is expressed in normal
       
   165   form using Isabelle's meta-variables).
       
   166 
       
   167   Similarly, introducing some assumption @{text \<chi>} has two effects.
       
   168   On the one hand, a local theorem is created that may be used as a
       
   169   fact in subsequent proof steps.  On the other hand, any result
       
   170   @{text "\<chi> \<turnstile> \<phi>"} exported from the context becomes conditional wrt.\
       
   171   the assumption: @{text "\<turnstile> \<chi> \<Longrightarrow> \<phi>"}.  Thus, solving an enclosing goal
       
   172   using such a result would basically introduce a new subgoal stemming
       
   173   from the assumption.  How this situation is handled depends on the
       
   174   version of assumption command used: while @{command "assume"}
       
   175   insists on solving the subgoal by unification with some premise of
       
   176   the goal, @{command "presume"} leaves the subgoal unchanged in order
       
   177   to be proved later by the user.
       
   178 
       
   179   Local definitions, introduced by ``@{command "def"}~@{text "x \<equiv>
       
   180   t"}'', are achieved by combining ``@{command "fix"}~@{text x}'' with
       
   181   another version of assumption that causes any hypothetical equation
       
   182   @{text "x \<equiv> t"} to be eliminated by the reflexivity rule.  Thus,
       
   183   exporting some result @{text "x \<equiv> t \<turnstile> \<phi>[x]"} yields @{text "\<turnstile>
       
   184   \<phi>[t]"}.
       
   185 
       
   186   @{rail \<open>
       
   187     @@{command fix} (@{syntax vars} + @'and')
       
   188     ;
       
   189     (@@{command assume} | @@{command presume}) (@{syntax props} + @'and')
       
   190     ;
       
   191     @@{command def} (def + @'and')
       
   192     ;
       
   193     def: @{syntax thmdecl}? \<newline>
       
   194       @{syntax name} ('==' | '\<equiv>') @{syntax term} @{syntax term_pat}?
       
   195   \<close>}
       
   196 
       
   197   \begin{description}
       
   198   
       
   199   \item @{command "fix"}~@{text x} introduces a local variable @{text
       
   200   x} that is \emph{arbitrary, but fixed.}
       
   201   
       
   202   \item @{command "assume"}~@{text "a: \<phi>"} and @{command
       
   203   "presume"}~@{text "a: \<phi>"} introduce a local fact @{text "\<phi> \<turnstile> \<phi>"} by
       
   204   assumption.  Subsequent results applied to an enclosing goal (e.g.\
       
   205   by @{command_ref "show"}) are handled as follows: @{command
       
   206   "assume"} expects to be able to unify with existing premises in the
       
   207   goal, while @{command "presume"} leaves @{text \<phi>} as new subgoals.
       
   208   
       
   209   Several lists of assumptions may be given (separated by
       
   210   @{keyword_ref "and"}; the resulting list of current facts consists
       
   211   of all of these concatenated.
       
   212   
       
   213   \item @{command "def"}~@{text "x \<equiv> t"} introduces a local
       
   214   (non-polymorphic) definition.  In results exported from the context,
       
   215   @{text x} is replaced by @{text t}.  Basically, ``@{command
       
   216   "def"}~@{text "x \<equiv> t"}'' abbreviates ``@{command "fix"}~@{text
       
   217   x}~@{command "assume"}~@{text "x \<equiv> t"}'', with the resulting
       
   218   hypothetical equation solved by reflexivity.
       
   219   
       
   220   The default name for the definitional equation is @{text x_def}.
       
   221   Several simultaneous definitions may be given at the same time.
       
   222 
       
   223   \end{description}
       
   224 
       
   225   The special name @{fact_ref prems} refers to all assumptions of the
       
   226   current context as a list of theorems.  This feature should be used
       
   227   with great care!  It is better avoided in final proof texts.
       
   228 *}
       
   229 
       
   230 
       
   231 subsection {* Term abbreviations \label{sec:term-abbrev} *}
       
   232 
       
   233 text {*
       
   234   \begin{matharray}{rcl}
       
   235     @{command_def "let"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
   236     @{keyword_def "is"} & : & syntax \\
       
   237   \end{matharray}
       
   238 
       
   239   Abbreviations may be either bound by explicit @{command
       
   240   "let"}~@{text "p \<equiv> t"} statements, or by annotating assumptions or
       
   241   goal statements with a list of patterns ``@{text "(\<IS> p\<^sub>1 \<dots>
       
   242   p\<^sub>n)"}''.  In both cases, higher-order matching is invoked to
       
   243   bind extra-logical term variables, which may be either named
       
   244   schematic variables of the form @{text ?x}, or nameless dummies
       
   245   ``@{variable _}'' (underscore). Note that in the @{command "let"}
       
   246   form the patterns occur on the left-hand side, while the @{keyword
       
   247   "is"} patterns are in postfix position.
       
   248 
       
   249   Polymorphism of term bindings is handled in Hindley-Milner style,
       
   250   similar to ML.  Type variables referring to local assumptions or
       
   251   open goal statements are \emph{fixed}, while those of finished
       
   252   results or bound by @{command "let"} may occur in \emph{arbitrary}
       
   253   instances later.  Even though actual polymorphism should be rarely
       
   254   used in practice, this mechanism is essential to achieve proper
       
   255   incremental type-inference, as the user proceeds to build up the
       
   256   Isar proof text from left to right.
       
   257 
       
   258   \medskip Term abbreviations are quite different from local
       
   259   definitions as introduced via @{command "def"} (see
       
   260   \secref{sec:proof-context}).  The latter are visible within the
       
   261   logic as actual equations, while abbreviations disappear during the
       
   262   input process just after type checking.  Also note that @{command
       
   263   "def"} does not support polymorphism.
       
   264 
       
   265   @{rail \<open>
       
   266     @@{command let} ((@{syntax term} + @'and') '=' @{syntax term} + @'and')
       
   267   \<close>}
       
   268 
       
   269   The syntax of @{keyword "is"} patterns follows @{syntax term_pat} or
       
   270   @{syntax prop_pat} (see \secref{sec:term-decls}).
       
   271 
       
   272   \begin{description}
       
   273 
       
   274   \item @{command "let"}~@{text "p\<^sub>1 = t\<^sub>1 \<AND> \<dots> p\<^sub>n = t\<^sub>n"} binds any
       
   275   text variables in patterns @{text "p\<^sub>1, \<dots>, p\<^sub>n"} by simultaneous
       
   276   higher-order matching against terms @{text "t\<^sub>1, \<dots>, t\<^sub>n"}.
       
   277 
       
   278   \item @{text "(\<IS> p\<^sub>1 \<dots> p\<^sub>n)"} resembles @{command "let"}, but
       
   279   matches @{text "p\<^sub>1, \<dots>, p\<^sub>n"} against the preceding statement.  Also
       
   280   note that @{keyword "is"} is not a separate command, but part of
       
   281   others (such as @{command "assume"}, @{command "have"} etc.).
       
   282 
       
   283   \end{description}
       
   284 
       
   285   Some \emph{implicit} term abbreviations\index{term abbreviations}
       
   286   for goals and facts are available as well.  For any open goal,
       
   287   @{variable_ref thesis} refers to its object-level statement,
       
   288   abstracted over any meta-level parameters (if present).  Likewise,
       
   289   @{variable_ref this} is bound for fact statements resulting from
       
   290   assumptions or finished goals.  In case @{variable this} refers to
       
   291   an object-logic statement that is an application @{text "f t"}, then
       
   292   @{text t} is bound to the special text variable ``@{variable "\<dots>"}''
       
   293   (three dots).  The canonical application of this convenience are
       
   294   calculational proofs (see \secref{sec:calculation}).
       
   295 *}
       
   296 
       
   297 
       
   298 subsection {* Facts and forward chaining \label{sec:proof-facts} *}
       
   299 
       
   300 text {*
       
   301   \begin{matharray}{rcl}
       
   302     @{command_def "note"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
   303     @{command_def "then"} & : & @{text "proof(state) \<rightarrow> proof(chain)"} \\
       
   304     @{command_def "from"} & : & @{text "proof(state) \<rightarrow> proof(chain)"} \\
       
   305     @{command_def "with"} & : & @{text "proof(state) \<rightarrow> proof(chain)"} \\
       
   306     @{command_def "using"} & : & @{text "proof(prove) \<rightarrow> proof(prove)"} \\
       
   307     @{command_def "unfolding"} & : & @{text "proof(prove) \<rightarrow> proof(prove)"} \\
       
   308   \end{matharray}
       
   309 
       
   310   New facts are established either by assumption or proof of local
       
   311   statements.  Any fact will usually be involved in further proofs,
       
   312   either as explicit arguments of proof methods, or when forward
       
   313   chaining towards the next goal via @{command "then"} (and variants);
       
   314   @{command "from"} and @{command "with"} are composite forms
       
   315   involving @{command "note"}.  The @{command "using"} elements
       
   316   augments the collection of used facts \emph{after} a goal has been
       
   317   stated.  Note that the special theorem name @{fact_ref this} refers
       
   318   to the most recently established facts, but only \emph{before}
       
   319   issuing a follow-up claim.
       
   320 
       
   321   @{rail \<open>
       
   322     @@{command note} (@{syntax thmdef}? @{syntax thmrefs} + @'and')
       
   323     ;
       
   324     (@@{command from} | @@{command with} | @@{command using} | @@{command unfolding})
       
   325       (@{syntax thmrefs} + @'and')
       
   326   \<close>}
       
   327 
       
   328   \begin{description}
       
   329 
       
   330   \item @{command "note"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"} recalls existing facts
       
   331   @{text "b\<^sub>1, \<dots>, b\<^sub>n"}, binding the result as @{text a}.  Note that
       
   332   attributes may be involved as well, both on the left and right hand
       
   333   sides.
       
   334 
       
   335   \item @{command "then"} indicates forward chaining by the current
       
   336   facts in order to establish the goal to be claimed next.  The
       
   337   initial proof method invoked to refine that will be offered the
       
   338   facts to do ``anything appropriate'' (see also
       
   339   \secref{sec:proof-steps}).  For example, method @{method (Pure) rule}
       
   340   (see \secref{sec:pure-meth-att}) would typically do an elimination
       
   341   rather than an introduction.  Automatic methods usually insert the
       
   342   facts into the goal state before operation.  This provides a simple
       
   343   scheme to control relevance of facts in automated proof search.
       
   344   
       
   345   \item @{command "from"}~@{text b} abbreviates ``@{command
       
   346   "note"}~@{text b}~@{command "then"}''; thus @{command "then"} is
       
   347   equivalent to ``@{command "from"}~@{text this}''.
       
   348   
       
   349   \item @{command "with"}~@{text "b\<^sub>1 \<dots> b\<^sub>n"} abbreviates ``@{command
       
   350   "from"}~@{text "b\<^sub>1 \<dots> b\<^sub>n \<AND> this"}''; thus the forward chaining
       
   351   is from earlier facts together with the current ones.
       
   352   
       
   353   \item @{command "using"}~@{text "b\<^sub>1 \<dots> b\<^sub>n"} augments the facts being
       
   354   currently indicated for use by a subsequent refinement step (such as
       
   355   @{command_ref "apply"} or @{command_ref "proof"}).
       
   356   
       
   357   \item @{command "unfolding"}~@{text "b\<^sub>1 \<dots> b\<^sub>n"} is structurally
       
   358   similar to @{command "using"}, but unfolds definitional equations
       
   359   @{text "b\<^sub>1, \<dots> b\<^sub>n"} throughout the goal state and facts.
       
   360 
       
   361   \end{description}
       
   362 
       
   363   Forward chaining with an empty list of theorems is the same as not
       
   364   chaining at all.  Thus ``@{command "from"}~@{text nothing}'' has no
       
   365   effect apart from entering @{text "prove(chain)"} mode, since
       
   366   @{fact_ref nothing} is bound to the empty list of theorems.
       
   367 
       
   368   Basic proof methods (such as @{method_ref (Pure) rule}) expect multiple
       
   369   facts to be given in their proper order, corresponding to a prefix
       
   370   of the premises of the rule involved.  Note that positions may be
       
   371   easily skipped using something like @{command "from"}~@{text "_
       
   372   \<AND> a \<AND> b"}, for example.  This involves the trivial rule
       
   373   @{text "PROP \<psi> \<Longrightarrow> PROP \<psi>"}, which is bound in Isabelle/Pure as
       
   374   ``@{fact_ref "_"}'' (underscore).
       
   375 
       
   376   Automated methods (such as @{method simp} or @{method auto}) just
       
   377   insert any given facts before their usual operation.  Depending on
       
   378   the kind of procedure involved, the order of facts is less
       
   379   significant here.
       
   380 *}
       
   381 
       
   382 
       
   383 subsection {* Goals \label{sec:goals} *}
       
   384 
       
   385 text {*
       
   386   \begin{matharray}{rcl}
       
   387     @{command_def "lemma"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
       
   388     @{command_def "theorem"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
       
   389     @{command_def "corollary"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
       
   390     @{command_def "schematic_lemma"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
       
   391     @{command_def "schematic_theorem"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
       
   392     @{command_def "schematic_corollary"} & : & @{text "local_theory \<rightarrow> proof(prove)"} \\
       
   393     @{command_def "have"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
       
   394     @{command_def "show"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
       
   395     @{command_def "hence"} & : & @{text "proof(state) \<rightarrow> proof(prove)"} \\
       
   396     @{command_def "thus"} & : & @{text "proof(state) \<rightarrow> proof(prove)"} \\
       
   397     @{command_def "print_statement"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
       
   398   \end{matharray}
       
   399 
       
   400   From a theory context, proof mode is entered by an initial goal
       
   401   command such as @{command "lemma"}, @{command "theorem"}, or
       
   402   @{command "corollary"}.  Within a proof, new claims may be
       
   403   introduced locally as well; four variants are available here to
       
   404   indicate whether forward chaining of facts should be performed
       
   405   initially (via @{command_ref "then"}), and whether the final result
       
   406   is meant to solve some pending goal.
       
   407 
       
   408   Goals may consist of multiple statements, resulting in a list of
       
   409   facts eventually.  A pending multi-goal is internally represented as
       
   410   a meta-level conjunction (@{text "&&&"}), which is usually
       
   411   split into the corresponding number of sub-goals prior to an initial
       
   412   method application, via @{command_ref "proof"}
       
   413   (\secref{sec:proof-steps}) or @{command_ref "apply"}
       
   414   (\secref{sec:tactic-commands}).  The @{method_ref induct} method
       
   415   covered in \secref{sec:cases-induct} acts on multiple claims
       
   416   simultaneously.
       
   417 
       
   418   Claims at the theory level may be either in short or long form.  A
       
   419   short goal merely consists of several simultaneous propositions
       
   420   (often just one).  A long goal includes an explicit context
       
   421   specification for the subsequent conclusion, involving local
       
   422   parameters and assumptions.  Here the role of each part of the
       
   423   statement is explicitly marked by separate keywords (see also
       
   424   \secref{sec:locale}); the local assumptions being introduced here
       
   425   are available as @{fact_ref assms} in the proof.  Moreover, there
       
   426   are two kinds of conclusions: @{element_def "shows"} states several
       
   427   simultaneous propositions (essentially a big conjunction), while
       
   428   @{element_def "obtains"} claims several simultaneous simultaneous
       
   429   contexts of (essentially a big disjunction of eliminated parameters
       
   430   and assumptions, cf.\ \secref{sec:obtain}).
       
   431 
       
   432   @{rail \<open>
       
   433     (@@{command lemma} | @@{command theorem} | @@{command corollary} |
       
   434      @@{command schematic_lemma} | @@{command schematic_theorem} |
       
   435      @@{command schematic_corollary}) @{syntax target}? (goal | longgoal)
       
   436     ;
       
   437     (@@{command have} | @@{command show} | @@{command hence} | @@{command thus}) goal
       
   438     ;
       
   439     @@{command print_statement} @{syntax modes}? @{syntax thmrefs}
       
   440     ;
       
   441   
       
   442     goal: (@{syntax props} + @'and')
       
   443     ;
       
   444     longgoal: @{syntax thmdecl}? (@{syntax_ref "includes"}?) (@{syntax context_elem} * ) conclusion
       
   445     ;
       
   446     conclusion: @'shows' goal | @'obtains' (@{syntax parname}? case + '|')
       
   447     ;
       
   448     case: (@{syntax vars} + @'and') @'where' (@{syntax props} + @'and')
       
   449   \<close>}
       
   450 
       
   451   \begin{description}
       
   452   
       
   453   \item @{command "lemma"}~@{text "a: \<phi>"} enters proof mode with
       
   454   @{text \<phi>} as main goal, eventually resulting in some fact @{text "\<turnstile>
       
   455   \<phi>"} to be put back into the target context.  An additional @{syntax
       
   456   context} specification may build up an initial proof context for the
       
   457   subsequent claim; this includes local definitions and syntax as
       
   458   well, see also @{syntax "includes"} in \secref{sec:bundle} and
       
   459   @{syntax context_elem} in \secref{sec:locale}.
       
   460   
       
   461   \item @{command "theorem"}~@{text "a: \<phi>"} and @{command
       
   462   "corollary"}~@{text "a: \<phi>"} are essentially the same as @{command
       
   463   "lemma"}~@{text "a: \<phi>"}, but the facts are internally marked as
       
   464   being of a different kind.  This discrimination acts like a formal
       
   465   comment.
       
   466 
       
   467   \item @{command "schematic_lemma"}, @{command "schematic_theorem"},
       
   468   @{command "schematic_corollary"} are similar to @{command "lemma"},
       
   469   @{command "theorem"}, @{command "corollary"}, respectively but allow
       
   470   the statement to contain unbound schematic variables.
       
   471 
       
   472   Under normal circumstances, an Isar proof text needs to specify
       
   473   claims explicitly.  Schematic goals are more like goals in Prolog,
       
   474   where certain results are synthesized in the course of reasoning.
       
   475   With schematic statements, the inherent compositionality of Isar
       
   476   proofs is lost, which also impacts performance, because proof
       
   477   checking is forced into sequential mode.
       
   478   
       
   479   \item @{command "have"}~@{text "a: \<phi>"} claims a local goal,
       
   480   eventually resulting in a fact within the current logical context.
       
   481   This operation is completely independent of any pending sub-goals of
       
   482   an enclosing goal statements, so @{command "have"} may be freely
       
   483   used for experimental exploration of potential results within a
       
   484   proof body.
       
   485   
       
   486   \item @{command "show"}~@{text "a: \<phi>"} is like @{command
       
   487   "have"}~@{text "a: \<phi>"} plus a second stage to refine some pending
       
   488   sub-goal for each one of the finished result, after having been
       
   489   exported into the corresponding context (at the head of the
       
   490   sub-proof of this @{command "show"} command).
       
   491   
       
   492   To accommodate interactive debugging, resulting rules are printed
       
   493   before being applied internally.  Even more, interactive execution
       
   494   of @{command "show"} predicts potential failure and displays the
       
   495   resulting error as a warning beforehand.  Watch out for the
       
   496   following message:
       
   497 
       
   498   %FIXME proper antiquitation
       
   499   \begin{ttbox}
       
   500   Problem! Local statement will fail to solve any pending goal
       
   501   \end{ttbox}
       
   502   
       
   503   \item @{command "hence"} abbreviates ``@{command "then"}~@{command
       
   504   "have"}'', i.e.\ claims a local goal to be proven by forward
       
   505   chaining the current facts.  Note that @{command "hence"} is also
       
   506   equivalent to ``@{command "from"}~@{text this}~@{command "have"}''.
       
   507   
       
   508   \item @{command "thus"} abbreviates ``@{command "then"}~@{command
       
   509   "show"}''.  Note that @{command "thus"} is also equivalent to
       
   510   ``@{command "from"}~@{text this}~@{command "show"}''.
       
   511   
       
   512   \item @{command "print_statement"}~@{text a} prints facts from the
       
   513   current theory or proof context in long statement form, according to
       
   514   the syntax for @{command "lemma"} given above.
       
   515 
       
   516   \end{description}
       
   517 
       
   518   Any goal statement causes some term abbreviations (such as
       
   519   @{variable_ref "?thesis"}) to be bound automatically, see also
       
   520   \secref{sec:term-abbrev}.
       
   521 
       
   522   The optional case names of @{element_ref "obtains"} have a twofold
       
   523   meaning: (1) during the of this claim they refer to the the local
       
   524   context introductions, (2) the resulting rule is annotated
       
   525   accordingly to support symbolic case splits when used with the
       
   526   @{method_ref cases} method (cf.\ \secref{sec:cases-induct}).
       
   527 *}
       
   528 
       
   529 
       
   530 section {* Refinement steps *}
       
   531 
       
   532 subsection {* Proof method expressions \label{sec:proof-meth} *}
       
   533 
       
   534 text {* Proof methods are either basic ones, or expressions composed
       
   535   of methods via ``@{verbatim ","}'' (sequential composition),
       
   536   ``@{verbatim "|"}'' (alternative choices), ``@{verbatim "?"}''
       
   537   (try), ``@{verbatim "+"}'' (repeat at least once), ``@{verbatim
       
   538   "["}@{text n}@{verbatim "]"}'' (restriction to first @{text n}
       
   539   sub-goals, with default @{text "n = 1"}).  In practice, proof
       
   540   methods are usually just a comma separated list of @{syntax
       
   541   nameref}~@{syntax args} specifications.  Note that parentheses may
       
   542   be dropped for single method specifications (with no arguments).
       
   543 
       
   544   @{rail \<open>
       
   545     @{syntax_def method}:
       
   546       (@{syntax nameref} | '(' methods ')') (() | '?' | '+' | '[' @{syntax nat}? ']')
       
   547     ;
       
   548     methods: (@{syntax nameref} @{syntax args} | @{syntax method}) + (',' | '|')
       
   549   \<close>}
       
   550 
       
   551   Proper Isar proof methods do \emph{not} admit arbitrary goal
       
   552   addressing, but refer either to the first sub-goal or all sub-goals
       
   553   uniformly.  The goal restriction operator ``@{text "[n]"}''
       
   554   evaluates a method expression within a sandbox consisting of the
       
   555   first @{text n} sub-goals (which need to exist).  For example, the
       
   556   method ``@{text "simp_all[3]"}'' simplifies the first three
       
   557   sub-goals, while ``@{text "(rule foo, simp_all)[]"}'' simplifies all
       
   558   new goals that emerge from applying rule @{text "foo"} to the
       
   559   originally first one.
       
   560 
       
   561   Improper methods, notably tactic emulations, offer a separate
       
   562   low-level goal addressing scheme as explicit argument to the
       
   563   individual tactic being involved.  Here ``@{text "[!]"}'' refers to
       
   564   all goals, and ``@{text "[n-]"}'' to all goals starting from @{text
       
   565   "n"}.
       
   566 
       
   567   @{rail \<open>
       
   568     @{syntax_def goal_spec}:
       
   569       '[' (@{syntax nat} '-' @{syntax nat} | @{syntax nat} '-' | @{syntax nat} | '!' ) ']'
       
   570   \<close>}
       
   571 *}
       
   572 
       
   573 
       
   574 subsection {* Initial and terminal proof steps \label{sec:proof-steps} *}
       
   575 
       
   576 text {*
       
   577   \begin{matharray}{rcl}
       
   578     @{command_def "proof"} & : & @{text "proof(prove) \<rightarrow> proof(state)"} \\
       
   579     @{command_def "qed"} & : & @{text "proof(state) \<rightarrow> proof(state) | local_theory | theory"} \\
       
   580     @{command_def "by"} & : & @{text "proof(prove) \<rightarrow> proof(state) | local_theory | theory"} \\
       
   581     @{command_def ".."} & : & @{text "proof(prove) \<rightarrow> proof(state) | local_theory | theory"} \\
       
   582     @{command_def "."} & : & @{text "proof(prove) \<rightarrow> proof(state) | local_theory | theory"} \\
       
   583     @{command_def "sorry"} & : & @{text "proof(prove) \<rightarrow> proof(state) | local_theory | theory"} \\
       
   584   \end{matharray}
       
   585 
       
   586   Arbitrary goal refinement via tactics is considered harmful.
       
   587   Structured proof composition in Isar admits proof methods to be
       
   588   invoked in two places only.
       
   589 
       
   590   \begin{enumerate}
       
   591 
       
   592   \item An \emph{initial} refinement step @{command_ref
       
   593   "proof"}~@{text "m\<^sub>1"} reduces a newly stated goal to a number
       
   594   of sub-goals that are to be solved later.  Facts are passed to
       
   595   @{text "m\<^sub>1"} for forward chaining, if so indicated by @{text
       
   596   "proof(chain)"} mode.
       
   597   
       
   598   \item A \emph{terminal} conclusion step @{command_ref "qed"}~@{text
       
   599   "m\<^sub>2"} is intended to solve remaining goals.  No facts are
       
   600   passed to @{text "m\<^sub>2"}.
       
   601 
       
   602   \end{enumerate}
       
   603 
       
   604   The only other (proper) way to affect pending goals in a proof body
       
   605   is by @{command_ref "show"}, which involves an explicit statement of
       
   606   what is to be solved eventually.  Thus we avoid the fundamental
       
   607   problem of unstructured tactic scripts that consist of numerous
       
   608   consecutive goal transformations, with invisible effects.
       
   609 
       
   610   \medskip As a general rule of thumb for good proof style, initial
       
   611   proof methods should either solve the goal completely, or constitute
       
   612   some well-understood reduction to new sub-goals.  Arbitrary
       
   613   automatic proof tools that are prone leave a large number of badly
       
   614   structured sub-goals are no help in continuing the proof document in
       
   615   an intelligible manner.
       
   616 
       
   617   Unless given explicitly by the user, the default initial method is
       
   618   @{method_ref (Pure) rule} (or its classical variant @{method_ref
       
   619   rule}), which applies a single standard elimination or introduction
       
   620   rule according to the topmost symbol involved.  There is no separate
       
   621   default terminal method.  Any remaining goals are always solved by
       
   622   assumption in the very last step.
       
   623 
       
   624   @{rail \<open>
       
   625     @@{command proof} method?
       
   626     ;
       
   627     @@{command qed} method?
       
   628     ;
       
   629     @@{command "by"} method method?
       
   630     ;
       
   631     (@@{command "."} | @@{command ".."} | @@{command sorry})
       
   632   \<close>}
       
   633 
       
   634   \begin{description}
       
   635   
       
   636   \item @{command "proof"}~@{text "m\<^sub>1"} refines the goal by proof
       
   637   method @{text "m\<^sub>1"}; facts for forward chaining are passed if so
       
   638   indicated by @{text "proof(chain)"} mode.
       
   639   
       
   640   \item @{command "qed"}~@{text "m\<^sub>2"} refines any remaining goals by
       
   641   proof method @{text "m\<^sub>2"} and concludes the sub-proof by assumption.
       
   642   If the goal had been @{text "show"} (or @{text "thus"}), some
       
   643   pending sub-goal is solved as well by the rule resulting from the
       
   644   result \emph{exported} into the enclosing goal context.  Thus @{text
       
   645   "qed"} may fail for two reasons: either @{text "m\<^sub>2"} fails, or the
       
   646   resulting rule does not fit to any pending goal\footnote{This
       
   647   includes any additional ``strong'' assumptions as introduced by
       
   648   @{command "assume"}.} of the enclosing context.  Debugging such a
       
   649   situation might involve temporarily changing @{command "show"} into
       
   650   @{command "have"}, or weakening the local context by replacing
       
   651   occurrences of @{command "assume"} by @{command "presume"}.
       
   652   
       
   653   \item @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} is a \emph{terminal
       
   654   proof}\index{proof!terminal}; it abbreviates @{command
       
   655   "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"}, but with
       
   656   backtracking across both methods.  Debugging an unsuccessful
       
   657   @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} command can be done by expanding its
       
   658   definition; in many cases @{command "proof"}~@{text "m\<^sub>1"} (or even
       
   659   @{text "apply"}~@{text "m\<^sub>1"}) is already sufficient to see the
       
   660   problem.
       
   661 
       
   662   \item ``@{command ".."}'' is a \emph{default
       
   663   proof}\index{proof!default}; it abbreviates @{command "by"}~@{text
       
   664   "rule"}.
       
   665 
       
   666   \item ``@{command "."}'' is a \emph{trivial
       
   667   proof}\index{proof!trivial}; it abbreviates @{command "by"}~@{text
       
   668   "this"}.
       
   669   
       
   670   \item @{command "sorry"} is a \emph{fake proof}\index{proof!fake}
       
   671   pretending to solve the pending claim without further ado.  This
       
   672   only works in interactive development, or if the @{attribute
       
   673   quick_and_dirty} is enabled.  Facts emerging from fake
       
   674   proofs are not the real thing.  Internally, the derivation object is
       
   675   tainted by an oracle invocation, which may be inspected via the
       
   676   theorem status \cite{isabelle-implementation}.
       
   677   
       
   678   The most important application of @{command "sorry"} is to support
       
   679   experimentation and top-down proof development.
       
   680 
       
   681   \end{description}
       
   682 *}
       
   683 
       
   684 
       
   685 subsection {* Fundamental methods and attributes \label{sec:pure-meth-att} *}
       
   686 
       
   687 text {*
       
   688   The following proof methods and attributes refer to basic logical
       
   689   operations of Isar.  Further methods and attributes are provided by
       
   690   several generic and object-logic specific tools and packages (see
       
   691   \chref{ch:gen-tools} and \partref{part:hol}).
       
   692 
       
   693   \begin{matharray}{rcl}
       
   694     @{command_def "print_rules"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\[0.5ex]
       
   695     @{method_def "-"} & : & @{text method} \\
       
   696     @{method_def "fact"} & : & @{text method} \\
       
   697     @{method_def "assumption"} & : & @{text method} \\
       
   698     @{method_def "this"} & : & @{text method} \\
       
   699     @{method_def (Pure) "rule"} & : & @{text method} \\
       
   700     @{attribute_def (Pure) "intro"} & : & @{text attribute} \\
       
   701     @{attribute_def (Pure) "elim"} & : & @{text attribute} \\
       
   702     @{attribute_def (Pure) "dest"} & : & @{text attribute} \\
       
   703     @{attribute_def (Pure) "rule"} & : & @{text attribute} \\[0.5ex]
       
   704     @{attribute_def "OF"} & : & @{text attribute} \\
       
   705     @{attribute_def "of"} & : & @{text attribute} \\
       
   706     @{attribute_def "where"} & : & @{text attribute} \\
       
   707   \end{matharray}
       
   708 
       
   709   @{rail \<open>
       
   710     @@{method fact} @{syntax thmrefs}?
       
   711     ;
       
   712     @@{method (Pure) rule} @{syntax thmrefs}?
       
   713     ;
       
   714     rulemod: ('intro' | 'elim' | 'dest')
       
   715       ((('!' | () | '?') @{syntax nat}?) | 'del') ':' @{syntax thmrefs}
       
   716     ;
       
   717     (@@{attribute intro} | @@{attribute elim} | @@{attribute dest})
       
   718       ('!' | () | '?') @{syntax nat}?
       
   719     ;
       
   720     @@{attribute (Pure) rule} 'del'
       
   721     ;
       
   722     @@{attribute OF} @{syntax thmrefs}
       
   723     ;
       
   724     @@{attribute of} @{syntax insts} ('concl' ':' @{syntax insts})? \<newline>
       
   725       (@'for' (@{syntax vars} + @'and'))?
       
   726     ;
       
   727     @@{attribute "where"}
       
   728       ((@{syntax name} | @{syntax var} | @{syntax typefree} | @{syntax typevar}) '='
       
   729       (@{syntax type} | @{syntax term}) * @'and') \<newline>
       
   730       (@'for' (@{syntax vars} + @'and'))?
       
   731   \<close>}
       
   732 
       
   733   \begin{description}
       
   734   
       
   735   \item @{command "print_rules"} prints rules declared via attributes
       
   736   @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute
       
   737   (Pure) dest} of Isabelle/Pure.
       
   738 
       
   739   See also the analogous @{command "print_claset"} command for similar
       
   740   rule declarations of the classical reasoner
       
   741   (\secref{sec:classical}).
       
   742 
       
   743   \item ``@{method "-"}'' (minus) does nothing but insert the forward
       
   744   chaining facts as premises into the goal.  Note that command
       
   745   @{command_ref "proof"} without any method actually performs a single
       
   746   reduction step using the @{method_ref (Pure) rule} method; thus a plain
       
   747   \emph{do-nothing} proof step would be ``@{command "proof"}~@{text
       
   748   "-"}'' rather than @{command "proof"} alone.
       
   749   
       
   750   \item @{method "fact"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} composes some fact from
       
   751   @{text "a\<^sub>1, \<dots>, a\<^sub>n"} (or implicitly from the current proof context)
       
   752   modulo unification of schematic type and term variables.  The rule
       
   753   structure is not taken into account, i.e.\ meta-level implication is
       
   754   considered atomic.  This is the same principle underlying literal
       
   755   facts (cf.\ \secref{sec:syn-att}): ``@{command "have"}~@{text
       
   756   "\<phi>"}~@{command "by"}~@{text fact}'' is equivalent to ``@{command
       
   757   "note"}~@{verbatim "`"}@{text \<phi>}@{verbatim "`"}'' provided that
       
   758   @{text "\<turnstile> \<phi>"} is an instance of some known @{text "\<turnstile> \<phi>"} in the
       
   759   proof context.
       
   760   
       
   761   \item @{method assumption} solves some goal by a single assumption
       
   762   step.  All given facts are guaranteed to participate in the
       
   763   refinement; this means there may be only 0 or 1 in the first place.
       
   764   Recall that @{command "qed"} (\secref{sec:proof-steps}) already
       
   765   concludes any remaining sub-goals by assumption, so structured
       
   766   proofs usually need not quote the @{method assumption} method at
       
   767   all.
       
   768   
       
   769   \item @{method this} applies all of the current facts directly as
       
   770   rules.  Recall that ``@{command "."}'' (dot) abbreviates ``@{command
       
   771   "by"}~@{text this}''.
       
   772   
       
   773   \item @{method (Pure) rule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} applies some rule given as
       
   774   argument in backward manner; facts are used to reduce the rule
       
   775   before applying it to the goal.  Thus @{method (Pure) rule} without facts
       
   776   is plain introduction, while with facts it becomes elimination.
       
   777   
       
   778   When no arguments are given, the @{method (Pure) rule} method tries to pick
       
   779   appropriate rules automatically, as declared in the current context
       
   780   using the @{attribute (Pure) intro}, @{attribute (Pure) elim},
       
   781   @{attribute (Pure) dest} attributes (see below).  This is the
       
   782   default behavior of @{command "proof"} and ``@{command ".."}'' 
       
   783   (double-dot) steps (see \secref{sec:proof-steps}).
       
   784   
       
   785   \item @{attribute (Pure) intro}, @{attribute (Pure) elim}, and
       
   786   @{attribute (Pure) dest} declare introduction, elimination, and
       
   787   destruct rules, to be used with method @{method (Pure) rule}, and similar
       
   788   tools.  Note that the latter will ignore rules declared with
       
   789   ``@{text "?"}'', while ``@{text "!"}''  are used most aggressively.
       
   790   
       
   791   The classical reasoner (see \secref{sec:classical}) introduces its
       
   792   own variants of these attributes; use qualified names to access the
       
   793   present versions of Isabelle/Pure, i.e.\ @{attribute (Pure)
       
   794   "Pure.intro"}.
       
   795   
       
   796   \item @{attribute (Pure) rule}~@{text del} undeclares introduction,
       
   797   elimination, or destruct rules.
       
   798 
       
   799   \item @{attribute OF}~@{text "a\<^sub>1 \<dots> a\<^sub>n"} applies some theorem to all
       
   800   of the given rules @{text "a\<^sub>1, \<dots>, a\<^sub>n"} in canonical right-to-left
       
   801   order, which means that premises stemming from the @{text "a\<^sub>i"}
       
   802   emerge in parallel in the result, without interfering with each
       
   803   other.  In many practical situations, the @{text "a\<^sub>i"} do not have
       
   804   premises themselves, so @{text "rule [OF a\<^sub>1 \<dots> a\<^sub>n]"} can be actually
       
   805   read as functional application (modulo unification).
       
   806 
       
   807   Argument positions may be effectively skipped by using ``@{text _}''
       
   808   (underscore), which refers to the propositional identity rule in the
       
   809   Pure theory.
       
   810   
       
   811   \item @{attribute of}~@{text "t\<^sub>1 \<dots> t\<^sub>n"} performs positional
       
   812   instantiation of term variables.  The terms @{text "t\<^sub>1, \<dots>, t\<^sub>n"} are
       
   813   substituted for any schematic variables occurring in a theorem from
       
   814   left to right; ``@{text _}'' (underscore) indicates to skip a
       
   815   position.  Arguments following a ``@{text "concl:"}'' specification
       
   816   refer to positions of the conclusion of a rule.
       
   817 
       
   818   An optional context of local variables @{text "\<FOR> x\<^sub>1 \<dots> x\<^sub>m"} may
       
   819   be specified: the instantiated theorem is exported, and these
       
   820   variables become schematic (usually with some shifting of indices).
       
   821   
       
   822   \item @{attribute "where"}~@{text "x\<^sub>1 = t\<^sub>1 \<AND> \<dots> x\<^sub>n = t\<^sub>n"}
       
   823   performs named instantiation of schematic type and term variables
       
   824   occurring in a theorem.  Schematic variables have to be specified on
       
   825   the left-hand side (e.g.\ @{text "?x1.3"}).  The question mark may
       
   826   be omitted if the variable name is a plain identifier without index.
       
   827   As type instantiations are inferred from term instantiations,
       
   828   explicit type instantiations are seldom necessary.
       
   829 
       
   830   An optional context of local variables @{text "\<FOR> x\<^sub>1 \<dots> x\<^sub>m"} may
       
   831   be specified as for @{attribute "of"} above.
       
   832 
       
   833   \end{description}
       
   834 *}
       
   835 
       
   836 
       
   837 subsection {* Emulating tactic scripts \label{sec:tactic-commands} *}
       
   838 
       
   839 text {*
       
   840   The Isar provides separate commands to accommodate tactic-style
       
   841   proof scripts within the same system.  While being outside the
       
   842   orthodox Isar proof language, these might come in handy for
       
   843   interactive exploration and debugging, or even actual tactical proof
       
   844   within new-style theories (to benefit from document preparation, for
       
   845   example).  See also \secref{sec:tactics} for actual tactics, that
       
   846   have been encapsulated as proof methods.  Proper proof methods may
       
   847   be used in scripts, too.
       
   848 
       
   849   \begin{matharray}{rcl}
       
   850     @{command_def "apply"}@{text "\<^sup>*"} & : & @{text "proof(prove) \<rightarrow> proof(prove)"} \\
       
   851     @{command_def "apply_end"}@{text "\<^sup>*"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
   852     @{command_def "done"}@{text "\<^sup>*"} & : & @{text "proof(prove) \<rightarrow> proof(state) | local_theory | theory"} \\
       
   853     @{command_def "defer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow> proof"} \\
       
   854     @{command_def "prefer"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow> proof"} \\
       
   855     @{command_def "back"}@{text "\<^sup>*"} & : & @{text "proof \<rightarrow> proof"} \\
       
   856   \end{matharray}
       
   857 
       
   858   @{rail \<open>
       
   859     ( @@{command apply} | @@{command apply_end} ) @{syntax method}
       
   860     ;
       
   861     @@{command defer} @{syntax nat}?
       
   862     ;
       
   863     @@{command prefer} @{syntax nat}
       
   864   \<close>}
       
   865 
       
   866   \begin{description}
       
   867 
       
   868   \item @{command "apply"}~@{text m} applies proof method @{text m} in
       
   869   initial position, but unlike @{command "proof"} it retains ``@{text
       
   870   "proof(prove)"}'' mode.  Thus consecutive method applications may be
       
   871   given just as in tactic scripts.
       
   872   
       
   873   Facts are passed to @{text m} as indicated by the goal's
       
   874   forward-chain mode, and are \emph{consumed} afterwards.  Thus any
       
   875   further @{command "apply"} command would always work in a purely
       
   876   backward manner.
       
   877   
       
   878   \item @{command "apply_end"}~@{text "m"} applies proof method @{text
       
   879   m} as if in terminal position.  Basically, this simulates a
       
   880   multi-step tactic script for @{command "qed"}, but may be given
       
   881   anywhere within the proof body.
       
   882   
       
   883   No facts are passed to @{text m} here.  Furthermore, the static
       
   884   context is that of the enclosing goal (as for actual @{command
       
   885   "qed"}).  Thus the proof method may not refer to any assumptions
       
   886   introduced in the current body, for example.
       
   887   
       
   888   \item @{command "done"} completes a proof script, provided that the
       
   889   current goal state is solved completely.  Note that actual
       
   890   structured proof commands (e.g.\ ``@{command "."}'' or @{command
       
   891   "sorry"}) may be used to conclude proof scripts as well.
       
   892 
       
   893   \item @{command "defer"}~@{text n} and @{command "prefer"}~@{text n}
       
   894   shuffle the list of pending goals: @{command "defer"} puts off
       
   895   sub-goal @{text n} to the end of the list (@{text "n = 1"} by
       
   896   default), while @{command "prefer"} brings sub-goal @{text n} to the
       
   897   front.
       
   898   
       
   899   \item @{command "back"} does back-tracking over the result sequence
       
   900   of the latest proof command.  Any proof command may return multiple
       
   901   results, and this command explores the possibilities step-by-step.
       
   902   It is mainly useful for experimentation and interactive exploration,
       
   903   and should be avoided in finished proofs.
       
   904   
       
   905   \end{description}
       
   906 
       
   907   Any proper Isar proof method may be used with tactic script commands
       
   908   such as @{command "apply"}.  A few additional emulations of actual
       
   909   tactics are provided as well; these would be never used in actual
       
   910   structured proofs, of course.
       
   911 *}
       
   912 
       
   913 
       
   914 subsection {* Defining proof methods *}
       
   915 
       
   916 text {*
       
   917   \begin{matharray}{rcl}
       
   918     @{command_def "method_setup"} & : & @{text "theory \<rightarrow> theory"} \\
       
   919   \end{matharray}
       
   920 
       
   921   @{rail \<open>
       
   922     @@{command method_setup} @{syntax name} '=' @{syntax text} @{syntax text}?
       
   923   \<close>}
       
   924 
       
   925   \begin{description}
       
   926 
       
   927   \item @{command "method_setup"}~@{text "name = text description"}
       
   928   defines a proof method in the current theory.  The given @{text
       
   929   "text"} has to be an ML expression of type
       
   930   @{ML_type "(Proof.context -> Proof.method) context_parser"}, cf.\
       
   931   basic parsers defined in structure @{ML_structure Args} and @{ML_structure
       
   932   Attrib}.  There are also combinators like @{ML METHOD} and @{ML
       
   933   SIMPLE_METHOD} to turn certain tactic forms into official proof
       
   934   methods; the primed versions refer to tactics with explicit goal
       
   935   addressing.
       
   936 
       
   937   Here are some example method definitions:
       
   938 
       
   939   \end{description}
       
   940 *}
       
   941 
       
   942   method_setup my_method1 = {*
       
   943     Scan.succeed (K (SIMPLE_METHOD' (fn i: int => no_tac)))
       
   944   *}  "my first method (without any arguments)"
       
   945 
       
   946   method_setup my_method2 = {*
       
   947     Scan.succeed (fn ctxt: Proof.context =>
       
   948       SIMPLE_METHOD' (fn i: int => no_tac))
       
   949   *}  "my second method (with context)"
       
   950 
       
   951   method_setup my_method3 = {*
       
   952     Attrib.thms >> (fn thms: thm list => fn ctxt: Proof.context =>
       
   953       SIMPLE_METHOD' (fn i: int => no_tac))
       
   954   *}  "my third method (with theorem arguments and context)"
       
   955 
       
   956 
       
   957 section {* Generalized elimination \label{sec:obtain} *}
       
   958 
       
   959 text {*
       
   960   \begin{matharray}{rcl}
       
   961     @{command_def "obtain"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
       
   962     @{command_def "guess"}@{text "\<^sup>*"} & : & @{text "proof(state) | proof(chain) \<rightarrow> proof(prove)"} \\
       
   963   \end{matharray}
       
   964 
       
   965   Generalized elimination means that additional elements with certain
       
   966   properties may be introduced in the current context, by virtue of a
       
   967   locally proven ``soundness statement''.  Technically speaking, the
       
   968   @{command "obtain"} language element is like a declaration of
       
   969   @{command "fix"} and @{command "assume"} (see also see
       
   970   \secref{sec:proof-context}), together with a soundness proof of its
       
   971   additional claim.  According to the nature of existential reasoning,
       
   972   assumptions get eliminated from any result exported from the context
       
   973   later, provided that the corresponding parameters do \emph{not}
       
   974   occur in the conclusion.
       
   975 
       
   976   @{rail \<open>
       
   977     @@{command obtain} @{syntax parname}? (@{syntax vars} + @'and')
       
   978       @'where' (@{syntax props} + @'and')
       
   979     ;
       
   980     @@{command guess} (@{syntax vars} + @'and')
       
   981   \<close>}
       
   982 
       
   983   The derived Isar command @{command "obtain"} is defined as follows
       
   984   (where @{text "b\<^sub>1, \<dots>, b\<^sub>k"} shall refer to (optional)
       
   985   facts indicated for forward chaining).
       
   986   \begin{matharray}{l}
       
   987     @{text "\<langle>using b\<^sub>1 \<dots> b\<^sub>k\<rangle>"}~~@{command "obtain"}~@{text "x\<^sub>1 \<dots> x\<^sub>m \<WHERE> a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n  \<langle>proof\<rangle> \<equiv>"} \\[1ex]
       
   988     \quad @{command "have"}~@{text "\<And>thesis. (\<And>x\<^sub>1 \<dots> x\<^sub>m. \<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> thesis) \<Longrightarrow> thesis"} \\
       
   989     \quad @{command "proof"}~@{method succeed} \\
       
   990     \qquad @{command "fix"}~@{text thesis} \\
       
   991     \qquad @{command "assume"}~@{text "that [Pure.intro?]: \<And>x\<^sub>1 \<dots> x\<^sub>m. \<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> thesis"} \\
       
   992     \qquad @{command "then"}~@{command "show"}~@{text thesis} \\
       
   993     \quad\qquad @{command "apply"}~@{text -} \\
       
   994     \quad\qquad @{command "using"}~@{text "b\<^sub>1 \<dots> b\<^sub>k  \<langle>proof\<rangle>"} \\
       
   995     \quad @{command "qed"} \\
       
   996     \quad @{command "fix"}~@{text "x\<^sub>1 \<dots> x\<^sub>m"}~@{command "assume"}@{text "\<^sup>* a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} \\
       
   997   \end{matharray}
       
   998 
       
   999   Typically, the soundness proof is relatively straight-forward, often
       
  1000   just by canonical automated tools such as ``@{command "by"}~@{text
       
  1001   simp}'' or ``@{command "by"}~@{text blast}''.  Accordingly, the
       
  1002   ``@{text that}'' reduction above is declared as simplification and
       
  1003   introduction rule.
       
  1004 
       
  1005   In a sense, @{command "obtain"} represents at the level of Isar
       
  1006   proofs what would be meta-logical existential quantifiers and
       
  1007   conjunctions.  This concept has a broad range of useful
       
  1008   applications, ranging from plain elimination (or introduction) of
       
  1009   object-level existential and conjunctions, to elimination over
       
  1010   results of symbolic evaluation of recursive definitions, for
       
  1011   example.  Also note that @{command "obtain"} without parameters acts
       
  1012   much like @{command "have"}, where the result is treated as a
       
  1013   genuine assumption.
       
  1014 
       
  1015   An alternative name to be used instead of ``@{text that}'' above may
       
  1016   be given in parentheses.
       
  1017 
       
  1018   \medskip The improper variant @{command "guess"} is similar to
       
  1019   @{command "obtain"}, but derives the obtained statement from the
       
  1020   course of reasoning!  The proof starts with a fixed goal @{text
       
  1021   thesis}.  The subsequent proof may refine this to anything of the
       
  1022   form like @{text "\<And>x\<^sub>1 \<dots> x\<^sub>m. \<phi>\<^sub>1 \<Longrightarrow> \<dots>
       
  1023   \<phi>\<^sub>n \<Longrightarrow> thesis"}, but must not introduce new subgoals.  The
       
  1024   final goal state is then used as reduction rule for the obtain
       
  1025   scheme described above.  Obtained parameters @{text "x\<^sub>1, \<dots>,
       
  1026   x\<^sub>m"} are marked as internal by default, which prevents the
       
  1027   proof context from being polluted by ad-hoc variables.  The variable
       
  1028   names and type constraints given as arguments for @{command "guess"}
       
  1029   specify a prefix of obtained parameters explicitly in the text.
       
  1030 
       
  1031   It is important to note that the facts introduced by @{command
       
  1032   "obtain"} and @{command "guess"} may not be polymorphic: any
       
  1033   type-variables occurring here are fixed in the present context!
       
  1034 *}
       
  1035 
       
  1036 
       
  1037 section {* Calculational reasoning \label{sec:calculation} *}
       
  1038 
       
  1039 text {*
       
  1040   \begin{matharray}{rcl}
       
  1041     @{command_def "also"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
  1042     @{command_def "finally"} & : & @{text "proof(state) \<rightarrow> proof(chain)"} \\
       
  1043     @{command_def "moreover"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
  1044     @{command_def "ultimately"} & : & @{text "proof(state) \<rightarrow> proof(chain)"} \\
       
  1045     @{command_def "print_trans_rules"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
       
  1046     @{attribute trans} & : & @{text attribute} \\
       
  1047     @{attribute sym} & : & @{text attribute} \\
       
  1048     @{attribute symmetric} & : & @{text attribute} \\
       
  1049   \end{matharray}
       
  1050 
       
  1051   Calculational proof is forward reasoning with implicit application
       
  1052   of transitivity rules (such those of @{text "="}, @{text "\<le>"},
       
  1053   @{text "<"}).  Isabelle/Isar maintains an auxiliary fact register
       
  1054   @{fact_ref calculation} for accumulating results obtained by
       
  1055   transitivity composed with the current result.  Command @{command
       
  1056   "also"} updates @{fact calculation} involving @{fact this}, while
       
  1057   @{command "finally"} exhibits the final @{fact calculation} by
       
  1058   forward chaining towards the next goal statement.  Both commands
       
  1059   require valid current facts, i.e.\ may occur only after commands
       
  1060   that produce theorems such as @{command "assume"}, @{command
       
  1061   "note"}, or some finished proof of @{command "have"}, @{command
       
  1062   "show"} etc.  The @{command "moreover"} and @{command "ultimately"}
       
  1063   commands are similar to @{command "also"} and @{command "finally"},
       
  1064   but only collect further results in @{fact calculation} without
       
  1065   applying any rules yet.
       
  1066 
       
  1067   Also note that the implicit term abbreviation ``@{text "\<dots>"}'' has
       
  1068   its canonical application with calculational proofs.  It refers to
       
  1069   the argument of the preceding statement. (The argument of a curried
       
  1070   infix expression happens to be its right-hand side.)
       
  1071 
       
  1072   Isabelle/Isar calculations are implicitly subject to block structure
       
  1073   in the sense that new threads of calculational reasoning are
       
  1074   commenced for any new block (as opened by a local goal, for
       
  1075   example).  This means that, apart from being able to nest
       
  1076   calculations, there is no separate \emph{begin-calculation} command
       
  1077   required.
       
  1078 
       
  1079   \medskip The Isar calculation proof commands may be defined as
       
  1080   follows:\footnote{We suppress internal bookkeeping such as proper
       
  1081   handling of block-structure.}
       
  1082 
       
  1083   \begin{matharray}{rcl}
       
  1084     @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\
       
  1085     @{command "also"}@{text "\<^sub>n+1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex]
       
  1086     @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex]
       
  1087     @{command "moreover"} & \equiv & @{command "note"}~@{text "calculation = calculation this"} \\
       
  1088     @{command "ultimately"} & \equiv & @{command "moreover"}~@{command "from"}~@{text calculation} \\
       
  1089   \end{matharray}
       
  1090 
       
  1091   @{rail \<open>
       
  1092     (@@{command also} | @@{command finally}) ('(' @{syntax thmrefs} ')')?
       
  1093     ;
       
  1094     @@{attribute trans} (() | 'add' | 'del')
       
  1095   \<close>}
       
  1096 
       
  1097   \begin{description}
       
  1098 
       
  1099   \item @{command "also"}~@{text "(a\<^sub>1 \<dots> a\<^sub>n)"} maintains the auxiliary
       
  1100   @{fact calculation} register as follows.  The first occurrence of
       
  1101   @{command "also"} in some calculational thread initializes @{fact
       
  1102   calculation} by @{fact this}. Any subsequent @{command "also"} on
       
  1103   the same level of block-structure updates @{fact calculation} by
       
  1104   some transitivity rule applied to @{fact calculation} and @{fact
       
  1105   this} (in that order).  Transitivity rules are picked from the
       
  1106   current context, unless alternative rules are given as explicit
       
  1107   arguments.
       
  1108 
       
  1109   \item @{command "finally"}~@{text "(a\<^sub>1 \<dots> a\<^sub>n)"} maintaining @{fact
       
  1110   calculation} in the same way as @{command "also"}, and concludes the
       
  1111   current calculational thread.  The final result is exhibited as fact
       
  1112   for forward chaining towards the next goal. Basically, @{command
       
  1113   "finally"} just abbreviates @{command "also"}~@{command
       
  1114   "from"}~@{fact calculation}.  Typical idioms for concluding
       
  1115   calculational proofs are ``@{command "finally"}~@{command
       
  1116   "show"}~@{text ?thesis}~@{command "."}'' and ``@{command
       
  1117   "finally"}~@{command "have"}~@{text \<phi>}~@{command "."}''.
       
  1118 
       
  1119   \item @{command "moreover"} and @{command "ultimately"} are
       
  1120   analogous to @{command "also"} and @{command "finally"}, but collect
       
  1121   results only, without applying rules.
       
  1122 
       
  1123   \item @{command "print_trans_rules"} prints the list of transitivity
       
  1124   rules (for calculational commands @{command "also"} and @{command
       
  1125   "finally"}) and symmetry rules (for the @{attribute symmetric}
       
  1126   operation and single step elimination patters) of the current
       
  1127   context.
       
  1128 
       
  1129   \item @{attribute trans} declares theorems as transitivity rules.
       
  1130 
       
  1131   \item @{attribute sym} declares symmetry rules, as well as
       
  1132   @{attribute "Pure.elim"}@{text "?"} rules.
       
  1133 
       
  1134   \item @{attribute symmetric} resolves a theorem with some rule
       
  1135   declared as @{attribute sym} in the current context.  For example,
       
  1136   ``@{command "assume"}~@{text "[symmetric]: x = y"}'' produces a
       
  1137   swapped fact derived from that assumption.
       
  1138 
       
  1139   In structured proof texts it is often more appropriate to use an
       
  1140   explicit single-step elimination proof, such as ``@{command
       
  1141   "assume"}~@{text "x = y"}~@{command "then"}~@{command "have"}~@{text
       
  1142   "y = x"}~@{command ".."}''.
       
  1143 
       
  1144   \end{description}
       
  1145 *}
       
  1146 
       
  1147 
       
  1148 section {* Proof by cases and induction \label{sec:cases-induct} *}
       
  1149 
       
  1150 subsection {* Rule contexts *}
       
  1151 
       
  1152 text {*
       
  1153   \begin{matharray}{rcl}
       
  1154     @{command_def "case"} & : & @{text "proof(state) \<rightarrow> proof(state)"} \\
       
  1155     @{command_def "print_cases"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
       
  1156     @{attribute_def case_names} & : & @{text attribute} \\
       
  1157     @{attribute_def case_conclusion} & : & @{text attribute} \\
       
  1158     @{attribute_def params} & : & @{text attribute} \\
       
  1159     @{attribute_def consumes} & : & @{text attribute} \\
       
  1160   \end{matharray}
       
  1161 
       
  1162   The puristic way to build up Isar proof contexts is by explicit
       
  1163   language elements like @{command "fix"}, @{command "assume"},
       
  1164   @{command "let"} (see \secref{sec:proof-context}).  This is adequate
       
  1165   for plain natural deduction, but easily becomes unwieldy in concrete
       
  1166   verification tasks, which typically involve big induction rules with
       
  1167   several cases.
       
  1168 
       
  1169   The @{command "case"} command provides a shorthand to refer to a
       
  1170   local context symbolically: certain proof methods provide an
       
  1171   environment of named ``cases'' of the form @{text "c: x\<^sub>1, \<dots>,
       
  1172   x\<^sub>m, \<phi>\<^sub>1, \<dots>, \<phi>\<^sub>n"}; the effect of ``@{command
       
  1173   "case"}~@{text c}'' is then equivalent to ``@{command "fix"}~@{text
       
  1174   "x\<^sub>1 \<dots> x\<^sub>m"}~@{command "assume"}~@{text "c: \<phi>\<^sub>1 \<dots>
       
  1175   \<phi>\<^sub>n"}''.  Term bindings may be covered as well, notably
       
  1176   @{variable ?case} for the main conclusion.
       
  1177 
       
  1178   By default, the ``terminology'' @{text "x\<^sub>1, \<dots>, x\<^sub>m"} of
       
  1179   a case value is marked as hidden, i.e.\ there is no way to refer to
       
  1180   such parameters in the subsequent proof text.  After all, original
       
  1181   rule parameters stem from somewhere outside of the current proof
       
  1182   text.  By using the explicit form ``@{command "case"}~@{text "(c
       
  1183   y\<^sub>1 \<dots> y\<^sub>m)"}'' instead, the proof author is able to
       
  1184   chose local names that fit nicely into the current context.
       
  1185 
       
  1186   \medskip It is important to note that proper use of @{command
       
  1187   "case"} does not provide means to peek at the current goal state,
       
  1188   which is not directly observable in Isar!  Nonetheless, goal
       
  1189   refinement commands do provide named cases @{text "goal\<^sub>i"}
       
  1190   for each subgoal @{text "i = 1, \<dots>, n"} of the resulting goal state.
       
  1191   Using this extra feature requires great care, because some bits of
       
  1192   the internal tactical machinery intrude the proof text.  In
       
  1193   particular, parameter names stemming from the left-over of automated
       
  1194   reasoning tools are usually quite unpredictable.
       
  1195 
       
  1196   Under normal circumstances, the text of cases emerge from standard
       
  1197   elimination or induction rules, which in turn are derived from
       
  1198   previous theory specifications in a canonical way (say from
       
  1199   @{command "inductive"} definitions).
       
  1200 
       
  1201   \medskip Proper cases are only available if both the proof method
       
  1202   and the rules involved support this.  By using appropriate
       
  1203   attributes, case names, conclusions, and parameters may be also
       
  1204   declared by hand.  Thus variant versions of rules that have been
       
  1205   derived manually become ready to use in advanced case analysis
       
  1206   later.
       
  1207 
       
  1208   @{rail \<open>
       
  1209     @@{command case} (caseref | '(' caseref (('_' | @{syntax name}) *) ')')
       
  1210     ;
       
  1211     caseref: nameref attributes?
       
  1212     ;
       
  1213 
       
  1214     @@{attribute case_names} ((@{syntax name} ( '[' (('_' | @{syntax name}) +) ']' ) ? ) +)
       
  1215     ;
       
  1216     @@{attribute case_conclusion} @{syntax name} (@{syntax name} * )
       
  1217     ;
       
  1218     @@{attribute params} ((@{syntax name} * ) + @'and')
       
  1219     ;
       
  1220     @@{attribute consumes} @{syntax int}?
       
  1221   \<close>}
       
  1222 
       
  1223   \begin{description}
       
  1224   
       
  1225   \item @{command "case"}~@{text "(c x\<^sub>1 \<dots> x\<^sub>m)"} invokes a named local
       
  1226   context @{text "c: x\<^sub>1, \<dots>, x\<^sub>m, \<phi>\<^sub>1, \<dots>, \<phi>\<^sub>m"}, as provided by an
       
  1227   appropriate proof method (such as @{method_ref cases} and
       
  1228   @{method_ref induct}).  The command ``@{command "case"}~@{text "(c
       
  1229   x\<^sub>1 \<dots> x\<^sub>m)"}'' abbreviates ``@{command "fix"}~@{text "x\<^sub>1 \<dots>
       
  1230   x\<^sub>m"}~@{command "assume"}~@{text "c: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"}''.
       
  1231 
       
  1232   \item @{command "print_cases"} prints all local contexts of the
       
  1233   current state, using Isar proof language notation.
       
  1234   
       
  1235   \item @{attribute case_names}~@{text "c\<^sub>1 \<dots> c\<^sub>k"} declares names for
       
  1236   the local contexts of premises of a theorem; @{text "c\<^sub>1, \<dots>, c\<^sub>k"}
       
  1237   refers to the \emph{prefix} of the list of premises. Each of the
       
  1238   cases @{text "c\<^sub>i"} can be of the form @{text "c[h\<^sub>1 \<dots> h\<^sub>n]"} where
       
  1239   the @{text "h\<^sub>1 \<dots> h\<^sub>n"} are the names of the hypotheses in case @{text "c\<^sub>i"}
       
  1240   from left to right.
       
  1241   
       
  1242   \item @{attribute case_conclusion}~@{text "c d\<^sub>1 \<dots> d\<^sub>k"} declares
       
  1243   names for the conclusions of a named premise @{text c}; here @{text
       
  1244   "d\<^sub>1, \<dots>, d\<^sub>k"} refers to the prefix of arguments of a logical formula
       
  1245   built by nesting a binary connective (e.g.\ @{text "\<or>"}).
       
  1246   
       
  1247   Note that proof methods such as @{method induct} and @{method
       
  1248   coinduct} already provide a default name for the conclusion as a
       
  1249   whole.  The need to name subformulas only arises with cases that
       
  1250   split into several sub-cases, as in common co-induction rules.
       
  1251 
       
  1252   \item @{attribute params}~@{text "p\<^sub>1 \<dots> p\<^sub>m \<AND> \<dots> q\<^sub>1 \<dots> q\<^sub>n"} renames
       
  1253   the innermost parameters of premises @{text "1, \<dots>, n"} of some
       
  1254   theorem.  An empty list of names may be given to skip positions,
       
  1255   leaving the present parameters unchanged.
       
  1256   
       
  1257   Note that the default usage of case rules does \emph{not} directly
       
  1258   expose parameters to the proof context.
       
  1259   
       
  1260   \item @{attribute consumes}~@{text n} declares the number of ``major
       
  1261   premises'' of a rule, i.e.\ the number of facts to be consumed when
       
  1262   it is applied by an appropriate proof method.  The default value of
       
  1263   @{attribute consumes} is @{text "n = 1"}, which is appropriate for
       
  1264   the usual kind of cases and induction rules for inductive sets (cf.\
       
  1265   \secref{sec:hol-inductive}).  Rules without any @{attribute
       
  1266   consumes} declaration given are treated as if @{attribute
       
  1267   consumes}~@{text 0} had been specified.
       
  1268 
       
  1269   A negative @{text n} is interpreted relatively to the total number
       
  1270   of premises of the rule in the target context.  Thus its absolute
       
  1271   value specifies the remaining number of premises, after subtracting
       
  1272   the prefix of major premises as indicated above. This form of
       
  1273   declaration has the technical advantage of being stable under more
       
  1274   morphisms, notably those that export the result from a nested
       
  1275   @{command_ref context} with additional assumptions.
       
  1276 
       
  1277   Note that explicit @{attribute consumes} declarations are only
       
  1278   rarely needed; this is already taken care of automatically by the
       
  1279   higher-level @{attribute cases}, @{attribute induct}, and
       
  1280   @{attribute coinduct} declarations.
       
  1281 
       
  1282   \end{description}
       
  1283 *}
       
  1284 
       
  1285 
       
  1286 subsection {* Proof methods *}
       
  1287 
       
  1288 text {*
       
  1289   \begin{matharray}{rcl}
       
  1290     @{method_def cases} & : & @{text method} \\
       
  1291     @{method_def induct} & : & @{text method} \\
       
  1292     @{method_def induction} & : & @{text method} \\
       
  1293     @{method_def coinduct} & : & @{text method} \\
       
  1294   \end{matharray}
       
  1295 
       
  1296   The @{method cases}, @{method induct}, @{method induction},
       
  1297   and @{method coinduct}
       
  1298   methods provide a uniform interface to common proof techniques over
       
  1299   datatypes, inductive predicates (or sets), recursive functions etc.
       
  1300   The corresponding rules may be specified and instantiated in a
       
  1301   casual manner.  Furthermore, these methods provide named local
       
  1302   contexts that may be invoked via the @{command "case"} proof command
       
  1303   within the subsequent proof text.  This accommodates compact proof
       
  1304   texts even when reasoning about large specifications.
       
  1305 
       
  1306   The @{method induct} method also provides some additional
       
  1307   infrastructure in order to be applicable to structure statements
       
  1308   (either using explicit meta-level connectives, or including facts
       
  1309   and parameters separately).  This avoids cumbersome encoding of
       
  1310   ``strengthened'' inductive statements within the object-logic.
       
  1311 
       
  1312   Method @{method induction} differs from @{method induct} only in
       
  1313   the names of the facts in the local context invoked by the @{command "case"}
       
  1314   command.
       
  1315 
       
  1316   @{rail \<open>
       
  1317     @@{method cases} ('(' 'no_simp' ')')? \<newline>
       
  1318       (@{syntax insts} * @'and') rule?
       
  1319     ;
       
  1320     (@@{method induct} | @@{method induction})
       
  1321       ('(' 'no_simp' ')')? (definsts * @'and') \<newline> arbitrary? taking? rule?
       
  1322     ;
       
  1323     @@{method coinduct} @{syntax insts} taking rule?
       
  1324     ;
       
  1325 
       
  1326     rule: ('type' | 'pred' | 'set') ':' (@{syntax nameref} +) | 'rule' ':' (@{syntax thmref} +)
       
  1327     ;
       
  1328     definst: @{syntax name} ('==' | '\<equiv>') @{syntax term} | '(' @{syntax term} ')' | @{syntax inst}
       
  1329     ;
       
  1330     definsts: ( definst * )
       
  1331     ;
       
  1332     arbitrary: 'arbitrary' ':' ((@{syntax term} * ) @'and' +)
       
  1333     ;
       
  1334     taking: 'taking' ':' @{syntax insts}
       
  1335   \<close>}
       
  1336 
       
  1337   \begin{description}
       
  1338 
       
  1339   \item @{method cases}~@{text "insts R"} applies method @{method
       
  1340   rule} with an appropriate case distinction theorem, instantiated to
       
  1341   the subjects @{text insts}.  Symbolic case names are bound according
       
  1342   to the rule's local contexts.
       
  1343 
       
  1344   The rule is determined as follows, according to the facts and
       
  1345   arguments passed to the @{method cases} method:
       
  1346 
       
  1347   \medskip
       
  1348   \begin{tabular}{llll}
       
  1349     facts           &                 & arguments   & rule \\\hline
       
  1350                     & @{method cases} &             & classical case split \\
       
  1351                     & @{method cases} & @{text t}   & datatype exhaustion (type of @{text t}) \\
       
  1352     @{text "\<turnstile> A t"} & @{method cases} & @{text "\<dots>"} & inductive predicate/set elimination (of @{text A}) \\
       
  1353     @{text "\<dots>"}     & @{method cases} & @{text "\<dots> rule: R"} & explicit rule @{text R} \\
       
  1354   \end{tabular}
       
  1355   \medskip
       
  1356 
       
  1357   Several instantiations may be given, referring to the \emph{suffix}
       
  1358   of premises of the case rule; within each premise, the \emph{prefix}
       
  1359   of variables is instantiated.  In most situations, only a single
       
  1360   term needs to be specified; this refers to the first variable of the
       
  1361   last premise (it is usually the same for all cases).  The @{text
       
  1362   "(no_simp)"} option can be used to disable pre-simplification of
       
  1363   cases (see the description of @{method induct} below for details).
       
  1364 
       
  1365   \item @{method induct}~@{text "insts R"} and
       
  1366   @{method induction}~@{text "insts R"} are analogous to the
       
  1367   @{method cases} method, but refer to induction rules, which are
       
  1368   determined as follows:
       
  1369 
       
  1370   \medskip
       
  1371   \begin{tabular}{llll}
       
  1372     facts           &                  & arguments            & rule \\\hline
       
  1373                     & @{method induct} & @{text "P x"}        & datatype induction (type of @{text x}) \\
       
  1374     @{text "\<turnstile> A x"} & @{method induct} & @{text "\<dots>"}          & predicate/set induction (of @{text A}) \\
       
  1375     @{text "\<dots>"}     & @{method induct} & @{text "\<dots> rule: R"} & explicit rule @{text R} \\
       
  1376   \end{tabular}
       
  1377   \medskip
       
  1378   
       
  1379   Several instantiations may be given, each referring to some part of
       
  1380   a mutual inductive definition or datatype --- only related partial
       
  1381   induction rules may be used together, though.  Any of the lists of
       
  1382   terms @{text "P, x, \<dots>"} refers to the \emph{suffix} of variables
       
  1383   present in the induction rule.  This enables the writer to specify
       
  1384   only induction variables, or both predicates and variables, for
       
  1385   example.
       
  1386 
       
  1387   Instantiations may be definitional: equations @{text "x \<equiv> t"}
       
  1388   introduce local definitions, which are inserted into the claim and
       
  1389   discharged after applying the induction rule.  Equalities reappear
       
  1390   in the inductive cases, but have been transformed according to the
       
  1391   induction principle being involved here.  In order to achieve
       
  1392   practically useful induction hypotheses, some variables occurring in
       
  1393   @{text t} need to be fixed (see below).  Instantiations of the form
       
  1394   @{text t}, where @{text t} is not a variable, are taken as a
       
  1395   shorthand for \mbox{@{text "x \<equiv> t"}}, where @{text x} is a fresh
       
  1396   variable. If this is not intended, @{text t} has to be enclosed in
       
  1397   parentheses.  By default, the equalities generated by definitional
       
  1398   instantiations are pre-simplified using a specific set of rules,
       
  1399   usually consisting of distinctness and injectivity theorems for
       
  1400   datatypes. This pre-simplification may cause some of the parameters
       
  1401   of an inductive case to disappear, or may even completely delete
       
  1402   some of the inductive cases, if one of the equalities occurring in
       
  1403   their premises can be simplified to @{text False}.  The @{text
       
  1404   "(no_simp)"} option can be used to disable pre-simplification.
       
  1405   Additional rules to be used in pre-simplification can be declared
       
  1406   using the @{attribute_def induct_simp} attribute.
       
  1407 
       
  1408   The optional ``@{text "arbitrary: x\<^sub>1 \<dots> x\<^sub>m"}''
       
  1409   specification generalizes variables @{text "x\<^sub>1, \<dots>,
       
  1410   x\<^sub>m"} of the original goal before applying induction.  One can
       
  1411   separate variables by ``@{text "and"}'' to generalize them in other
       
  1412   goals then the first. Thus induction hypotheses may become
       
  1413   sufficiently general to get the proof through.  Together with
       
  1414   definitional instantiations, one may effectively perform induction
       
  1415   over expressions of a certain structure.
       
  1416   
       
  1417   The optional ``@{text "taking: t\<^sub>1 \<dots> t\<^sub>n"}''
       
  1418   specification provides additional instantiations of a prefix of
       
  1419   pending variables in the rule.  Such schematic induction rules
       
  1420   rarely occur in practice, though.
       
  1421 
       
  1422   \item @{method coinduct}~@{text "inst R"} is analogous to the
       
  1423   @{method induct} method, but refers to coinduction rules, which are
       
  1424   determined as follows:
       
  1425 
       
  1426   \medskip
       
  1427   \begin{tabular}{llll}
       
  1428     goal          &                    & arguments & rule \\\hline
       
  1429                   & @{method coinduct} & @{text x} & type coinduction (type of @{text x}) \\
       
  1430     @{text "A x"} & @{method coinduct} & @{text "\<dots>"} & predicate/set coinduction (of @{text A}) \\
       
  1431     @{text "\<dots>"}   & @{method coinduct} & @{text "\<dots> rule: R"} & explicit rule @{text R} \\
       
  1432   \end{tabular}
       
  1433   
       
  1434   Coinduction is the dual of induction.  Induction essentially
       
  1435   eliminates @{text "A x"} towards a generic result @{text "P x"},
       
  1436   while coinduction introduces @{text "A x"} starting with @{text "B
       
  1437   x"}, for a suitable ``bisimulation'' @{text B}.  The cases of a
       
  1438   coinduct rule are typically named after the predicates or sets being
       
  1439   covered, while the conclusions consist of several alternatives being
       
  1440   named after the individual destructor patterns.
       
  1441   
       
  1442   The given instantiation refers to the \emph{suffix} of variables
       
  1443   occurring in the rule's major premise, or conclusion if unavailable.
       
  1444   An additional ``@{text "taking: t\<^sub>1 \<dots> t\<^sub>n"}''
       
  1445   specification may be required in order to specify the bisimulation
       
  1446   to be used in the coinduction step.
       
  1447 
       
  1448   \end{description}
       
  1449 
       
  1450   Above methods produce named local contexts, as determined by the
       
  1451   instantiated rule as given in the text.  Beyond that, the @{method
       
  1452   induct} and @{method coinduct} methods guess further instantiations
       
  1453   from the goal specification itself.  Any persisting unresolved
       
  1454   schematic variables of the resulting rule will render the the
       
  1455   corresponding case invalid.  The term binding @{variable ?case} for
       
  1456   the conclusion will be provided with each case, provided that term
       
  1457   is fully specified.
       
  1458 
       
  1459   The @{command "print_cases"} command prints all named cases present
       
  1460   in the current proof state.
       
  1461 
       
  1462   \medskip Despite the additional infrastructure, both @{method cases}
       
  1463   and @{method coinduct} merely apply a certain rule, after
       
  1464   instantiation, while conforming due to the usual way of monotonic
       
  1465   natural deduction: the context of a structured statement @{text
       
  1466   "\<And>x\<^sub>1 \<dots> x\<^sub>m. \<phi>\<^sub>1 \<Longrightarrow> \<dots> \<phi>\<^sub>n \<Longrightarrow> \<dots>"}
       
  1467   reappears unchanged after the case split.
       
  1468 
       
  1469   The @{method induct} method is fundamentally different in this
       
  1470   respect: the meta-level structure is passed through the
       
  1471   ``recursive'' course involved in the induction.  Thus the original
       
  1472   statement is basically replaced by separate copies, corresponding to
       
  1473   the induction hypotheses and conclusion; the original goal context
       
  1474   is no longer available.  Thus local assumptions, fixed parameters
       
  1475   and definitions effectively participate in the inductive rephrasing
       
  1476   of the original statement.
       
  1477 
       
  1478   In @{method induct} proofs, local assumptions introduced by cases are split
       
  1479   into two different kinds: @{text hyps} stemming from the rule and
       
  1480   @{text prems} from the goal statement.  This is reflected in the
       
  1481   extracted cases accordingly, so invoking ``@{command "case"}~@{text
       
  1482   c}'' will provide separate facts @{text c.hyps} and @{text c.prems},
       
  1483   as well as fact @{text c} to hold the all-inclusive list.
       
  1484 
       
  1485   In @{method induction} proofs, local assumptions introduced by cases are
       
  1486   split into three different kinds: @{text IH}, the induction hypotheses,
       
  1487   @{text hyps}, the remaining hypotheses stemming from the rule, and
       
  1488   @{text prems}, the assumptions from the goal statement. The names are
       
  1489   @{text c.IH}, @{text c.hyps} and @{text c.prems}, as above.
       
  1490 
       
  1491 
       
  1492   \medskip Facts presented to either method are consumed according to
       
  1493   the number of ``major premises'' of the rule involved, which is
       
  1494   usually 0 for plain cases and induction rules of datatypes etc.\ and
       
  1495   1 for rules of inductive predicates or sets and the like.  The
       
  1496   remaining facts are inserted into the goal verbatim before the
       
  1497   actual @{text cases}, @{text induct}, or @{text coinduct} rule is
       
  1498   applied.
       
  1499 *}
       
  1500 
       
  1501 
       
  1502 subsection {* Declaring rules *}
       
  1503 
       
  1504 text {*
       
  1505   \begin{matharray}{rcl}
       
  1506     @{command_def "print_induct_rules"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\
       
  1507     @{attribute_def cases} & : & @{text attribute} \\
       
  1508     @{attribute_def induct} & : & @{text attribute} \\
       
  1509     @{attribute_def coinduct} & : & @{text attribute} \\
       
  1510   \end{matharray}
       
  1511 
       
  1512   @{rail \<open>
       
  1513     @@{attribute cases} spec
       
  1514     ;
       
  1515     @@{attribute induct} spec
       
  1516     ;
       
  1517     @@{attribute coinduct} spec
       
  1518     ;
       
  1519 
       
  1520     spec: (('type' | 'pred' | 'set') ':' @{syntax nameref}) | 'del'
       
  1521   \<close>}
       
  1522 
       
  1523   \begin{description}
       
  1524 
       
  1525   \item @{command "print_induct_rules"} prints cases and induct rules
       
  1526   for predicates (or sets) and types of the current context.
       
  1527 
       
  1528   \item @{attribute cases}, @{attribute induct}, and @{attribute
       
  1529   coinduct} (as attributes) declare rules for reasoning about
       
  1530   (co)inductive predicates (or sets) and types, using the
       
  1531   corresponding methods of the same name.  Certain definitional
       
  1532   packages of object-logics usually declare emerging cases and
       
  1533   induction rules as expected, so users rarely need to intervene.
       
  1534 
       
  1535   Rules may be deleted via the @{text "del"} specification, which
       
  1536   covers all of the @{text "type"}/@{text "pred"}/@{text "set"}
       
  1537   sub-categories simultaneously.  For example, @{attribute
       
  1538   cases}~@{text del} removes any @{attribute cases} rules declared for
       
  1539   some type, predicate, or set.
       
  1540   
       
  1541   Manual rule declarations usually refer to the @{attribute
       
  1542   case_names} and @{attribute params} attributes to adjust names of
       
  1543   cases and parameters of a rule; the @{attribute consumes}
       
  1544   declaration is taken care of automatically: @{attribute
       
  1545   consumes}~@{text 0} is specified for ``type'' rules and @{attribute
       
  1546   consumes}~@{text 1} for ``predicate'' / ``set'' rules.
       
  1547 
       
  1548   \end{description}
       
  1549 *}
       
  1550 
       
  1551 end