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