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