doc-src/IsarRef/Thy/Proof.thy
author wenzelm
Wed, 14 May 2008 20:31:17 +0200
changeset 26894 1120f6cc10b0
parent 26888 9942cd184c48
child 26901 d1694ef6e7a7
permissions -rw-r--r--
proper checking of various Isar elements;
wenzelm@26869
     1
(* $Id$ *)
wenzelm@26869
     2
wenzelm@26869
     3
theory Proof
wenzelm@26869
     4
imports Main
wenzelm@26869
     5
begin
wenzelm@26869
     6
wenzelm@26869
     7
chapter {* Proofs *}
wenzelm@26869
     8
wenzelm@26870
     9
text {*
wenzelm@26870
    10
  Proof commands perform transitions of Isar/VM machine
wenzelm@26870
    11
  configurations, which are block-structured, consisting of a stack of
wenzelm@26870
    12
  nodes with three main components: logical proof context, current
wenzelm@26870
    13
  facts, and open goals.  Isar/VM transitions are \emph{typed}
wenzelm@26870
    14
  according to the following three different modes of operation:
wenzelm@26870
    15
wenzelm@26870
    16
  \begin{descr}
wenzelm@26870
    17
wenzelm@26870
    18
  \item [@{text "proof(prove)"}] means that a new goal has just been
wenzelm@26870
    19
  stated that is now to be \emph{proven}; the next command may refine
wenzelm@26870
    20
  it by some proof method, and enter a sub-proof to establish the
wenzelm@26870
    21
  actual result.
wenzelm@26870
    22
wenzelm@26870
    23
  \item [@{text "proof(state)"}] is like a nested theory mode: the
wenzelm@26870
    24
  context may be augmented by \emph{stating} additional assumptions,
wenzelm@26870
    25
  intermediate results etc.
wenzelm@26870
    26
wenzelm@26870
    27
  \item [@{text "proof(chain)"}] is intermediate between @{text
wenzelm@26870
    28
  "proof(state)"} and @{text "proof(prove)"}: existing facts (i.e.\
wenzelm@26870
    29
  the contents of the special ``@{fact_ref this}'' register) have been
wenzelm@26870
    30
  just picked up in order to be used when refining the goal claimed
wenzelm@26870
    31
  next.
wenzelm@26870
    32
wenzelm@26870
    33
  \end{descr}
wenzelm@26870
    34
wenzelm@26870
    35
  The proof mode indicator may be read as a verb telling the writer
wenzelm@26870
    36
  what kind of operation may be performed next.  The corresponding
wenzelm@26870
    37
  typings of proof commands restricts the shape of well-formed proof
wenzelm@26870
    38
  texts to particular command sequences.  So dynamic arrangements of
wenzelm@26870
    39
  commands eventually turn out as static texts of a certain structure.
wenzelm@26870
    40
  \Appref{ap:refcard} gives a simplified grammar of the overall
wenzelm@26870
    41
  (extensible) language emerging that way.
wenzelm@26870
    42
*}
wenzelm@26870
    43
wenzelm@26870
    44
wenzelm@26870
    45
section {* Context elements \label{sec:proof-context} *}
wenzelm@26870
    46
wenzelm@26870
    47
text {*
wenzelm@26870
    48
  \begin{matharray}{rcl}
wenzelm@26870
    49
    @{command_def "fix"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
    50
    @{command_def "assume"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
    51
    @{command_def "presume"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
    52
    @{command_def "def"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
    53
  \end{matharray}
wenzelm@26870
    54
wenzelm@26870
    55
  The logical proof context consists of fixed variables and
wenzelm@26870
    56
  assumptions.  The former closely correspond to Skolem constants, or
wenzelm@26870
    57
  meta-level universal quantification as provided by the Isabelle/Pure
wenzelm@26870
    58
  logical framework.  Introducing some \emph{arbitrary, but fixed}
wenzelm@26870
    59
  variable via ``@{command "fix"}~@{text x}'' results in a local value
wenzelm@26870
    60
  that may be used in the subsequent proof as any other variable or
wenzelm@26870
    61
  constant.  Furthermore, any result @{text "\<turnstile> \<phi>[x]"} exported from
wenzelm@26870
    62
  the context will be universally closed wrt.\ @{text x} at the
wenzelm@26870
    63
  outermost level: @{text "\<turnstile> \<And>x. \<phi>[x]"} (this is expressed in normal
wenzelm@26870
    64
  form using Isabelle's meta-variables).
wenzelm@26870
    65
wenzelm@26870
    66
  Similarly, introducing some assumption @{text \<chi>} has two effects.
wenzelm@26870
    67
  On the one hand, a local theorem is created that may be used as a
wenzelm@26870
    68
  fact in subsequent proof steps.  On the other hand, any result
wenzelm@26870
    69
  @{text "\<chi> \<turnstile> \<phi>"} exported from the context becomes conditional wrt.\
wenzelm@26870
    70
  the assumption: @{text "\<turnstile> \<chi> \<Longrightarrow> \<phi>"}.  Thus, solving an enclosing goal
wenzelm@26870
    71
  using such a result would basically introduce a new subgoal stemming
wenzelm@26870
    72
  from the assumption.  How this situation is handled depends on the
wenzelm@26870
    73
  version of assumption command used: while @{command "assume"}
wenzelm@26870
    74
  insists on solving the subgoal by unification with some premise of
wenzelm@26870
    75
  the goal, @{command "presume"} leaves the subgoal unchanged in order
wenzelm@26870
    76
  to be proved later by the user.
wenzelm@26870
    77
wenzelm@26870
    78
  Local definitions, introduced by ``@{command "def"}~@{text "x \<equiv>
wenzelm@26870
    79
  t"}'', are achieved by combining ``@{command "fix"}~@{text x}'' with
wenzelm@26870
    80
  another version of assumption that causes any hypothetical equation
wenzelm@26870
    81
  @{text "x \<equiv> t"} to be eliminated by the reflexivity rule.  Thus,
wenzelm@26870
    82
  exporting some result @{text "x \<equiv> t \<turnstile> \<phi>[x]"} yields @{text "\<turnstile>
wenzelm@26870
    83
  \<phi>[t]"}.
wenzelm@26870
    84
wenzelm@26870
    85
  \begin{rail}
wenzelm@26870
    86
    'fix' (vars + 'and')
wenzelm@26870
    87
    ;
wenzelm@26870
    88
    ('assume' | 'presume') (props + 'and')
wenzelm@26870
    89
    ;
wenzelm@26870
    90
    'def' (def + 'and')
wenzelm@26870
    91
    ;
wenzelm@26870
    92
    def: thmdecl? \\ name ('==' | equiv) term termpat?
wenzelm@26870
    93
    ;
wenzelm@26870
    94
  \end{rail}
wenzelm@26870
    95
wenzelm@26870
    96
  \begin{descr}
wenzelm@26870
    97
  
wenzelm@26870
    98
  \item [@{command "fix"}~@{text x}] introduces a local variable
wenzelm@26870
    99
  @{text x} that is \emph{arbitrary, but fixed.}
wenzelm@26870
   100
  
wenzelm@26870
   101
  \item [@{command "assume"}~@{text "a: \<phi>"} and @{command
wenzelm@26870
   102
  "presume"}~@{text "a: \<phi>"}] introduce a local fact @{text "\<phi> \<turnstile> \<phi>"} by
wenzelm@26870
   103
  assumption.  Subsequent results applied to an enclosing goal (e.g.\
wenzelm@26870
   104
  by @{command_ref "show"}) are handled as follows: @{command
wenzelm@26870
   105
  "assume"} expects to be able to unify with existing premises in the
wenzelm@26870
   106
  goal, while @{command "presume"} leaves @{text \<phi>} as new subgoals.
wenzelm@26870
   107
  
wenzelm@26870
   108
  Several lists of assumptions may be given (separated by
wenzelm@26870
   109
  @{keyword_ref "and"}; the resulting list of current facts consists
wenzelm@26870
   110
  of all of these concatenated.
wenzelm@26870
   111
  
wenzelm@26870
   112
  \item [@{command "def"}~@{text "x \<equiv> t"}] introduces a local
wenzelm@26870
   113
  (non-polymorphic) definition.  In results exported from the context,
wenzelm@26870
   114
  @{text x} is replaced by @{text t}.  Basically, ``@{command
wenzelm@26870
   115
  "def"}~@{text "x \<equiv> t"}'' abbreviates ``@{command "fix"}~@{text
wenzelm@26870
   116
  x}~@{command "assume"}~@{text "x \<equiv> t"}'', with the resulting
wenzelm@26870
   117
  hypothetical equation solved by reflexivity.
wenzelm@26870
   118
  
wenzelm@26870
   119
  The default name for the definitional equation is @{text x_def}.
wenzelm@26870
   120
  Several simultaneous definitions may be given at the same time.
wenzelm@26870
   121
wenzelm@26870
   122
  \end{descr}
wenzelm@26870
   123
wenzelm@26870
   124
  The special name @{fact_ref prems} refers to all assumptions of the
wenzelm@26870
   125
  current context as a list of theorems.  This feature should be used
wenzelm@26870
   126
  with great care!  It is better avoided in final proof texts.
wenzelm@26870
   127
*}
wenzelm@26870
   128
wenzelm@26870
   129
wenzelm@26870
   130
section {* Facts and forward chaining *}
wenzelm@26870
   131
wenzelm@26870
   132
text {*
wenzelm@26870
   133
  \begin{matharray}{rcl}
wenzelm@26870
   134
    @{command_def "note"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   135
    @{command_def "then"} & : & \isartrans{proof(state)}{proof(chain)} \\
wenzelm@26870
   136
    @{command_def "from"} & : & \isartrans{proof(state)}{proof(chain)} \\
wenzelm@26870
   137
    @{command_def "with"} & : & \isartrans{proof(state)}{proof(chain)} \\
wenzelm@26870
   138
    @{command_def "using"} & : & \isartrans{proof(prove)}{proof(prove)} \\
wenzelm@26870
   139
    @{command_def "unfolding"} & : & \isartrans{proof(prove)}{proof(prove)} \\
wenzelm@26870
   140
  \end{matharray}
wenzelm@26870
   141
wenzelm@26870
   142
  New facts are established either by assumption or proof of local
wenzelm@26870
   143
  statements.  Any fact will usually be involved in further proofs,
wenzelm@26870
   144
  either as explicit arguments of proof methods, or when forward
wenzelm@26870
   145
  chaining towards the next goal via @{command "then"} (and variants);
wenzelm@26870
   146
  @{command "from"} and @{command "with"} are composite forms
wenzelm@26870
   147
  involving @{command "note"}.  The @{command "using"} elements
wenzelm@26870
   148
  augments the collection of used facts \emph{after} a goal has been
wenzelm@26870
   149
  stated.  Note that the special theorem name @{fact_ref this} refers
wenzelm@26870
   150
  to the most recently established facts, but only \emph{before}
wenzelm@26870
   151
  issuing a follow-up claim.
wenzelm@26870
   152
wenzelm@26870
   153
  \begin{rail}
wenzelm@26870
   154
    'note' (thmdef? thmrefs + 'and')
wenzelm@26870
   155
    ;
wenzelm@26870
   156
    ('from' | 'with' | 'using' | 'unfolding') (thmrefs + 'and')
wenzelm@26870
   157
    ;
wenzelm@26870
   158
  \end{rail}
wenzelm@26870
   159
wenzelm@26870
   160
  \begin{descr}
wenzelm@26870
   161
wenzelm@26870
   162
  \item [@{command "note"}~@{text "a = b\<^sub>1 \<dots> b\<^sub>n"}]
wenzelm@26870
   163
  recalls existing facts @{text "b\<^sub>1, \<dots>, b\<^sub>n"}, binding
wenzelm@26870
   164
  the result as @{text a}.  Note that attributes may be involved as
wenzelm@26870
   165
  well, both on the left and right hand sides.
wenzelm@26870
   166
wenzelm@26870
   167
  \item [@{command "then"}] indicates forward chaining by the current
wenzelm@26870
   168
  facts in order to establish the goal to be claimed next.  The
wenzelm@26870
   169
  initial proof method invoked to refine that will be offered the
wenzelm@26870
   170
  facts to do ``anything appropriate'' (see also
wenzelm@26870
   171
  \secref{sec:proof-steps}).  For example, method @{method_ref rule}
wenzelm@26870
   172
  (see \secref{sec:pure-meth-att}) would typically do an elimination
wenzelm@26870
   173
  rather than an introduction.  Automatic methods usually insert the
wenzelm@26870
   174
  facts into the goal state before operation.  This provides a simple
wenzelm@26870
   175
  scheme to control relevance of facts in automated proof search.
wenzelm@26870
   176
  
wenzelm@26870
   177
  \item [@{command "from"}~@{text b}] abbreviates ``@{command
wenzelm@26870
   178
  "note"}~@{text b}~@{command "then"}''; thus @{command "then"} is
wenzelm@26870
   179
  equivalent to ``@{command "from"}~@{text this}''.
wenzelm@26870
   180
  
wenzelm@26870
   181
  \item [@{command "with"}~@{text "b\<^sub>1 \<dots> b\<^sub>n"}]
wenzelm@26870
   182
  abbreviates ``@{command "from"}~@{text "b\<^sub>1 \<dots> b\<^sub>n \<AND>
wenzelm@26870
   183
  this"}''; thus the forward chaining is from earlier facts together
wenzelm@26870
   184
  with the current ones.
wenzelm@26870
   185
  
wenzelm@26870
   186
  \item [@{command "using"}~@{text "b\<^sub>1 \<dots> b\<^sub>n"}] augments
wenzelm@26870
   187
  the facts being currently indicated for use by a subsequent
wenzelm@26870
   188
  refinement step (such as @{command_ref "apply"} or @{command_ref
wenzelm@26870
   189
  "proof"}).
wenzelm@26870
   190
  
wenzelm@26870
   191
  \item [@{command "unfolding"}~@{text "b\<^sub>1 \<dots> b\<^sub>n"}] is
wenzelm@26870
   192
  structurally similar to @{command "using"}, but unfolds definitional
wenzelm@26870
   193
  equations @{text "b\<^sub>1, \<dots> b\<^sub>n"} throughout the goal state
wenzelm@26870
   194
  and facts.
wenzelm@26870
   195
wenzelm@26870
   196
  \end{descr}
wenzelm@26870
   197
wenzelm@26870
   198
  Forward chaining with an empty list of theorems is the same as not
wenzelm@26870
   199
  chaining at all.  Thus ``@{command "from"}~@{text nothing}'' has no
wenzelm@26870
   200
  effect apart from entering @{text "prove(chain)"} mode, since
wenzelm@26870
   201
  @{fact_ref nothing} is bound to the empty list of theorems.
wenzelm@26870
   202
wenzelm@26870
   203
  Basic proof methods (such as @{method_ref rule}) expect multiple
wenzelm@26870
   204
  facts to be given in their proper order, corresponding to a prefix
wenzelm@26870
   205
  of the premises of the rule involved.  Note that positions may be
wenzelm@26870
   206
  easily skipped using something like @{command "from"}~@{text "_
wenzelm@26870
   207
  \<AND> a \<AND> b"}, for example.  This involves the trivial rule
wenzelm@26870
   208
  @{text "PROP \<psi> \<Longrightarrow> PROP \<psi>"}, which is bound in Isabelle/Pure as
wenzelm@26870
   209
  ``@{fact_ref "_"}'' (underscore).
wenzelm@26870
   210
wenzelm@26870
   211
  Automated methods (such as @{method simp} or @{method auto}) just
wenzelm@26870
   212
  insert any given facts before their usual operation.  Depending on
wenzelm@26870
   213
  the kind of procedure involved, the order of facts is less
wenzelm@26870
   214
  significant here.
wenzelm@26870
   215
*}
wenzelm@26870
   216
wenzelm@26870
   217
wenzelm@26870
   218
section {* Goal statements \label{sec:goals} *}
wenzelm@26870
   219
wenzelm@26870
   220
text {*
wenzelm@26870
   221
  \begin{matharray}{rcl}
wenzelm@26870
   222
    @{command_def "lemma"} & : & \isartrans{local{\dsh}theory}{proof(prove)} \\
wenzelm@26870
   223
    @{command_def "theorem"} & : & \isartrans{local{\dsh}theory}{proof(prove)} \\
wenzelm@26870
   224
    @{command_def "corollary"} & : & \isartrans{local{\dsh}theory}{proof(prove)} \\
wenzelm@26870
   225
    @{command_def "have"} & : & \isartrans{proof(state) ~|~ proof(chain)}{proof(prove)} \\
wenzelm@26870
   226
    @{command_def "show"} & : & \isartrans{proof(state) ~|~ proof(chain)}{proof(prove)} \\
wenzelm@26870
   227
    @{command_def "hence"} & : & \isartrans{proof(state)}{proof(prove)} \\
wenzelm@26870
   228
    @{command_def "thus"} & : & \isartrans{proof(state)}{proof(prove)} \\
wenzelm@26870
   229
    @{command_def "print_statement"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
wenzelm@26870
   230
  \end{matharray}
wenzelm@26870
   231
wenzelm@26870
   232
  From a theory context, proof mode is entered by an initial goal
wenzelm@26870
   233
  command such as @{command "lemma"}, @{command "theorem"}, or
wenzelm@26870
   234
  @{command "corollary"}.  Within a proof, new claims may be
wenzelm@26870
   235
  introduced locally as well; four variants are available here to
wenzelm@26870
   236
  indicate whether forward chaining of facts should be performed
wenzelm@26870
   237
  initially (via @{command_ref "then"}), and whether the final result
wenzelm@26870
   238
  is meant to solve some pending goal.
wenzelm@26870
   239
wenzelm@26870
   240
  Goals may consist of multiple statements, resulting in a list of
wenzelm@26870
   241
  facts eventually.  A pending multi-goal is internally represented as
wenzelm@26870
   242
  a meta-level conjunction (printed as @{text "&&"}), which is usually
wenzelm@26870
   243
  split into the corresponding number of sub-goals prior to an initial
wenzelm@26870
   244
  method application, via @{command_ref "proof"}
wenzelm@26870
   245
  (\secref{sec:proof-steps}) or @{command_ref "apply"}
wenzelm@26870
   246
  (\secref{sec:tactic-commands}).  The @{method_ref induct} method
wenzelm@26870
   247
  covered in \secref{sec:cases-induct} acts on multiple claims
wenzelm@26870
   248
  simultaneously.
wenzelm@26870
   249
wenzelm@26870
   250
  Claims at the theory level may be either in short or long form.  A
wenzelm@26870
   251
  short goal merely consists of several simultaneous propositions
wenzelm@26870
   252
  (often just one).  A long goal includes an explicit context
wenzelm@26870
   253
  specification for the subsequent conclusion, involving local
wenzelm@26870
   254
  parameters and assumptions.  Here the role of each part of the
wenzelm@26870
   255
  statement is explicitly marked by separate keywords (see also
wenzelm@26870
   256
  \secref{sec:locale}); the local assumptions being introduced here
wenzelm@26870
   257
  are available as @{fact_ref assms} in the proof.  Moreover, there
wenzelm@26870
   258
  are two kinds of conclusions: @{element_def "shows"} states several
wenzelm@26870
   259
  simultaneous propositions (essentially a big conjunction), while
wenzelm@26870
   260
  @{element_def "obtains"} claims several simultaneous simultaneous
wenzelm@26870
   261
  contexts of (essentially a big disjunction of eliminated parameters
wenzelm@26870
   262
  and assumptions, cf.\ \secref{sec:obtain}).
wenzelm@26870
   263
wenzelm@26870
   264
  \begin{rail}
wenzelm@26870
   265
    ('lemma' | 'theorem' | 'corollary') target? (goal | longgoal)
wenzelm@26870
   266
    ;
wenzelm@26870
   267
    ('have' | 'show' | 'hence' | 'thus') goal
wenzelm@26870
   268
    ;
wenzelm@26870
   269
    'print\_statement' modes? thmrefs
wenzelm@26870
   270
    ;
wenzelm@26870
   271
  
wenzelm@26870
   272
    goal: (props + 'and')
wenzelm@26870
   273
    ;
wenzelm@26870
   274
    longgoal: thmdecl? (contextelem *) conclusion
wenzelm@26870
   275
    ;
wenzelm@26870
   276
    conclusion: 'shows' goal | 'obtains' (parname? case + '|')
wenzelm@26870
   277
    ;
wenzelm@26870
   278
    case: (vars + 'and') 'where' (props + 'and')
wenzelm@26870
   279
    ;
wenzelm@26870
   280
  \end{rail}
wenzelm@26870
   281
wenzelm@26870
   282
  \begin{descr}
wenzelm@26870
   283
  
wenzelm@26870
   284
  \item [@{command "lemma"}~@{text "a: \<phi>"}] enters proof mode with
wenzelm@26870
   285
  @{text \<phi>} as main goal, eventually resulting in some fact @{text "\<turnstile>
wenzelm@26870
   286
  \<phi>"} to be put back into the target context.  An additional
wenzelm@26870
   287
  \railnonterm{context} specification may build up an initial proof
wenzelm@26870
   288
  context for the subsequent claim; this includes local definitions
wenzelm@26870
   289
  and syntax as well, see the definition of @{syntax contextelem} in
wenzelm@26870
   290
  \secref{sec:locale}.
wenzelm@26870
   291
  
wenzelm@26870
   292
  \item [@{command "theorem"}~@{text "a: \<phi>"} and @{command
wenzelm@26870
   293
  "corollary"}~@{text "a: \<phi>"}] are essentially the same as @{command
wenzelm@26870
   294
  "lemma"}~@{text "a: \<phi>"}, but the facts are internally marked as
wenzelm@26870
   295
  being of a different kind.  This discrimination acts like a formal
wenzelm@26870
   296
  comment.
wenzelm@26870
   297
  
wenzelm@26870
   298
  \item [@{command "have"}~@{text "a: \<phi>"}] claims a local goal,
wenzelm@26870
   299
  eventually resulting in a fact within the current logical context.
wenzelm@26870
   300
  This operation is completely independent of any pending sub-goals of
wenzelm@26870
   301
  an enclosing goal statements, so @{command "have"} may be freely
wenzelm@26870
   302
  used for experimental exploration of potential results within a
wenzelm@26870
   303
  proof body.
wenzelm@26870
   304
  
wenzelm@26870
   305
  \item [@{command "show"}~@{text "a: \<phi>"}] is like @{command
wenzelm@26870
   306
  "have"}~@{text "a: \<phi>"} plus a second stage to refine some pending
wenzelm@26870
   307
  sub-goal for each one of the finished result, after having been
wenzelm@26870
   308
  exported into the corresponding context (at the head of the
wenzelm@26870
   309
  sub-proof of this @{command "show"} command).
wenzelm@26870
   310
  
wenzelm@26870
   311
  To accommodate interactive debugging, resulting rules are printed
wenzelm@26870
   312
  before being applied internally.  Even more, interactive execution
wenzelm@26870
   313
  of @{command "show"} predicts potential failure and displays the
wenzelm@26870
   314
  resulting error as a warning beforehand.  Watch out for the
wenzelm@26870
   315
  following message:
wenzelm@26870
   316
wenzelm@26870
   317
  %FIXME proper antiquitation
wenzelm@26870
   318
  \begin{ttbox}
wenzelm@26870
   319
  Problem! Local statement will fail to solve any pending goal
wenzelm@26870
   320
  \end{ttbox}
wenzelm@26870
   321
  
wenzelm@26870
   322
  \item [@{command "hence"}] abbreviates ``@{command "then"}~@{command
wenzelm@26870
   323
  "have"}'', i.e.\ claims a local goal to be proven by forward
wenzelm@26870
   324
  chaining the current facts.  Note that @{command "hence"} is also
wenzelm@26870
   325
  equivalent to ``@{command "from"}~@{text this}~@{command "have"}''.
wenzelm@26870
   326
  
wenzelm@26870
   327
  \item [@{command "thus"}] abbreviates ``@{command "then"}~@{command
wenzelm@26870
   328
  "show"}''.  Note that @{command "thus"} is also equivalent to
wenzelm@26870
   329
  ``@{command "from"}~@{text this}~@{command "show"}''.
wenzelm@26870
   330
  
wenzelm@26870
   331
  \item [@{command "print_statement"}~@{text a}] prints facts from the
wenzelm@26870
   332
  current theory or proof context in long statement form, according to
wenzelm@26870
   333
  the syntax for @{command "lemma"} given above.
wenzelm@26870
   334
wenzelm@26870
   335
  \end{descr}
wenzelm@26870
   336
wenzelm@26870
   337
  Any goal statement causes some term abbreviations (such as
wenzelm@26870
   338
  @{variable_ref "?thesis"}) to be bound automatically, see also
wenzelm@26870
   339
  \secref{sec:term-abbrev}.  Furthermore, the local context of a
wenzelm@26870
   340
  (non-atomic) goal is provided via the @{case_ref rule_context} case.
wenzelm@26870
   341
wenzelm@26870
   342
  The optional case names of @{element_ref "obtains"} have a twofold
wenzelm@26870
   343
  meaning: (1) during the of this claim they refer to the the local
wenzelm@26870
   344
  context introductions, (2) the resulting rule is annotated
wenzelm@26870
   345
  accordingly to support symbolic case splits when used with the
wenzelm@26870
   346
  @{method_ref cases} method (cf.  \secref{sec:cases-induct}).
wenzelm@26870
   347
wenzelm@26870
   348
  \medskip
wenzelm@26870
   349
wenzelm@26870
   350
  \begin{warn}
wenzelm@26870
   351
    Isabelle/Isar suffers theory-level goal statements to contain
wenzelm@26870
   352
    \emph{unbound schematic variables}, although this does not conform
wenzelm@26870
   353
    to the aim of human-readable proof documents!  The main problem
wenzelm@26870
   354
    with schematic goals is that the actual outcome is usually hard to
wenzelm@26870
   355
    predict, depending on the behavior of the proof methods applied
wenzelm@26870
   356
    during the course of reasoning.  Note that most semi-automated
wenzelm@26870
   357
    methods heavily depend on several kinds of implicit rule
wenzelm@26870
   358
    declarations within the current theory context.  As this would
wenzelm@26870
   359
    also result in non-compositional checking of sub-proofs,
wenzelm@26870
   360
    \emph{local goals} are not allowed to be schematic at all.
wenzelm@26870
   361
    Nevertheless, schematic goals do have their use in Prolog-style
wenzelm@26870
   362
    interactive synthesis of proven results, usually by stepwise
wenzelm@26870
   363
    refinement via emulation of traditional Isabelle tactic scripts
wenzelm@26870
   364
    (see also \secref{sec:tactic-commands}).  In any case, users
wenzelm@26870
   365
    should know what they are doing.
wenzelm@26870
   366
  \end{warn}
wenzelm@26870
   367
*}
wenzelm@26870
   368
wenzelm@26870
   369
wenzelm@26870
   370
section {* Initial and terminal proof steps \label{sec:proof-steps} *}
wenzelm@26870
   371
wenzelm@26870
   372
text {*
wenzelm@26870
   373
  \begin{matharray}{rcl}
wenzelm@26870
   374
    @{command_def "proof"} & : & \isartrans{proof(prove)}{proof(state)} \\
wenzelm@26870
   375
    @{command_def "qed"} & : & \isartrans{proof(state)}{proof(state) ~|~ theory} \\
wenzelm@26870
   376
    @{command_def "by"} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
wenzelm@26870
   377
    @{command_def ".."} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
wenzelm@26870
   378
    @{command_def "."} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
wenzelm@26870
   379
    @{command_def "sorry"} & : & \isartrans{proof(prove)}{proof(state) ~|~ theory} \\
wenzelm@26870
   380
  \end{matharray}
wenzelm@26870
   381
wenzelm@26870
   382
  Arbitrary goal refinement via tactics is considered harmful.
wenzelm@26870
   383
  Structured proof composition in Isar admits proof methods to be
wenzelm@26870
   384
  invoked in two places only.
wenzelm@26870
   385
wenzelm@26870
   386
  \begin{enumerate}
wenzelm@26870
   387
wenzelm@26870
   388
  \item An \emph{initial} refinement step @{command_ref
wenzelm@26870
   389
  "proof"}~@{text "m\<^sub>1"} reduces a newly stated goal to a number
wenzelm@26870
   390
  of sub-goals that are to be solved later.  Facts are passed to
wenzelm@26870
   391
  @{text "m\<^sub>1"} for forward chaining, if so indicated by @{text
wenzelm@26870
   392
  "proof(chain)"} mode.
wenzelm@26870
   393
  
wenzelm@26870
   394
  \item A \emph{terminal} conclusion step @{command_ref "qed"}~@{text
wenzelm@26870
   395
  "m\<^sub>2"} is intended to solve remaining goals.  No facts are
wenzelm@26870
   396
  passed to @{text "m\<^sub>2"}.
wenzelm@26870
   397
wenzelm@26870
   398
  \end{enumerate}
wenzelm@26870
   399
wenzelm@26870
   400
  The only other (proper) way to affect pending goals in a proof body
wenzelm@26870
   401
  is by @{command_ref "show"}, which involves an explicit statement of
wenzelm@26870
   402
  what is to be solved eventually.  Thus we avoid the fundamental
wenzelm@26870
   403
  problem of unstructured tactic scripts that consist of numerous
wenzelm@26870
   404
  consecutive goal transformations, with invisible effects.
wenzelm@26870
   405
wenzelm@26870
   406
  \medskip As a general rule of thumb for good proof style, initial
wenzelm@26870
   407
  proof methods should either solve the goal completely, or constitute
wenzelm@26870
   408
  some well-understood reduction to new sub-goals.  Arbitrary
wenzelm@26870
   409
  automatic proof tools that are prone leave a large number of badly
wenzelm@26870
   410
  structured sub-goals are no help in continuing the proof document in
wenzelm@26870
   411
  an intelligible manner.
wenzelm@26870
   412
wenzelm@26870
   413
  Unless given explicitly by the user, the default initial method is
wenzelm@26870
   414
  ``@{method_ref rule}'', which applies a single standard elimination
wenzelm@26870
   415
  or introduction rule according to the topmost symbol involved.
wenzelm@26870
   416
  There is no separate default terminal method.  Any remaining goals
wenzelm@26870
   417
  are always solved by assumption in the very last step.
wenzelm@26870
   418
wenzelm@26870
   419
  \begin{rail}
wenzelm@26870
   420
    'proof' method?
wenzelm@26870
   421
    ;
wenzelm@26870
   422
    'qed' method?
wenzelm@26870
   423
    ;
wenzelm@26870
   424
    'by' method method?
wenzelm@26870
   425
    ;
wenzelm@26870
   426
    ('.' | '..' | 'sorry')
wenzelm@26870
   427
    ;
wenzelm@26870
   428
  \end{rail}
wenzelm@26870
   429
wenzelm@26870
   430
  \begin{descr}
wenzelm@26870
   431
  
wenzelm@26870
   432
  \item [@{command "proof"}~@{text "m\<^sub>1"}] refines the goal by
wenzelm@26870
   433
  proof method @{text "m\<^sub>1"}; facts for forward chaining are
wenzelm@26870
   434
  passed if so indicated by @{text "proof(chain)"} mode.
wenzelm@26870
   435
  
wenzelm@26870
   436
  \item [@{command "qed"}~@{text "m\<^sub>2"}] refines any remaining
wenzelm@26870
   437
  goals by proof method @{text "m\<^sub>2"} and concludes the
wenzelm@26870
   438
  sub-proof by assumption.  If the goal had been @{text "show"} (or
wenzelm@26870
   439
  @{text "thus"}), some pending sub-goal is solved as well by the rule
wenzelm@26870
   440
  resulting from the result \emph{exported} into the enclosing goal
wenzelm@26870
   441
  context.  Thus @{text "qed"} may fail for two reasons: either @{text
wenzelm@26870
   442
  "m\<^sub>2"} fails, or the resulting rule does not fit to any
wenzelm@26870
   443
  pending goal\footnote{This includes any additional ``strong''
wenzelm@26870
   444
  assumptions as introduced by @{command "assume"}.} of the enclosing
wenzelm@26870
   445
  context.  Debugging such a situation might involve temporarily
wenzelm@26870
   446
  changing @{command "show"} into @{command "have"}, or weakening the
wenzelm@26870
   447
  local context by replacing occurrences of @{command "assume"} by
wenzelm@26870
   448
  @{command "presume"}.
wenzelm@26870
   449
  
wenzelm@26870
   450
  \item [@{command "by"}~@{text "m\<^sub>1 m\<^sub>2"}] is a
wenzelm@26870
   451
  \emph{terminal proof}\index{proof!terminal}; it abbreviates
wenzelm@26870
   452
  @{command "proof"}~@{text "m\<^sub>1"}~@{text "qed"}~@{text
wenzelm@26870
   453
  "m\<^sub>2"}, but with backtracking across both methods.  Debugging
wenzelm@26870
   454
  an unsuccessful @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"}
wenzelm@26870
   455
  command can be done by expanding its definition; in many cases
wenzelm@26870
   456
  @{command "proof"}~@{text "m\<^sub>1"} (or even @{text
wenzelm@26870
   457
  "apply"}~@{text "m\<^sub>1"}) is already sufficient to see the
wenzelm@26870
   458
  problem.
wenzelm@26870
   459
wenzelm@26870
   460
  \item [``@{command ".."}''] is a \emph{default
wenzelm@26870
   461
  proof}\index{proof!default}; it abbreviates @{command "by"}~@{text
wenzelm@26870
   462
  "rule"}.
wenzelm@26870
   463
wenzelm@26870
   464
  \item [``@{command "."}''] is a \emph{trivial
wenzelm@26870
   465
  proof}\index{proof!trivial}; it abbreviates @{command "by"}~@{text
wenzelm@26870
   466
  "this"}.
wenzelm@26870
   467
  
wenzelm@26870
   468
  \item [@{command "sorry"}] is a \emph{fake proof}\index{proof!fake}
wenzelm@26870
   469
  pretending to solve the pending claim without further ado.  This
wenzelm@26870
   470
  only works in interactive development, or if the @{ML
wenzelm@26870
   471
  quick_and_dirty} flag is enabled (in ML).  Facts emerging from fake
wenzelm@26870
   472
  proofs are not the real thing.  Internally, each theorem container
wenzelm@26870
   473
  is tainted by an oracle invocation, which is indicated as ``@{text
wenzelm@26870
   474
  "[!]"}'' in the printed result.
wenzelm@26870
   475
  
wenzelm@26870
   476
  The most important application of @{command "sorry"} is to support
wenzelm@26870
   477
  experimentation and top-down proof development.
wenzelm@26870
   478
wenzelm@26870
   479
  \end{descr}
wenzelm@26870
   480
*}
wenzelm@26870
   481
wenzelm@26870
   482
wenzelm@26870
   483
section {* Fundamental methods and attributes \label{sec:pure-meth-att} *}
wenzelm@26870
   484
wenzelm@26870
   485
text {*
wenzelm@26870
   486
  The following proof methods and attributes refer to basic logical
wenzelm@26870
   487
  operations of Isar.  Further methods and attributes are provided by
wenzelm@26870
   488
  several generic and object-logic specific tools and packages (see
wenzelm@26870
   489
  \chref{ch:gen-tools} and \chref{ch:hol}).
wenzelm@26870
   490
wenzelm@26870
   491
  \begin{matharray}{rcl}
wenzelm@26870
   492
    @{method_def "-"} & : & \isarmeth \\
wenzelm@26870
   493
    @{method_def "fact"} & : & \isarmeth \\
wenzelm@26870
   494
    @{method_def "assumption"} & : & \isarmeth \\
wenzelm@26870
   495
    @{method_def "this"} & : & \isarmeth \\
wenzelm@26870
   496
    @{method_def "rule"} & : & \isarmeth \\
wenzelm@26870
   497
    @{method_def "iprover"} & : & \isarmeth \\[0.5ex]
wenzelm@26870
   498
    @{attribute_def "intro"} & : & \isaratt \\
wenzelm@26870
   499
    @{attribute_def "elim"} & : & \isaratt \\
wenzelm@26870
   500
    @{attribute_def "dest"} & : & \isaratt \\
wenzelm@26870
   501
    @{attribute_def "rule"} & : & \isaratt \\[0.5ex]
wenzelm@26870
   502
    @{attribute_def "OF"} & : & \isaratt \\
wenzelm@26870
   503
    @{attribute_def "of"} & : & \isaratt \\
wenzelm@26870
   504
    @{attribute_def "where"} & : & \isaratt \\
wenzelm@26870
   505
  \end{matharray}
wenzelm@26870
   506
wenzelm@26870
   507
  \begin{rail}
wenzelm@26870
   508
    'fact' thmrefs?
wenzelm@26870
   509
    ;
wenzelm@26870
   510
    'rule' thmrefs?
wenzelm@26870
   511
    ;
wenzelm@26870
   512
    'iprover' ('!' ?) (rulemod *)
wenzelm@26870
   513
    ;
wenzelm@26870
   514
    rulemod: ('intro' | 'elim' | 'dest') ((('!' | () | '?') nat?) | 'del') ':' thmrefs
wenzelm@26870
   515
    ;
wenzelm@26870
   516
    ('intro' | 'elim' | 'dest') ('!' | () | '?') nat?
wenzelm@26870
   517
    ;
wenzelm@26870
   518
    'rule' 'del'
wenzelm@26870
   519
    ;
wenzelm@26870
   520
    'OF' thmrefs
wenzelm@26870
   521
    ;
wenzelm@26870
   522
    'of' insts ('concl' ':' insts)?
wenzelm@26870
   523
    ;
wenzelm@26870
   524
    'where' ((name | var | typefree | typevar) '=' (type | term) * 'and')
wenzelm@26870
   525
    ;
wenzelm@26870
   526
  \end{rail}
wenzelm@26870
   527
wenzelm@26870
   528
  \begin{descr}
wenzelm@26870
   529
  
wenzelm@26870
   530
  \item [``@{method "-"}'' (minus)] does nothing but insert the
wenzelm@26870
   531
  forward chaining facts as premises into the goal.  Note that command
wenzelm@26870
   532
  @{command_ref "proof"} without any method actually performs a single
wenzelm@26870
   533
  reduction step using the @{method_ref rule} method; thus a plain
wenzelm@26870
   534
  \emph{do-nothing} proof step would be ``@{command "proof"}~@{text
wenzelm@26870
   535
  "-"}'' rather than @{command "proof"} alone.
wenzelm@26870
   536
  
wenzelm@26870
   537
  \item [@{method "fact"}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}] composes
wenzelm@26870
   538
  some fact from @{text "a\<^sub>1, \<dots>, a\<^sub>n"} (or implicitly from
wenzelm@26870
   539
  the current proof context) modulo unification of schematic type and
wenzelm@26870
   540
  term variables.  The rule structure is not taken into account, i.e.\
wenzelm@26870
   541
  meta-level implication is considered atomic.  This is the same
wenzelm@26870
   542
  principle underlying literal facts (cf.\ \secref{sec:syn-att}):
wenzelm@26870
   543
  ``@{command "have"}~@{text "\<phi>"}~@{command "by"}~@{text fact}'' is
wenzelm@26870
   544
  equivalent to ``@{command "note"}~@{verbatim "`"}@{text \<phi>}@{verbatim
wenzelm@26870
   545
  "`"}'' provided that @{text "\<turnstile> \<phi>"} is an instance of some known
wenzelm@26870
   546
  @{text "\<turnstile> \<phi>"} in the proof context.
wenzelm@26870
   547
  
wenzelm@26870
   548
  \item [@{method assumption}] solves some goal by a single assumption
wenzelm@26870
   549
  step.  All given facts are guaranteed to participate in the
wenzelm@26870
   550
  refinement; this means there may be only 0 or 1 in the first place.
wenzelm@26870
   551
  Recall that @{command "qed"} (\secref{sec:proof-steps}) already
wenzelm@26870
   552
  concludes any remaining sub-goals by assumption, so structured
wenzelm@26870
   553
  proofs usually need not quote the @{method assumption} method at
wenzelm@26870
   554
  all.
wenzelm@26870
   555
  
wenzelm@26870
   556
  \item [@{method this}] applies all of the current facts directly as
wenzelm@26870
   557
  rules.  Recall that ``@{command "."}'' (dot) abbreviates ``@{command
wenzelm@26870
   558
  "by"}~@{text this}''.
wenzelm@26870
   559
  
wenzelm@26870
   560
  \item [@{method rule}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}] applies some
wenzelm@26870
   561
  rule given as argument in backward manner; facts are used to reduce
wenzelm@26870
   562
  the rule before applying it to the goal.  Thus @{method rule}
wenzelm@26870
   563
  without facts is plain introduction, while with facts it becomes
wenzelm@26870
   564
  elimination.
wenzelm@26870
   565
  
wenzelm@26870
   566
  When no arguments are given, the @{method rule} method tries to pick
wenzelm@26870
   567
  appropriate rules automatically, as declared in the current context
wenzelm@26870
   568
  using the @{attribute intro}, @{attribute elim}, @{attribute dest}
wenzelm@26870
   569
  attributes (see below).  This is the default behavior of @{command
wenzelm@26870
   570
  "proof"} and ``@{command ".."}'' (double-dot) steps (see
wenzelm@26870
   571
  \secref{sec:proof-steps}).
wenzelm@26870
   572
  
wenzelm@26870
   573
  \item [@{method iprover}] performs intuitionistic proof search,
wenzelm@26870
   574
  depending on specifically declared rules from the context, or given
wenzelm@26870
   575
  as explicit arguments.  Chained facts are inserted into the goal
wenzelm@26870
   576
  before commencing proof search; ``@{method iprover}@{text "!"}'' 
wenzelm@26870
   577
  means to include the current @{fact prems} as well.
wenzelm@26870
   578
  
wenzelm@26870
   579
  Rules need to be classified as @{attribute intro}, @{attribute
wenzelm@26870
   580
  elim}, or @{attribute dest}; here the ``@{text "!"}'' indicator
wenzelm@26870
   581
  refers to ``safe'' rules, which may be applied aggressively (without
wenzelm@26870
   582
  considering back-tracking later).  Rules declared with ``@{text
wenzelm@26870
   583
  "?"}'' are ignored in proof search (the single-step @{method rule}
wenzelm@26870
   584
  method still observes these).  An explicit weight annotation may be
wenzelm@26870
   585
  given as well; otherwise the number of rule premises will be taken
wenzelm@26870
   586
  into account here.
wenzelm@26870
   587
  
wenzelm@26870
   588
  \item [@{attribute intro}, @{attribute elim}, and @{attribute dest}]
wenzelm@26870
   589
  declare introduction, elimination, and destruct rules, to be used
wenzelm@26870
   590
  with the @{method rule} and @{method iprover} methods.  Note that
wenzelm@26870
   591
  the latter will ignore rules declared with ``@{text "?"}'', while
wenzelm@26870
   592
  ``@{text "!"}''  are used most aggressively.
wenzelm@26870
   593
  
wenzelm@26870
   594
  The classical reasoner (see \secref{sec:classical}) introduces its
wenzelm@26870
   595
  own variants of these attributes; use qualified names to access the
wenzelm@26870
   596
  present versions of Isabelle/Pure, i.e.\ @{attribute "Pure.intro"}.
wenzelm@26870
   597
  
wenzelm@26870
   598
  \item [@{attribute rule}~@{text del}] undeclares introduction,
wenzelm@26870
   599
  elimination, or destruct rules.
wenzelm@26870
   600
  
wenzelm@26870
   601
  \item [@{attribute OF}~@{text "a\<^sub>1 \<dots> a\<^sub>n"}] applies some
wenzelm@26870
   602
  theorem to all of the given rules @{text "a\<^sub>1, \<dots>, a\<^sub>n"}
wenzelm@26870
   603
  (in parallel).  This corresponds to the @{ML "op MRS"} operation in
wenzelm@26870
   604
  ML, but note the reversed order.  Positions may be effectively
wenzelm@26870
   605
  skipped by including ``@{text _}'' (underscore) as argument.
wenzelm@26870
   606
  
wenzelm@26870
   607
  \item [@{attribute of}~@{text "t\<^sub>1 \<dots> t\<^sub>n"}] performs
wenzelm@26870
   608
  positional instantiation of term variables.  The terms @{text
wenzelm@26870
   609
  "t\<^sub>1, \<dots>, t\<^sub>n"} are substituted for any schematic
wenzelm@26888
   610
  variables occurring in a theorem from left to right; ``@{text _}''
wenzelm@26888
   611
  (underscore) indicates to skip a position.  Arguments following a
wenzelm@26888
   612
  ``@{text "concl:"}'' specification refer to positions of the
wenzelm@26888
   613
  conclusion of a rule.
wenzelm@26870
   614
  
wenzelm@26870
   615
  \item [@{attribute "where"}~@{text "x\<^sub>1 = t\<^sub>1 \<AND> \<dots>
wenzelm@26870
   616
  x\<^sub>n = t\<^sub>n"}] performs named instantiation of schematic
wenzelm@26870
   617
  type and term variables occurring in a theorem.  Schematic variables
wenzelm@26870
   618
  have to be specified on the left-hand side (e.g.\ @{text "?x1.3"}).
wenzelm@26870
   619
  The question mark may be omitted if the variable name is a plain
wenzelm@26870
   620
  identifier without index.  As type instantiations are inferred from
wenzelm@26870
   621
  term instantiations, explicit type instantiations are seldom
wenzelm@26870
   622
  necessary.
wenzelm@26870
   623
wenzelm@26870
   624
  \end{descr}
wenzelm@26870
   625
*}
wenzelm@26870
   626
wenzelm@26870
   627
wenzelm@26870
   628
section {* Term abbreviations \label{sec:term-abbrev} *}
wenzelm@26870
   629
wenzelm@26870
   630
text {*
wenzelm@26870
   631
  \begin{matharray}{rcl}
wenzelm@26870
   632
    @{command_def "let"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   633
    @{keyword_def "is"} & : & syntax \\
wenzelm@26870
   634
  \end{matharray}
wenzelm@26870
   635
wenzelm@26870
   636
  Abbreviations may be either bound by explicit @{command
wenzelm@26870
   637
  "let"}~@{text "p \<equiv> t"} statements, or by annotating assumptions or
wenzelm@26870
   638
  goal statements with a list of patterns ``@{text "(\<IS> p\<^sub>1 \<dots>
wenzelm@26870
   639
  p\<^sub>n)"}''.  In both cases, higher-order matching is invoked to
wenzelm@26870
   640
  bind extra-logical term variables, which may be either named
wenzelm@26870
   641
  schematic variables of the form @{text ?x}, or nameless dummies
wenzelm@26870
   642
  ``@{variable _}'' (underscore). Note that in the @{command "let"}
wenzelm@26870
   643
  form the patterns occur on the left-hand side, while the @{keyword
wenzelm@26870
   644
  "is"} patterns are in postfix position.
wenzelm@26870
   645
wenzelm@26870
   646
  Polymorphism of term bindings is handled in Hindley-Milner style,
wenzelm@26870
   647
  similar to ML.  Type variables referring to local assumptions or
wenzelm@26870
   648
  open goal statements are \emph{fixed}, while those of finished
wenzelm@26870
   649
  results or bound by @{command "let"} may occur in \emph{arbitrary}
wenzelm@26870
   650
  instances later.  Even though actual polymorphism should be rarely
wenzelm@26870
   651
  used in practice, this mechanism is essential to achieve proper
wenzelm@26870
   652
  incremental type-inference, as the user proceeds to build up the
wenzelm@26870
   653
  Isar proof text from left to right.
wenzelm@26870
   654
wenzelm@26870
   655
  \medskip Term abbreviations are quite different from local
wenzelm@26870
   656
  definitions as introduced via @{command "def"} (see
wenzelm@26870
   657
  \secref{sec:proof-context}).  The latter are visible within the
wenzelm@26870
   658
  logic as actual equations, while abbreviations disappear during the
wenzelm@26870
   659
  input process just after type checking.  Also note that @{command
wenzelm@26870
   660
  "def"} does not support polymorphism.
wenzelm@26870
   661
wenzelm@26870
   662
  \begin{rail}
wenzelm@26870
   663
    'let' ((term + 'and') '=' term + 'and')
wenzelm@26870
   664
    ;  
wenzelm@26870
   665
  \end{rail}
wenzelm@26870
   666
wenzelm@26870
   667
  The syntax of @{keyword "is"} patterns follows \railnonterm{termpat}
wenzelm@26870
   668
  or \railnonterm{proppat} (see \secref{sec:term-decls}).
wenzelm@26870
   669
wenzelm@26870
   670
  \begin{descr}
wenzelm@26870
   671
wenzelm@26870
   672
  \item [@{command "let"}~@{text "p\<^sub>1 = t\<^sub>1 \<AND> \<dots>
wenzelm@26870
   673
  p\<^sub>n = t\<^sub>n"}] binds any text variables in patterns @{text
wenzelm@26870
   674
  "p\<^sub>1, \<dots>, p\<^sub>n"} by simultaneous higher-order matching
wenzelm@26870
   675
  against terms @{text "t\<^sub>1, \<dots>, t\<^sub>n"}.
wenzelm@26870
   676
wenzelm@26870
   677
  \item [@{text "(\<IS> p\<^sub>1 \<dots> p\<^sub>n)"}] resembles @{command
wenzelm@26870
   678
  "let"}, but matches @{text "p\<^sub>1, \<dots>, p\<^sub>n"} against the
wenzelm@26870
   679
  preceding statement.  Also note that @{keyword "is"} is not a
wenzelm@26870
   680
  separate command, but part of others (such as @{command "assume"},
wenzelm@26870
   681
  @{command "have"} etc.).
wenzelm@26870
   682
wenzelm@26870
   683
  \end{descr}
wenzelm@26870
   684
wenzelm@26870
   685
  Some \emph{implicit} term abbreviations\index{term abbreviations}
wenzelm@26870
   686
  for goals and facts are available as well.  For any open goal,
wenzelm@26870
   687
  @{variable_ref thesis} refers to its object-level statement,
wenzelm@26870
   688
  abstracted over any meta-level parameters (if present).  Likewise,
wenzelm@26870
   689
  @{variable_ref this} is bound for fact statements resulting from
wenzelm@26870
   690
  assumptions or finished goals.  In case @{variable this} refers to
wenzelm@26870
   691
  an object-logic statement that is an application @{text "f t"}, then
wenzelm@26870
   692
  @{text t} is bound to the special text variable ``@{variable "\<dots>"}''
wenzelm@26870
   693
  (three dots).  The canonical application of this convenience are
wenzelm@26870
   694
  calculational proofs (see \secref{sec:calculation}).
wenzelm@26870
   695
*}
wenzelm@26870
   696
wenzelm@26870
   697
wenzelm@26870
   698
section {* Block structure *}
wenzelm@26870
   699
wenzelm@26870
   700
text {*
wenzelm@26870
   701
  \begin{matharray}{rcl}
wenzelm@26870
   702
    @{command_def "next"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   703
    @{command_def "{"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   704
    @{command_def "}"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   705
  \end{matharray}
wenzelm@26870
   706
wenzelm@26870
   707
  While Isar is inherently block-structured, opening and closing
wenzelm@26870
   708
  blocks is mostly handled rather casually, with little explicit
wenzelm@26870
   709
  user-intervention.  Any local goal statement automatically opens
wenzelm@26870
   710
  \emph{two} internal blocks, which are closed again when concluding
wenzelm@26870
   711
  the sub-proof (by @{command "qed"} etc.).  Sections of different
wenzelm@26870
   712
  context within a sub-proof may be switched via @{command "next"},
wenzelm@26870
   713
  which is just a single block-close followed by block-open again.
wenzelm@26870
   714
  The effect of @{command "next"} is to reset the local proof context;
wenzelm@26870
   715
  there is no goal focus involved here!
wenzelm@26870
   716
wenzelm@26870
   717
  For slightly more advanced applications, there are explicit block
wenzelm@26870
   718
  parentheses as well.  These typically achieve a stronger forward
wenzelm@26870
   719
  style of reasoning.
wenzelm@26870
   720
wenzelm@26870
   721
  \begin{descr}
wenzelm@26870
   722
wenzelm@26870
   723
  \item [@{command "next"}] switches to a fresh block within a
wenzelm@26870
   724
  sub-proof, resetting the local context to the initial one.
wenzelm@26870
   725
wenzelm@26870
   726
  \item [@{command "{"} and @{command "}"}] explicitly open and close
wenzelm@26870
   727
  blocks.  Any current facts pass through ``@{command "{"}''
wenzelm@26870
   728
  unchanged, while ``@{command "}"}'' causes any result to be
wenzelm@26870
   729
  \emph{exported} into the enclosing context.  Thus fixed variables
wenzelm@26870
   730
  are generalized, assumptions discharged, and local definitions
wenzelm@26870
   731
  unfolded (cf.\ \secref{sec:proof-context}).  There is no difference
wenzelm@26870
   732
  of @{command "assume"} and @{command "presume"} in this mode of
wenzelm@26870
   733
  forward reasoning --- in contrast to plain backward reasoning with
wenzelm@26870
   734
  the result exported at @{command "show"} time.
wenzelm@26870
   735
wenzelm@26870
   736
  \end{descr}
wenzelm@26870
   737
*}
wenzelm@26870
   738
wenzelm@26870
   739
wenzelm@26870
   740
section {* Emulating tactic scripts \label{sec:tactic-commands} *}
wenzelm@26870
   741
wenzelm@26870
   742
text {*
wenzelm@26870
   743
  The Isar provides separate commands to accommodate tactic-style
wenzelm@26870
   744
  proof scripts within the same system.  While being outside the
wenzelm@26870
   745
  orthodox Isar proof language, these might come in handy for
wenzelm@26870
   746
  interactive exploration and debugging, or even actual tactical proof
wenzelm@26870
   747
  within new-style theories (to benefit from document preparation, for
wenzelm@26870
   748
  example).  See also \secref{sec:tactics} for actual tactics, that
wenzelm@26870
   749
  have been encapsulated as proof methods.  Proper proof methods may
wenzelm@26870
   750
  be used in scripts, too.
wenzelm@26870
   751
wenzelm@26870
   752
  \begin{matharray}{rcl}
wenzelm@26870
   753
    @{command_def "apply"}@{text "\<^sup>*"} & : & \isartrans{proof(prove)}{proof(prove)} \\
wenzelm@26870
   754
    @{command_def "apply_end"}@{text "\<^sup>*"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   755
    @{command_def "done"}@{text "\<^sup>*"} & : & \isartrans{proof(prove)}{proof(state)} \\
wenzelm@26870
   756
    @{command_def "defer"}@{text "\<^sup>*"} & : & \isartrans{proof}{proof} \\
wenzelm@26870
   757
    @{command_def "prefer"}@{text "\<^sup>*"} & : & \isartrans{proof}{proof} \\
wenzelm@26870
   758
    @{command_def "back"}@{text "\<^sup>*"} & : & \isartrans{proof}{proof} \\
wenzelm@26870
   759
  \end{matharray}
wenzelm@26870
   760
wenzelm@26870
   761
  \begin{rail}
wenzelm@26870
   762
    ( 'apply' | 'apply\_end' ) method
wenzelm@26870
   763
    ;
wenzelm@26870
   764
    'defer' nat?
wenzelm@26870
   765
    ;
wenzelm@26870
   766
    'prefer' nat
wenzelm@26870
   767
    ;
wenzelm@26870
   768
  \end{rail}
wenzelm@26870
   769
wenzelm@26870
   770
  \begin{descr}
wenzelm@26870
   771
wenzelm@26870
   772
  \item [@{command "apply"}~@{text m}] applies proof method @{text m}
wenzelm@26870
   773
  in initial position, but unlike @{command "proof"} it retains
wenzelm@26870
   774
  ``@{text "proof(prove)"}'' mode.  Thus consecutive method
wenzelm@26870
   775
  applications may be given just as in tactic scripts.
wenzelm@26870
   776
  
wenzelm@26870
   777
  Facts are passed to @{text m} as indicated by the goal's
wenzelm@26870
   778
  forward-chain mode, and are \emph{consumed} afterwards.  Thus any
wenzelm@26870
   779
  further @{command "apply"} command would always work in a purely
wenzelm@26870
   780
  backward manner.
wenzelm@26870
   781
  
wenzelm@26870
   782
  \item [@{command "apply_end"}~@{text "m"}] applies proof method
wenzelm@26870
   783
  @{text m} as if in terminal position.  Basically, this simulates a
wenzelm@26870
   784
  multi-step tactic script for @{command "qed"}, but may be given
wenzelm@26870
   785
  anywhere within the proof body.
wenzelm@26870
   786
  
wenzelm@26894
   787
  No facts are passed to @{text m} here.  Furthermore, the static
wenzelm@26870
   788
  context is that of the enclosing goal (as for actual @{command
wenzelm@26870
   789
  "qed"}).  Thus the proof method may not refer to any assumptions
wenzelm@26870
   790
  introduced in the current body, for example.
wenzelm@26870
   791
  
wenzelm@26870
   792
  \item [@{command "done"}] completes a proof script, provided that
wenzelm@26870
   793
  the current goal state is solved completely.  Note that actual
wenzelm@26870
   794
  structured proof commands (e.g.\ ``@{command "."}'' or @{command
wenzelm@26870
   795
  "sorry"}) may be used to conclude proof scripts as well.
wenzelm@26870
   796
wenzelm@26870
   797
  \item [@{command "defer"}~@{text n} and @{command "prefer"}~@{text
wenzelm@26870
   798
  n}] shuffle the list of pending goals: @{command "defer"} puts off
wenzelm@26870
   799
  sub-goal @{text n} to the end of the list (@{text "n = 1"} by
wenzelm@26870
   800
  default), while @{command "prefer"} brings sub-goal @{text n} to the
wenzelm@26870
   801
  front.
wenzelm@26870
   802
  
wenzelm@26870
   803
  \item [@{command "back"}] does back-tracking over the result
wenzelm@26870
   804
  sequence of the latest proof command.  Basically, any proof command
wenzelm@26870
   805
  may return multiple results.
wenzelm@26870
   806
  
wenzelm@26870
   807
  \end{descr}
wenzelm@26870
   808
wenzelm@26870
   809
  Any proper Isar proof method may be used with tactic script commands
wenzelm@26870
   810
  such as @{command "apply"}.  A few additional emulations of actual
wenzelm@26870
   811
  tactics are provided as well; these would be never used in actual
wenzelm@26870
   812
  structured proofs, of course.
wenzelm@26870
   813
*}
wenzelm@26870
   814
wenzelm@26870
   815
wenzelm@26870
   816
section {* Omitting proofs *}
wenzelm@26870
   817
wenzelm@26870
   818
text {*
wenzelm@26870
   819
  \begin{matharray}{rcl}
wenzelm@26870
   820
    @{command_def "oops"} & : & \isartrans{proof}{theory} \\
wenzelm@26870
   821
  \end{matharray}
wenzelm@26870
   822
wenzelm@26870
   823
  The @{command "oops"} command discontinues the current proof
wenzelm@26870
   824
  attempt, while considering the partial proof text as properly
wenzelm@26870
   825
  processed.  This is conceptually quite different from ``faking''
wenzelm@26870
   826
  actual proofs via @{command_ref "sorry"} (see
wenzelm@26870
   827
  \secref{sec:proof-steps}): @{command "oops"} does not observe the
wenzelm@26870
   828
  proof structure at all, but goes back right to the theory level.
wenzelm@26870
   829
  Furthermore, @{command "oops"} does not produce any result theorem
wenzelm@26870
   830
  --- there is no intended claim to be able to complete the proof
wenzelm@26870
   831
  anyhow.
wenzelm@26870
   832
wenzelm@26870
   833
  A typical application of @{command "oops"} is to explain Isar proofs
wenzelm@26870
   834
  \emph{within} the system itself, in conjunction with the document
wenzelm@26870
   835
  preparation tools of Isabelle described in \cite{isabelle-sys}.
wenzelm@26870
   836
  Thus partial or even wrong proof attempts can be discussed in a
wenzelm@26870
   837
  logically sound manner.  Note that the Isabelle {\LaTeX} macros can
wenzelm@26870
   838
  be easily adapted to print something like ``@{text "\<dots>"}'' instead of
wenzelm@26870
   839
  the keyword ``@{command "oops"}''.
wenzelm@26870
   840
wenzelm@26870
   841
  \medskip The @{command "oops"} command is undo-able, unlike
wenzelm@26870
   842
  @{command_ref "kill"} (see \secref{sec:history}).  The effect is to
wenzelm@26870
   843
  get back to the theory just before the opening of the proof.
wenzelm@26870
   844
*}
wenzelm@26870
   845
wenzelm@26870
   846
wenzelm@26870
   847
section {* Generalized elimination \label{sec:obtain} *}
wenzelm@26870
   848
wenzelm@26870
   849
text {*
wenzelm@26870
   850
  \begin{matharray}{rcl}
wenzelm@26870
   851
    @{command_def "obtain"} & : & \isartrans{proof(state)}{proof(prove)} \\
wenzelm@26870
   852
    @{command_def "guess"}@{text "\<^sup>*"} & : & \isartrans{proof(state)}{proof(prove)} \\
wenzelm@26870
   853
  \end{matharray}
wenzelm@26870
   854
wenzelm@26870
   855
  Generalized elimination means that additional elements with certain
wenzelm@26870
   856
  properties may be introduced in the current context, by virtue of a
wenzelm@26870
   857
  locally proven ``soundness statement''.  Technically speaking, the
wenzelm@26870
   858
  @{command "obtain"} language element is like a declaration of
wenzelm@26870
   859
  @{command "fix"} and @{command "assume"} (see also see
wenzelm@26870
   860
  \secref{sec:proof-context}), together with a soundness proof of its
wenzelm@26870
   861
  additional claim.  According to the nature of existential reasoning,
wenzelm@26870
   862
  assumptions get eliminated from any result exported from the context
wenzelm@26870
   863
  later, provided that the corresponding parameters do \emph{not}
wenzelm@26870
   864
  occur in the conclusion.
wenzelm@26870
   865
wenzelm@26870
   866
  \begin{rail}
wenzelm@26870
   867
    'obtain' parname? (vars + 'and') 'where' (props + 'and')
wenzelm@26870
   868
    ;
wenzelm@26870
   869
    'guess' (vars + 'and')
wenzelm@26870
   870
    ;
wenzelm@26870
   871
  \end{rail}
wenzelm@26870
   872
wenzelm@26870
   873
  The derived Isar command @{command "obtain"} is defined as follows
wenzelm@26870
   874
  (where @{text "b\<^sub>1, \<dots>, b\<^sub>k"} shall refer to (optional)
wenzelm@26870
   875
  facts indicated for forward chaining).
wenzelm@26870
   876
  \begin{matharray}{l}
wenzelm@26870
   877
    @{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]
wenzelm@26870
   878
    \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"} \\
wenzelm@26870
   879
    \quad @{command "proof"}~@{text succeed} \\
wenzelm@26870
   880
    \qquad @{command "fix"}~@{text thesis} \\
wenzelm@26870
   881
    \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"} \\
wenzelm@26870
   882
    \qquad @{command "then"}~@{command "show"}~@{text thesis} \\
wenzelm@26870
   883
    \quad\qquad @{command "apply"}~@{text -} \\
wenzelm@26870
   884
    \quad\qquad @{command "using"}~@{text "b\<^sub>1 \<dots> b\<^sub>k  \<langle>proof\<rangle>"} \\
wenzelm@26870
   885
    \quad @{command "qed"} \\
wenzelm@26870
   886
    \quad @{command "fix"}~@{text "x\<^sub>1 \<dots> x\<^sub>m"}~@{command "assume"}@{text "\<^sup>* a: \<phi>\<^sub>1 \<dots> \<phi>\<^sub>n"} \\
wenzelm@26870
   887
  \end{matharray}
wenzelm@26870
   888
wenzelm@26870
   889
  Typically, the soundness proof is relatively straight-forward, often
wenzelm@26870
   890
  just by canonical automated tools such as ``@{command "by"}~@{text
wenzelm@26870
   891
  simp}'' or ``@{command "by"}~@{text blast}''.  Accordingly, the
wenzelm@26870
   892
  ``@{text that}'' reduction above is declared as simplification and
wenzelm@26870
   893
  introduction rule.
wenzelm@26870
   894
wenzelm@26870
   895
  In a sense, @{command "obtain"} represents at the level of Isar
wenzelm@26870
   896
  proofs what would be meta-logical existential quantifiers and
wenzelm@26870
   897
  conjunctions.  This concept has a broad range of useful
wenzelm@26870
   898
  applications, ranging from plain elimination (or introduction) of
wenzelm@26870
   899
  object-level existential and conjunctions, to elimination over
wenzelm@26870
   900
  results of symbolic evaluation of recursive definitions, for
wenzelm@26870
   901
  example.  Also note that @{command "obtain"} without parameters acts
wenzelm@26870
   902
  much like @{command "have"}, where the result is treated as a
wenzelm@26870
   903
  genuine assumption.
wenzelm@26870
   904
wenzelm@26870
   905
  An alternative name to be used instead of ``@{text that}'' above may
wenzelm@26870
   906
  be given in parentheses.
wenzelm@26870
   907
wenzelm@26870
   908
  \medskip The improper variant @{command "guess"} is similar to
wenzelm@26870
   909
  @{command "obtain"}, but derives the obtained statement from the
wenzelm@26870
   910
  course of reasoning!  The proof starts with a fixed goal @{text
wenzelm@26870
   911
  thesis}.  The subsequent proof may refine this to anything of the
wenzelm@26870
   912
  form like @{text "\<And>x\<^sub>1 \<dots> x\<^sub>m. \<phi>\<^sub>1 \<Longrightarrow> \<dots>
wenzelm@26870
   913
  \<phi>\<^sub>n \<Longrightarrow> thesis"}, but must not introduce new subgoals.  The
wenzelm@26870
   914
  final goal state is then used as reduction rule for the obtain
wenzelm@26870
   915
  scheme described above.  Obtained parameters @{text "x\<^sub>1, \<dots>,
wenzelm@26870
   916
  x\<^sub>m"} are marked as internal by default, which prevents the
wenzelm@26870
   917
  proof context from being polluted by ad-hoc variables.  The variable
wenzelm@26870
   918
  names and type constraints given as arguments for @{command "guess"}
wenzelm@26870
   919
  specify a prefix of obtained parameters explicitly in the text.
wenzelm@26870
   920
wenzelm@26870
   921
  It is important to note that the facts introduced by @{command
wenzelm@26870
   922
  "obtain"} and @{command "guess"} may not be polymorphic: any
wenzelm@26870
   923
  type-variables occurring here are fixed in the present context!
wenzelm@26870
   924
*}
wenzelm@26870
   925
wenzelm@26870
   926
wenzelm@26870
   927
section {* Calculational reasoning \label{sec:calculation} *}
wenzelm@26870
   928
wenzelm@26870
   929
text {*
wenzelm@26870
   930
  \begin{matharray}{rcl}
wenzelm@26870
   931
    @{command_def "also"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   932
    @{command_def "finally"} & : & \isartrans{proof(state)}{proof(chain)} \\
wenzelm@26870
   933
    @{command_def "moreover"} & : & \isartrans{proof(state)}{proof(state)} \\
wenzelm@26870
   934
    @{command_def "ultimately"} & : & \isartrans{proof(state)}{proof(chain)} \\
wenzelm@26870
   935
    @{command_def "print_trans_rules"}@{text "\<^sup>*"} & : & \isarkeep{theory~|~proof} \\
wenzelm@26870
   936
    @{attribute trans} & : & \isaratt \\
wenzelm@26870
   937
    @{attribute sym} & : & \isaratt \\
wenzelm@26870
   938
    @{attribute symmetric} & : & \isaratt \\
wenzelm@26870
   939
  \end{matharray}
wenzelm@26870
   940
wenzelm@26870
   941
  Calculational proof is forward reasoning with implicit application
wenzelm@26870
   942
  of transitivity rules (such those of @{text "="}, @{text "\<le>"},
wenzelm@26870
   943
  @{text "<"}).  Isabelle/Isar maintains an auxiliary fact register
wenzelm@26870
   944
  @{fact_ref calculation} for accumulating results obtained by
wenzelm@26870
   945
  transitivity composed with the current result.  Command @{command
wenzelm@26870
   946
  "also"} updates @{fact calculation} involving @{fact this}, while
wenzelm@26870
   947
  @{command "finally"} exhibits the final @{fact calculation} by
wenzelm@26870
   948
  forward chaining towards the next goal statement.  Both commands
wenzelm@26870
   949
  require valid current facts, i.e.\ may occur only after commands
wenzelm@26870
   950
  that produce theorems such as @{command "assume"}, @{command
wenzelm@26870
   951
  "note"}, or some finished proof of @{command "have"}, @{command
wenzelm@26870
   952
  "show"} etc.  The @{command "moreover"} and @{command "ultimately"}
wenzelm@26870
   953
  commands are similar to @{command "also"} and @{command "finally"},
wenzelm@26870
   954
  but only collect further results in @{fact calculation} without
wenzelm@26870
   955
  applying any rules yet.
wenzelm@26870
   956
wenzelm@26870
   957
  Also note that the implicit term abbreviation ``@{text "\<dots>"}'' has
wenzelm@26870
   958
  its canonical application with calculational proofs.  It refers to
wenzelm@26870
   959
  the argument of the preceding statement. (The argument of a curried
wenzelm@26870
   960
  infix expression happens to be its right-hand side.)
wenzelm@26870
   961
wenzelm@26870
   962
  Isabelle/Isar calculations are implicitly subject to block structure
wenzelm@26870
   963
  in the sense that new threads of calculational reasoning are
wenzelm@26870
   964
  commenced for any new block (as opened by a local goal, for
wenzelm@26870
   965
  example).  This means that, apart from being able to nest
wenzelm@26870
   966
  calculations, there is no separate \emph{begin-calculation} command
wenzelm@26870
   967
  required.
wenzelm@26870
   968
wenzelm@26870
   969
  \medskip The Isar calculation proof commands may be defined as
wenzelm@26870
   970
  follows:\footnote{We suppress internal bookkeeping such as proper
wenzelm@26870
   971
  handling of block-structure.}
wenzelm@26870
   972
wenzelm@26870
   973
  \begin{matharray}{rcl}
wenzelm@26870
   974
    @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\
wenzelm@26870
   975
    @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex]
wenzelm@26870
   976
    @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex]
wenzelm@26870
   977
    @{command "moreover"} & \equiv & @{command "note"}~@{text "calculation = calculation this"} \\
wenzelm@26870
   978
    @{command "ultimately"} & \equiv & @{command "moreover"}~@{command "from"}~@{text calculation} \\
wenzelm@26870
   979
  \end{matharray}
wenzelm@26870
   980
wenzelm@26870
   981
  \begin{rail}
wenzelm@26870
   982
    ('also' | 'finally') ('(' thmrefs ')')?
wenzelm@26870
   983
    ;
wenzelm@26870
   984
    'trans' (() | 'add' | 'del')
wenzelm@26870
   985
    ;
wenzelm@26870
   986
  \end{rail}
wenzelm@26870
   987
wenzelm@26870
   988
  \begin{descr}
wenzelm@26870
   989
wenzelm@26870
   990
  \item [@{command "also"}~@{text "(a\<^sub>1 \<dots> a\<^sub>n)"}]
wenzelm@26870
   991
  maintains the auxiliary @{fact calculation} register as follows.
wenzelm@26870
   992
  The first occurrence of @{command "also"} in some calculational
wenzelm@26870
   993
  thread initializes @{fact calculation} by @{fact this}. Any
wenzelm@26870
   994
  subsequent @{command "also"} on the same level of block-structure
wenzelm@26870
   995
  updates @{fact calculation} by some transitivity rule applied to
wenzelm@26870
   996
  @{fact calculation} and @{fact this} (in that order).  Transitivity
wenzelm@26870
   997
  rules are picked from the current context, unless alternative rules
wenzelm@26870
   998
  are given as explicit arguments.
wenzelm@26870
   999
wenzelm@26870
  1000
  \item [@{command "finally"}~@{text "(a\<^sub>1 \<dots> a\<^sub>n)"}]
wenzelm@26870
  1001
  maintaining @{fact calculation} in the same way as @{command
wenzelm@26870
  1002
  "also"}, and concludes the current calculational thread.  The final
wenzelm@26870
  1003
  result is exhibited as fact for forward chaining towards the next
wenzelm@26870
  1004
  goal. Basically, @{command "finally"} just abbreviates @{command
wenzelm@26870
  1005
  "also"}~@{command "from"}~@{fact calculation}.  Typical idioms for
wenzelm@26870
  1006
  concluding calculational proofs are ``@{command "finally"}~@{command
wenzelm@26870
  1007
  "show"}~@{text ?thesis}~@{command "."}'' and ``@{command
wenzelm@26870
  1008
  "finally"}~@{command "have"}~@{text \<phi>}~@{command "."}''.
wenzelm@26870
  1009
wenzelm@26870
  1010
  \item [@{command "moreover"} and @{command "ultimately"}] are
wenzelm@26870
  1011
  analogous to @{command "also"} and @{command "finally"}, but collect
wenzelm@26870
  1012
  results only, without applying rules.
wenzelm@26870
  1013
wenzelm@26870
  1014
  \item [@{command "print_trans_rules"}] prints the list of
wenzelm@26870
  1015
  transitivity rules (for calculational commands @{command "also"} and
wenzelm@26870
  1016
  @{command "finally"}) and symmetry rules (for the @{attribute
wenzelm@26870
  1017
  symmetric} operation and single step elimination patters) of the
wenzelm@26870
  1018
  current context.
wenzelm@26870
  1019
wenzelm@26870
  1020
  \item [@{attribute trans}] declares theorems as transitivity rules.
wenzelm@26870
  1021
wenzelm@26870
  1022
  \item [@{attribute sym}] declares symmetry rules, as well as
wenzelm@26894
  1023
  @{attribute "Pure.elim"}@{text "?"} rules.
wenzelm@26870
  1024
wenzelm@26870
  1025
  \item [@{attribute symmetric}] resolves a theorem with some rule
wenzelm@26870
  1026
  declared as @{attribute sym} in the current context.  For example,
wenzelm@26870
  1027
  ``@{command "assume"}~@{text "[symmetric]: x = y"}'' produces a
wenzelm@26870
  1028
  swapped fact derived from that assumption.
wenzelm@26870
  1029
wenzelm@26870
  1030
  In structured proof texts it is often more appropriate to use an
wenzelm@26870
  1031
  explicit single-step elimination proof, such as ``@{command
wenzelm@26870
  1032
  "assume"}~@{text "x = y"}~@{command "then"}~@{command "have"}~@{text
wenzelm@26870
  1033
  "y = x"}~@{command ".."}''.
wenzelm@26870
  1034
wenzelm@26870
  1035
  \end{descr}
wenzelm@26870
  1036
*}
wenzelm@26870
  1037
wenzelm@26869
  1038
end