doc-src/IsarOverview/Isar/Logic.thy
author nipkow
Sun, 11 Jun 2006 19:36:10 +0200
changeset 19840 600c35fd1b5e
parent 17914 99ead7a7eb42
child 25412 6f56f0350f6c
permissions -rw-r--r--
added quoting via back quotes
wenzelm@17914
     1
(*<*)theory Logic imports Main begin(*>*)
kleing@13999
     2
kleing@13999
     3
section{*Logic \label{sec:Logic}*}
kleing@13999
     4
kleing@13999
     5
subsection{*Propositional logic*}
kleing@13999
     6
kleing@13999
     7
subsubsection{*Introduction rules*}
kleing@13999
     8
kleing@13999
     9
text{* We start with a really trivial toy proof to introduce the basic
kleing@13999
    10
features of structured proofs. *}
kleing@13999
    11
lemma "A \<longrightarrow> A"
kleing@13999
    12
proof (rule impI)
kleing@13999
    13
  assume a: "A"
kleing@13999
    14
  show "A" by(rule a)
kleing@13999
    15
qed
kleing@13999
    16
text{*\noindent
kleing@13999
    17
The operational reading: the \isakeyword{assume}-\isakeyword{show}
kleing@13999
    18
block proves @{prop"A \<Longrightarrow> A"} (@{term a} is a degenerate rule (no
kleing@13999
    19
assumptions) that proves @{term A} outright), which rule
kleing@13999
    20
@{thm[source]impI} (@{thm impI}) turns into the desired @{prop"A \<longrightarrow>
kleing@13999
    21
A"}.  However, this text is much too detailed for comfort. Therefore
kleing@13999
    22
Isar implements the following principle: \begin{quote}\em Command
kleing@13999
    23
\isakeyword{proof} automatically tries to select an introduction rule
kleing@13999
    24
based on the goal and a predefined list of rules.  \end{quote} Here
kleing@13999
    25
@{thm[source]impI} is applied automatically: *}
kleing@13999
    26
kleing@13999
    27
lemma "A \<longrightarrow> A"
kleing@13999
    28
proof
kleing@13999
    29
  assume a: A
kleing@13999
    30
  show A by(rule a)
kleing@13999
    31
qed
kleing@13999
    32
kleing@13999
    33
text{*\noindent Single-identifier formulae such as @{term A} need not
kleing@13999
    34
be enclosed in double quotes. However, we will continue to do so for
kleing@13999
    35
uniformity.
kleing@13999
    36
kleing@13999
    37
Trivial proofs, in particular those by assumption, should be trivial
kleing@13999
    38
to perform. Proof ``.'' does just that (and a bit more). Thus
kleing@13999
    39
naming of assumptions is often superfluous: *}
kleing@13999
    40
lemma "A \<longrightarrow> A"
kleing@13999
    41
proof
kleing@13999
    42
  assume "A"
kleing@13999
    43
  show "A" .
kleing@13999
    44
qed
kleing@13999
    45
kleing@13999
    46
text{* To hide proofs by assumption further, \isakeyword{by}@{text"(method)"}
kleing@13999
    47
first applies @{text method} and then tries to solve all remaining subgoals
kleing@13999
    48
by assumption: *}
kleing@13999
    49
lemma "A \<longrightarrow> A \<and> A"
kleing@13999
    50
proof
kleing@13999
    51
  assume "A"
kleing@13999
    52
  show "A \<and> A" by(rule conjI)
kleing@13999
    53
qed
kleing@13999
    54
text{*\noindent Rule @{thm[source]conjI} is of course @{thm conjI}.
kleing@13999
    55
A drawback of implicit proofs by assumption is that it
kleing@13999
    56
is no longer obvious where an assumption is used.
kleing@13999
    57
kleing@13999
    58
Proofs of the form \isakeyword{by}@{text"(rule"}~\emph{name}@{text")"}
kleing@13999
    59
can be abbreviated to ``..''  if \emph{name} refers to one of the
kleing@13999
    60
predefined introduction rules (or elimination rules, see below): *}
kleing@13999
    61
kleing@13999
    62
lemma "A \<longrightarrow> A \<and> A"
kleing@13999
    63
proof
kleing@13999
    64
  assume "A"
kleing@13999
    65
  show "A \<and> A" ..
kleing@13999
    66
qed
kleing@13999
    67
text{*\noindent
kleing@13999
    68
This is what happens: first the matching introduction rule @{thm[source]conjI}
kleing@13999
    69
is applied (first ``.''), then the two subgoals are solved by assumption
kleing@13999
    70
(second ``.''). *}
kleing@13999
    71
kleing@13999
    72
subsubsection{*Elimination rules*}
kleing@13999
    73
kleing@13999
    74
text{*A typical elimination rule is @{thm[source]conjE}, $\land$-elimination:
kleing@13999
    75
@{thm[display,indent=5]conjE}  In the following proof it is applied
kleing@13999
    76
by hand, after its first (\emph{major}) premise has been eliminated via
kleing@13999
    77
@{text"[OF AB]"}: *}
kleing@13999
    78
lemma "A \<and> B \<longrightarrow> B \<and> A"
kleing@13999
    79
proof
kleing@13999
    80
  assume AB: "A \<and> B"
kleing@13999
    81
  show "B \<and> A"
kleing@13999
    82
  proof (rule conjE[OF AB])  -- {*@{text"conjE[OF AB]"}: @{thm conjE[OF AB]} *}
kleing@13999
    83
    assume "A" "B"
kleing@13999
    84
    show ?thesis ..
kleing@13999
    85
  qed
kleing@13999
    86
qed
kleing@13999
    87
text{*\noindent Note that the term @{text"?thesis"} always stands for the
kleing@13999
    88
``current goal'', i.e.\ the enclosing \isakeyword{show} (or
kleing@13999
    89
\isakeyword{have}) statement.
kleing@13999
    90
kleing@13999
    91
This is too much proof text. Elimination rules should be selected
kleing@13999
    92
automatically based on their major premise, the formula or rather connective
kleing@13999
    93
to be eliminated. In Isar they are triggered by facts being fed
kleing@13999
    94
\emph{into} a proof. Syntax:
kleing@13999
    95
\begin{center}
kleing@13999
    96
\isakeyword{from} \emph{fact} \isakeyword{show} \emph{proposition} \emph{proof}
kleing@13999
    97
\end{center}
kleing@13999
    98
where \emph{fact} stands for the name of a previously proved
kleing@13999
    99
proposition, e.g.\ an assumption, an intermediate result or some global
kleing@13999
   100
theorem, which may also be modified with @{text OF} etc.
kleing@13999
   101
The \emph{fact} is ``piped'' into the \emph{proof}, which can deal with it
kleing@13999
   102
how it chooses. If the \emph{proof} starts with a plain \isakeyword{proof},
kleing@13999
   103
an elimination rule (from a predefined list) is applied
kleing@13999
   104
whose first premise is solved by the \emph{fact}. Thus the proof above
kleing@13999
   105
is equivalent to the following one: *}
kleing@13999
   106
kleing@13999
   107
lemma "A \<and> B \<longrightarrow> B \<and> A"
kleing@13999
   108
proof
kleing@13999
   109
  assume AB: "A \<and> B"
kleing@13999
   110
  from AB show "B \<and> A"
kleing@13999
   111
  proof
kleing@13999
   112
    assume "A" "B"
kleing@13999
   113
    show ?thesis ..
kleing@13999
   114
  qed
kleing@13999
   115
qed
kleing@13999
   116
kleing@13999
   117
text{* Now we come to a second important principle:
kleing@13999
   118
\begin{quote}\em
kleing@13999
   119
Try to arrange the sequence of propositions in a UNIX-like pipe,
kleing@13999
   120
such that the proof of each proposition builds on the previous proposition.
kleing@13999
   121
\end{quote}
kleing@13999
   122
The previous proposition can be referred to via the fact @{text this}.
kleing@13999
   123
This greatly reduces the need for explicit naming of propositions:
kleing@13999
   124
*}
kleing@13999
   125
lemma "A \<and> B \<longrightarrow> B \<and> A"
kleing@13999
   126
proof
kleing@13999
   127
  assume "A \<and> B"
kleing@13999
   128
  from this show "B \<and> A"
kleing@13999
   129
  proof
kleing@13999
   130
    assume "A" "B"
kleing@13999
   131
    show ?thesis ..
kleing@13999
   132
  qed
kleing@13999
   133
qed
kleing@13999
   134
kleing@13999
   135
text{*\noindent Because of the frequency of \isakeyword{from}~@{text
kleing@13999
   136
this}, Isar provides two abbreviations:
kleing@13999
   137
\begin{center}
kleing@13999
   138
\begin{tabular}{r@ {\quad=\quad}l}
kleing@13999
   139
\isakeyword{then} & \isakeyword{from} @{text this} \\
kleing@13999
   140
\isakeyword{thus} & \isakeyword{then} \isakeyword{show}
kleing@13999
   141
\end{tabular}
kleing@13999
   142
\end{center}
kleing@13999
   143
kleing@13999
   144
Here is an alternative proof that operates purely by forward reasoning: *}
kleing@13999
   145
lemma "A \<and> B \<longrightarrow> B \<and> A"
kleing@13999
   146
proof
kleing@13999
   147
  assume ab: "A \<and> B"
kleing@13999
   148
  from ab have a: "A" ..
kleing@13999
   149
  from ab have b: "B" ..
kleing@13999
   150
  from b a show "B \<and> A" ..
kleing@13999
   151
qed
kleing@13999
   152
kleing@13999
   153
text{*\noindent It is worth examining this text in detail because it
kleing@13999
   154
exhibits a number of new concepts.  For a start, it is the first time
kleing@13999
   155
we have proved intermediate propositions (\isakeyword{have}) on the
kleing@13999
   156
way to the final \isakeyword{show}. This is the norm in nontrivial
kleing@13999
   157
proofs where one cannot bridge the gap between the assumptions and the
kleing@13999
   158
conclusion in one step. To understand how the proof works we need to
kleing@13999
   159
explain more Isar details.
kleing@13999
   160
kleing@13999
   161
Method @{text rule} can be given a list of rules, in which case
kleing@13999
   162
@{text"(rule"}~\textit{rules}@{text")"} applies the first matching
kleing@13999
   163
rule in the list \textit{rules}. Command \isakeyword{from} can be
kleing@13999
   164
followed by any number of facts.  Given \isakeyword{from}~@{text
kleing@13999
   165
f}$_1$~\dots~@{text f}$_n$, the proof step
kleing@13999
   166
@{text"(rule"}~\textit{rules}@{text")"} following a \isakeyword{have}
kleing@13999
   167
or \isakeyword{show} searches \textit{rules} for a rule whose first
kleing@13999
   168
$n$ premises can be proved by @{text f}$_1$~\dots~@{text f}$_n$ in the
kleing@13999
   169
given order. Finally one needs to know that ``..'' is short for
kleing@13999
   170
@{text"by(rule"}~\textit{elim-rules intro-rules}@{text")"} (or
kleing@13999
   171
@{text"by(rule"}~\textit{intro-rules}@{text")"} if there are no facts
kleing@13999
   172
fed into the proof), i.e.\ elimination rules are tried before
kleing@13999
   173
introduction rules.
kleing@13999
   174
kleing@13999
   175
Thus in the above proof both \isakeyword{have}s are proved via
kleing@13999
   176
@{thm[source]conjE} triggered by \isakeyword{from}~@{text ab} whereas
kleing@13999
   177
in the \isakeyword{show} step no elimination rule is applicable and
kleing@13999
   178
the proof succeeds with @{thm[source]conjI}. The latter would fail had
kleing@13999
   179
we written \isakeyword{from}~@{text"a b"} instead of
kleing@13999
   180
\isakeyword{from}~@{text"b a"}.
kleing@13999
   181
kleing@13999
   182
Proofs starting with a plain @{text proof} behave the same because the
kleing@13999
   183
latter is short for @{text"proof (rule"}~\textit{elim-rules
kleing@13999
   184
intro-rules}@{text")"} (or @{text"proof
kleing@13999
   185
(rule"}~\textit{intro-rules}@{text")"} if there are no facts fed into
kleing@13999
   186
the proof). *}
kleing@13999
   187
kleing@13999
   188
subsection{*More constructs*}
kleing@13999
   189
kleing@13999
   190
text{* In the previous proof of @{prop"A \<and> B \<longrightarrow> B \<and> A"} we needed to feed
kleing@13999
   191
more than one fact into a proof step, a frequent situation. Then the
kleing@13999
   192
UNIX-pipe model appears to break down and we need to name the different
kleing@13999
   193
facts to refer to them. But this can be avoided:
kleing@13999
   194
*}
kleing@13999
   195
lemma "A \<and> B \<longrightarrow> B \<and> A"
kleing@13999
   196
proof
kleing@13999
   197
  assume ab: "A \<and> B"
kleing@13999
   198
  from ab have "B" ..
kleing@13999
   199
  moreover
kleing@13999
   200
  from ab have "A" ..
kleing@13999
   201
  ultimately show "B \<and> A" ..
kleing@13999
   202
qed
kleing@13999
   203
text{*\noindent You can combine any number of facts @{term A1} \dots\ @{term
kleing@13999
   204
An} into a sequence by separating their proofs with
kleing@13999
   205
\isakeyword{moreover}. After the final fact, \isakeyword{ultimately} stands
kleing@13999
   206
for \isakeyword{from}~@{term A1}~\dots~@{term An}.  This avoids having to
kleing@13999
   207
introduce names for all of the sequence elements.  *}
kleing@13999
   208
kleing@13999
   209
text{* Although we have only seen a few introduction and elimination rules so
kleing@13999
   210
far, Isar's predefined rules include all the usual natural deduction
kleing@13999
   211
rules. We conclude our exposition of propositional logic with an extended
kleing@13999
   212
example --- which rules are used implicitly where? *}
kleing@13999
   213
lemma "\<not> (A \<and> B) \<longrightarrow> \<not> A \<or> \<not> B"
kleing@13999
   214
proof
kleing@13999
   215
  assume n: "\<not> (A \<and> B)"
kleing@13999
   216
  show "\<not> A \<or> \<not> B"
kleing@13999
   217
  proof (rule ccontr)
kleing@13999
   218
    assume nn: "\<not> (\<not> A \<or> \<not> B)"
kleing@13999
   219
    have "\<not> A"
kleing@13999
   220
    proof
kleing@13999
   221
      assume "A"
kleing@13999
   222
      have "\<not> B"
kleing@13999
   223
      proof
kleing@13999
   224
        assume "B"
kleing@13999
   225
        have "A \<and> B" ..
kleing@13999
   226
        with n show False ..
kleing@13999
   227
      qed
kleing@13999
   228
      hence "\<not> A \<or> \<not> B" ..
kleing@13999
   229
      with nn show False ..
kleing@13999
   230
    qed
kleing@13999
   231
    hence "\<not> A \<or> \<not> B" ..
kleing@13999
   232
    with nn show False ..
kleing@13999
   233
  qed
kleing@13999
   234
qed
kleing@13999
   235
text{*\noindent
kleing@13999
   236
Rule @{thm[source]ccontr} (``classical contradiction'') is
kleing@13999
   237
@{thm ccontr[no_vars]}.
kleing@13999
   238
Apart from demonstrating the strangeness of classical
kleing@13999
   239
arguments by contradiction, this example also introduces two new
kleing@13999
   240
abbreviations:
kleing@13999
   241
\begin{center}
kleing@13999
   242
\begin{tabular}{l@ {\quad=\quad}l}
kleing@13999
   243
\isakeyword{hence} & \isakeyword{then} \isakeyword{have} \\
kleing@13999
   244
\isakeyword{with}~\emph{facts} &
kleing@13999
   245
\isakeyword{from}~\emph{facts} @{text this}
kleing@13999
   246
\end{tabular}
kleing@13999
   247
\end{center}
kleing@13999
   248
*}
kleing@13999
   249
nipkow@19840
   250
kleing@13999
   251
subsection{*Avoiding duplication*}
kleing@13999
   252
kleing@13999
   253
text{* So far our examples have been a bit unnatural: normally we want to
kleing@13999
   254
prove rules expressed with @{text"\<Longrightarrow>"}, not @{text"\<longrightarrow>"}. Here is an example:
kleing@13999
   255
*}
kleing@13999
   256
lemma "A \<and> B \<Longrightarrow> B \<and> A"
kleing@13999
   257
proof
kleing@13999
   258
  assume "A \<and> B" thus "B" ..
kleing@13999
   259
next
kleing@13999
   260
  assume "A \<and> B" thus "A" ..
kleing@13999
   261
qed
kleing@13999
   262
text{*\noindent The \isakeyword{proof} always works on the conclusion,
kleing@13999
   263
@{prop"B \<and> A"} in our case, thus selecting $\land$-introduction. Hence
kleing@13999
   264
we must show @{prop B} and @{prop A}; both are proved by
kleing@13999
   265
$\land$-elimination and the proofs are separated by \isakeyword{next}:
kleing@13999
   266
\begin{description}
kleing@13999
   267
\item[\isakeyword{next}] deals with multiple subgoals. For example,
kleing@13999
   268
when showing @{term"A \<and> B"} we need to show both @{term A} and @{term
kleing@13999
   269
B}.  Each subgoal is proved separately, in \emph{any} order. The
kleing@13999
   270
individual proofs are separated by \isakeyword{next}.  \footnote{Each
kleing@13999
   271
\isakeyword{show} must prove one of the pending subgoals.  If a
kleing@13999
   272
\isakeyword{show} matches multiple subgoals, e.g.\ if the subgoals
kleing@13999
   273
contain ?-variables, the first one is proved. Thus the order in which
kleing@13999
   274
the subgoals are proved can matter --- see
kleing@13999
   275
\S\ref{sec:CaseDistinction} for an example.}
kleing@13999
   276
kleing@13999
   277
Strictly speaking \isakeyword{next} is only required if the subgoals
kleing@13999
   278
are proved in different assumption contexts which need to be
kleing@13999
   279
separated, which is not the case above. For clarity we
kleing@13999
   280
have employed \isakeyword{next} anyway and will continue to do so.
kleing@13999
   281
\end{description}
kleing@13999
   282
kleing@13999
   283
This is all very well as long as formulae are small. Let us now look at some
kleing@13999
   284
devices to avoid repeating (possibly large) formulae. A very general method
kleing@13999
   285
is pattern matching: *}
kleing@13999
   286
kleing@13999
   287
lemma "large_A \<and> large_B \<Longrightarrow> large_B \<and> large_A"
kleing@13999
   288
      (is "?AB \<Longrightarrow> ?B \<and> ?A")
kleing@13999
   289
proof
kleing@13999
   290
  assume "?AB" thus "?B" ..
kleing@13999
   291
next
kleing@13999
   292
  assume "?AB" thus "?A" ..
kleing@13999
   293
qed
kleing@13999
   294
text{*\noindent Any formula may be followed by
kleing@13999
   295
@{text"("}\isakeyword{is}~\emph{pattern}@{text")"} which causes the pattern
kleing@13999
   296
to be matched against the formula, instantiating the @{text"?"}-variables in
kleing@13999
   297
the pattern. Subsequent uses of these variables in other terms causes
kleing@13999
   298
them to be replaced by the terms they stand for.
kleing@13999
   299
kleing@13999
   300
We can simplify things even more by stating the theorem by means of the
kleing@13999
   301
\isakeyword{assumes} and \isakeyword{shows} elements which allow direct
kleing@13999
   302
naming of assumptions: *}
kleing@13999
   303
kleing@13999
   304
lemma assumes AB: "large_A \<and> large_B"
kleing@13999
   305
  shows "large_B \<and> large_A" (is "?B \<and> ?A")
kleing@13999
   306
proof
kleing@13999
   307
  from AB show "?B" ..
kleing@13999
   308
next
kleing@13999
   309
  from AB show "?A" ..
kleing@13999
   310
qed
kleing@13999
   311
text{*\noindent Note the difference between @{text ?AB}, a term, and
kleing@13999
   312
@{text AB}, a fact.
kleing@13999
   313
kleing@13999
   314
Finally we want to start the proof with $\land$-elimination so we
kleing@13999
   315
don't have to perform it twice, as above. Here is a slick way to
kleing@13999
   316
achieve this: *}
kleing@13999
   317
kleing@13999
   318
lemma assumes AB: "large_A \<and> large_B"
kleing@13999
   319
  shows "large_B \<and> large_A" (is "?B \<and> ?A")
kleing@13999
   320
using AB
kleing@13999
   321
proof
kleing@13999
   322
  assume "?A" "?B" show ?thesis ..
kleing@13999
   323
qed
kleing@13999
   324
text{*\noindent Command \isakeyword{using} can appear before a proof
kleing@13999
   325
and adds further facts to those piped into the proof. Here @{text AB}
kleing@13999
   326
is the only such fact and it triggers $\land$-elimination. Another
kleing@13999
   327
frequent idiom is as follows:
kleing@13999
   328
\begin{center}
kleing@13999
   329
\isakeyword{from} \emph{major-facts}~
kleing@13999
   330
\isakeyword{show} \emph{proposition}~
kleing@13999
   331
\isakeyword{using} \emph{minor-facts}~
kleing@13999
   332
\emph{proof}
kleing@13999
   333
\end{center}
kleing@13999
   334
kleing@13999
   335
Sometimes it is necessary to suppress the implicit application of rules in a
kleing@13999
   336
\isakeyword{proof}. For example \isakeyword{show}~@{prop"A \<or> B"} would
kleing@13999
   337
trigger $\lor$-introduction, requiring us to prove @{prop A}. A simple
kleing@13999
   338
``@{text"-"}'' prevents this \emph{faux pas}: *}
kleing@13999
   339
kleing@13999
   340
lemma assumes AB: "A \<or> B" shows "B \<or> A"
kleing@13999
   341
proof -
kleing@13999
   342
  from AB show ?thesis
kleing@13999
   343
  proof
kleing@13999
   344
    assume A show ?thesis ..
kleing@13999
   345
  next
kleing@13999
   346
    assume B show ?thesis ..
kleing@13999
   347
  qed
kleing@13999
   348
qed
kleing@13999
   349
nipkow@19840
   350
text{* Too many names can easily clutter a proof.  We already learned
nipkow@19840
   351
about @{text this} as a means of avoiding explicit names. Another
nipkow@19840
   352
handy device is to refer to a fact not by name but by contents: for
nipkow@19840
   353
example, writing @{text "`A \<or> B`"} (enclosing the formula in back quotes)
nipkow@19840
   354
refers to the fact @{text"A \<or> B"}
nipkow@19840
   355
without the need to name it. Here is a simple example, a revised version
nipkow@19840
   356
of the previous proof *}
nipkow@19840
   357
nipkow@19840
   358
lemma assumes "A \<or> B" shows "B \<or> A"
nipkow@19840
   359
proof -
nipkow@19840
   360
  from `A \<or> B` show ?thesis
nipkow@19840
   361
(*<*)oops(*>*)
nipkow@19840
   362
text{*\noindent which continues as before.
nipkow@19840
   363
nipkow@19840
   364
Clearly, this device of quoting facts by contents is only advisable
nipkow@19840
   365
for small formulae. In such cases it is superior to naming because the
nipkow@19840
   366
reader immediately sees what the fact is without needing to search for
nipkow@19840
   367
it in the preceding proof text. *}
kleing@13999
   368
kleing@13999
   369
subsection{*Predicate calculus*}
kleing@13999
   370
kleing@13999
   371
text{* Command \isakeyword{fix} introduces new local variables into a
kleing@13999
   372
proof. The pair \isakeyword{fix}-\isakeyword{show} corresponds to @{text"\<And>"}
kleing@13999
   373
(the universal quantifier at the
kleing@13999
   374
meta-level) just like \isakeyword{assume}-\isakeyword{show} corresponds to
kleing@13999
   375
@{text"\<Longrightarrow>"}. Here is a sample proof, annotated with the rules that are
kleing@13999
   376
applied implicitly: *}
kleing@13999
   377
lemma assumes P: "\<forall>x. P x" shows "\<forall>x. P(f x)"
kleing@13999
   378
proof                       --"@{thm[source]allI}: @{thm allI}"
kleing@13999
   379
  fix a
kleing@13999
   380
  from P show "P(f a)" ..  --"@{thm[source]allE}: @{thm allE}"
kleing@13999
   381
qed
kleing@13999
   382
text{*\noindent Note that in the proof we have chosen to call the bound
kleing@13999
   383
variable @{term a} instead of @{term x} merely to show that the choice of
kleing@13999
   384
local names is irrelevant.
kleing@13999
   385
kleing@13999
   386
Next we look at @{text"\<exists>"} which is a bit more tricky.
kleing@13999
   387
*}
kleing@13999
   388
kleing@13999
   389
lemma assumes Pf: "\<exists>x. P(f x)" shows "\<exists>y. P y"
kleing@13999
   390
proof -
kleing@13999
   391
  from Pf show ?thesis
kleing@13999
   392
  proof              -- "@{thm[source]exE}: @{thm exE}"
kleing@13999
   393
    fix x
kleing@13999
   394
    assume "P(f x)"
kleing@13999
   395
    show ?thesis ..  -- "@{thm[source]exI}: @{thm exI}"
kleing@13999
   396
  qed
kleing@13999
   397
qed
kleing@13999
   398
text{*\noindent Explicit $\exists$-elimination as seen above can become
kleing@13999
   399
cumbersome in practice.  The derived Isar language element
kleing@13999
   400
\isakeyword{obtain} provides a more appealing form of generalised
kleing@13999
   401
existence reasoning: *}
kleing@13999
   402
kleing@13999
   403
lemma assumes Pf: "\<exists>x. P(f x)" shows "\<exists>y. P y"
kleing@13999
   404
proof -
kleing@13999
   405
  from Pf obtain x where "P(f x)" ..
kleing@13999
   406
  thus "\<exists>y. P y" ..
kleing@13999
   407
qed
kleing@13999
   408
text{*\noindent Note how the proof text follows the usual mathematical style
kleing@13999
   409
of concluding $P(x)$ from $\exists x. P(x)$, while carefully introducing $x$
kleing@13999
   410
as a new local variable.  Technically, \isakeyword{obtain} is similar to
kleing@13999
   411
\isakeyword{fix} and \isakeyword{assume} together with a soundness proof of
kleing@13999
   412
the elimination involved.
kleing@13999
   413
kleing@13999
   414
Here is a proof of a well known tautology.
kleing@13999
   415
Which rule is used where?  *}
kleing@13999
   416
kleing@13999
   417
lemma assumes ex: "\<exists>x. \<forall>y. P x y" shows "\<forall>y. \<exists>x. P x y"
kleing@13999
   418
proof
kleing@13999
   419
  fix y
kleing@13999
   420
  from ex obtain x where "\<forall>y. P x y" ..
kleing@13999
   421
  hence "P x y" ..
kleing@13999
   422
  thus "\<exists>x. P x y" ..
kleing@13999
   423
qed
kleing@13999
   424
kleing@13999
   425
subsection{*Making bigger steps*}
kleing@13999
   426
kleing@13999
   427
text{* So far we have confined ourselves to single step proofs. Of course
kleing@13999
   428
powerful automatic methods can be used just as well. Here is an example,
kleing@13999
   429
Cantor's theorem that there is no surjective function from a set to its
kleing@13999
   430
powerset: *}
kleing@13999
   431
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
kleing@13999
   432
proof
kleing@13999
   433
  let ?S = "{x. x \<notin> f x}"
kleing@13999
   434
  show "?S \<notin> range f"
kleing@13999
   435
  proof
kleing@13999
   436
    assume "?S \<in> range f"
nipkow@19840
   437
    then obtain y where "?S = f y" ..
kleing@13999
   438
    show False
kleing@13999
   439
    proof cases
kleing@13999
   440
      assume "y \<in> ?S"
nipkow@19840
   441
      with `?S = f y` show False by blast
kleing@13999
   442
    next
kleing@13999
   443
      assume "y \<notin> ?S"
nipkow@19840
   444
      with `?S = f y` show False by blast
kleing@13999
   445
    qed
kleing@13999
   446
  qed
kleing@13999
   447
qed
kleing@13999
   448
text{*\noindent
kleing@13999
   449
For a start, the example demonstrates two new constructs:
kleing@13999
   450
\begin{itemize}
kleing@13999
   451
\item \isakeyword{let} introduces an abbreviation for a term, in our case
kleing@13999
   452
the witness for the claim.
kleing@13999
   453
\item Proof by @{text"cases"} starts a proof by cases. Note that it remains
kleing@13999
   454
implicit what the two cases are: it is merely expected that the two subproofs
kleing@13999
   455
prove @{text"P \<Longrightarrow> ?thesis"} and @{text"\<not>P \<Longrightarrow> ?thesis"} (in that order)
kleing@13999
   456
for some @{term P}.
kleing@13999
   457
\end{itemize}
kleing@13999
   458
If you wonder how to \isakeyword{obtain} @{term y}:
kleing@13999
   459
via the predefined elimination rule @{thm rangeE[no_vars]}.
kleing@13999
   460
kleing@13999
   461
Method @{text blast} is used because the contradiction does not follow easily
kleing@13999
   462
by just a single rule. If you find the proof too cryptic for human
kleing@13999
   463
consumption, here is a more detailed version; the beginning up to
kleing@13999
   464
\isakeyword{obtain} stays unchanged. *}
kleing@13999
   465
kleing@13999
   466
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
kleing@13999
   467
proof
kleing@13999
   468
  let ?S = "{x. x \<notin> f x}"
kleing@13999
   469
  show "?S \<notin> range f"
kleing@13999
   470
  proof
kleing@13999
   471
    assume "?S \<in> range f"
nipkow@19840
   472
    then obtain y where "?S = f y" ..
kleing@13999
   473
    show False
kleing@13999
   474
    proof cases
kleing@13999
   475
      assume "y \<in> ?S"
kleing@13999
   476
      hence "y \<notin> f y"   by simp
nipkow@19840
   477
      hence "y \<notin> ?S"    by(simp add: `?S = f y`)
kleing@13999
   478
      thus False         by contradiction
kleing@13999
   479
    next
kleing@13999
   480
      assume "y \<notin> ?S"
kleing@13999
   481
      hence "y \<in> f y"   by simp
nipkow@19840
   482
      hence "y \<in> ?S"    by(simp add: `?S = f y`)
kleing@13999
   483
      thus False         by contradiction
kleing@13999
   484
    qed
kleing@13999
   485
  qed
kleing@13999
   486
qed
kleing@13999
   487
text{*\noindent Method @{text contradiction} succeeds if both $P$ and
kleing@13999
   488
$\neg P$ are among the assumptions and the facts fed into that step, in any order.
kleing@13999
   489
kleing@13999
   490
As it happens, Cantor's theorem can be proved automatically by best-first
kleing@13999
   491
search. Depth-first search would diverge, but best-first search successfully
kleing@13999
   492
navigates through the large search space:
kleing@13999
   493
*}
kleing@13999
   494
kleing@13999
   495
theorem "\<exists>S. S \<notin> range (f :: 'a \<Rightarrow> 'a set)"
kleing@13999
   496
by best
kleing@13999
   497
(* Of course this only works in the context of HOL's carefully
kleing@13999
   498
constructed introduction and elimination rules for set theory.*)
kleing@13999
   499
kleing@13999
   500
subsection{*Raw proof blocks*}
kleing@13999
   501
kleing@13999
   502
text{* Although we have shown how to employ powerful automatic methods like
kleing@13999
   503
@{text blast} to achieve bigger proof steps, there may still be the
kleing@13999
   504
tendency to use the default introduction and elimination rules to
kleing@13999
   505
decompose goals and facts. This can lead to very tedious proofs:
kleing@13999
   506
*}
kleing@13999
   507
(*<*)ML"set quick_and_dirty"(*>*)
kleing@13999
   508
lemma "\<forall>x y. A x y \<and> B x y \<longrightarrow> C x y"
kleing@13999
   509
proof
kleing@13999
   510
  fix x show "\<forall>y. A x y \<and> B x y \<longrightarrow> C x y"
kleing@13999
   511
  proof
kleing@13999
   512
    fix y show "A x y \<and> B x y \<longrightarrow> C x y"
kleing@13999
   513
    proof
kleing@13999
   514
      assume "A x y \<and> B x y"
kleing@13999
   515
      show "C x y" sorry
kleing@13999
   516
    qed
kleing@13999
   517
  qed
kleing@13999
   518
qed
kleing@13999
   519
text{*\noindent Since we are only interested in the decomposition and not the
kleing@13999
   520
actual proof, the latter has been replaced by
kleing@13999
   521
\isakeyword{sorry}. Command \isakeyword{sorry} proves anything but is
kleing@13999
   522
only allowed in quick and dirty mode, the default interactive mode. It
kleing@13999
   523
is very convenient for top down proof development.
kleing@13999
   524
kleing@13999
   525
Luckily we can avoid this step by step decomposition very easily: *}
kleing@13999
   526
kleing@13999
   527
lemma "\<forall>x y. A x y \<and> B x y \<longrightarrow> C x y"
kleing@13999
   528
proof -
kleing@13999
   529
  have "\<And>x y. \<lbrakk> A x y; B x y \<rbrakk> \<Longrightarrow> C x y"
kleing@13999
   530
  proof -
kleing@13999
   531
    fix x y assume "A x y" "B x y"
kleing@13999
   532
    show "C x y" sorry
kleing@13999
   533
  qed
kleing@13999
   534
  thus ?thesis by blast
kleing@13999
   535
qed
kleing@13999
   536
kleing@13999
   537
text{*\noindent
kleing@13999
   538
This can be simplified further by \emph{raw proof blocks}, i.e.\
kleing@13999
   539
proofs enclosed in braces: *}
kleing@13999
   540
kleing@13999
   541
lemma "\<forall>x y. A x y \<and> B x y \<longrightarrow> C x y"
kleing@13999
   542
proof -
kleing@13999
   543
  { fix x y assume "A x y" "B x y"
kleing@13999
   544
    have "C x y" sorry }
kleing@13999
   545
  thus ?thesis by blast
kleing@13999
   546
qed
kleing@13999
   547
kleing@13999
   548
text{*\noindent The result of the raw proof block is the same theorem
kleing@13999
   549
as above, namely @{prop"\<And>x y. \<lbrakk> A x y; B x y \<rbrakk> \<Longrightarrow> C x y"}.  Raw
kleing@13999
   550
proof blocks are like ordinary proofs except that they do not prove
kleing@13999
   551
some explicitly stated property but that the property emerges directly
kleing@13999
   552
out of the \isakeyword{fixe}s, \isakeyword{assume}s and
kleing@13999
   553
\isakeyword{have} in the block. Thus they again serve to avoid
kleing@13999
   554
duplication. Note that the conclusion of a raw proof block is stated with
kleing@13999
   555
\isakeyword{have} rather than \isakeyword{show} because it is not the
kleing@13999
   556
conclusion of some pending goal but some independent claim.
kleing@13999
   557
kleing@13999
   558
The general idea demonstrated in this subsection is very
kleing@13999
   559
important in Isar and distinguishes it from tactic-style proofs:
kleing@13999
   560
\begin{quote}\em
kleing@13999
   561
Do not manipulate the proof state into a particular form by applying
kleing@13999
   562
tactics but state the desired form explicitly and let the tactic verify
kleing@13999
   563
that from this form the original goal follows.
kleing@13999
   564
\end{quote}
nipkow@14617
   565
This yields more readable and also more robust proofs.
nipkow@14617
   566
nipkow@14617
   567
\subsubsection{General case distinctions}
nipkow@14617
   568
nipkow@14617
   569
As an important application of raw proof blocks we show how to deal
nipkow@14617
   570
with general case distinctions --- more specific kinds are treated in
nipkow@14617
   571
\S\ref{sec:CaseDistinction}. Imagine that you would like to prove some
nipkow@14617
   572
goal by distinguishing $n$ cases $P_1$, \dots, $P_n$. You show that
nipkow@14617
   573
the $n$ cases are exhaustive (i.e.\ $P_1 \lor \dots \lor P_n$) and
nipkow@14617
   574
that each case $P_i$ implies the goal. Taken together, this proves the
nipkow@14617
   575
goal. The corresponding Isar proof pattern (for $n = 3$) is very handy:
nipkow@14617
   576
*}
nipkow@14617
   577
text_raw{*\renewcommand{\isamarkupcmt}[1]{#1}*}
nipkow@14617
   578
(*<*)lemma "XXX"(*>*)
nipkow@14617
   579
proof -
nipkow@14617
   580
  have "P\<^isub>1 \<or> P\<^isub>2 \<or> P\<^isub>3" (*<*)sorry(*>*) -- {*\dots*}
nipkow@14617
   581
  moreover
nipkow@14617
   582
  { assume P\<^isub>1
nipkow@14617
   583
    -- {*\dots*}
nipkow@14617
   584
    have ?thesis (*<*)sorry(*>*) -- {*\dots*} }
nipkow@14617
   585
  moreover
nipkow@14617
   586
  { assume P\<^isub>2
nipkow@14617
   587
    -- {*\dots*}
nipkow@14617
   588
    have ?thesis (*<*)sorry(*>*) -- {*\dots*} }
nipkow@14617
   589
  moreover
nipkow@14617
   590
  { assume P\<^isub>3
nipkow@14617
   591
    -- {*\dots*}
nipkow@14617
   592
    have ?thesis (*<*)sorry(*>*) -- {*\dots*} }
nipkow@14617
   593
  ultimately show ?thesis by blast
nipkow@14617
   594
qed
nipkow@14617
   595
text_raw{*\renewcommand{\isamarkupcmt}[1]{{\isastylecmt--- #1}}*}
nipkow@14617
   596
kleing@13999
   597
kleing@13999
   598
subsection{*Further refinements*}
kleing@13999
   599
kleing@13999
   600
text{* This subsection discusses some further tricks that can make
kleing@13999
   601
life easier although they are not essential. *}
kleing@13999
   602
kleing@13999
   603
subsubsection{*\isakeyword{and}*}
kleing@13999
   604
kleing@13999
   605
text{* Propositions (following \isakeyword{assume} etc) may but need not be
kleing@13999
   606
separated by \isakeyword{and}. This is not just for readability
kleing@13999
   607
(\isakeyword{from} \isa{A} \isakeyword{and} \isa{B} looks nicer than
kleing@13999
   608
\isakeyword{from} \isa{A} \isa{B}) but for structuring lists of propositions
kleing@13999
   609
into possibly named blocks. In
kleing@13999
   610
\begin{center}
kleing@13999
   611
\isakeyword{assume} \isa{A:} $A_1$ $A_2$ \isakeyword{and} \isa{B:} $A_3$
kleing@13999
   612
\isakeyword{and} $A_4$
kleing@13999
   613
\end{center}
kleing@13999
   614
label \isa{A} refers to the list of propositions $A_1$ $A_2$ and
kleing@13999
   615
label \isa{B} to $A_3$. *}
kleing@13999
   616
kleing@13999
   617
subsubsection{*\isakeyword{note}*}
kleing@13999
   618
text{* If you want to remember intermediate fact(s) that cannot be
kleing@13999
   619
named directly, use \isakeyword{note}. For example the result of raw
kleing@13999
   620
proof block can be named by following it with
kleing@13999
   621
\isakeyword{note}~@{text"some_name = this"}.  As a side effect,
kleing@13999
   622
@{text this} is set to the list of facts on the right-hand side. You
kleing@13999
   623
can also say @{text"note some_fact"}, which simply sets @{text this},
kleing@13999
   624
i.e.\ recalls @{text"some_fact"}, e.g.\ in a \isakeyword{moreover} sequence. *}
kleing@13999
   625
kleing@13999
   626
kleing@13999
   627
subsubsection{*\isakeyword{fixes}*}
kleing@13999
   628
kleing@13999
   629
text{* Sometimes it is necessary to decorate a proposition with type
kleing@13999
   630
constraints, as in Cantor's theorem above. These type constraints tend
kleing@13999
   631
to make the theorem less readable. The situation can be improved a
kleing@13999
   632
little by combining the type constraint with an outer @{text"\<And>"}: *}
kleing@13999
   633
kleing@13999
   634
theorem "\<And>f :: 'a \<Rightarrow> 'a set. \<exists>S. S \<notin> range f"
kleing@13999
   635
(*<*)oops(*>*)
kleing@13999
   636
text{*\noindent However, now @{term f} is bound and we need a
kleing@13999
   637
\isakeyword{fix}~\isa{f} in the proof before we can refer to @{term f}.
kleing@13999
   638
This is avoided by \isakeyword{fixes}: *}
kleing@13999
   639
kleing@13999
   640
theorem fixes f :: "'a \<Rightarrow> 'a set" shows "\<exists>S. S \<notin> range f"
kleing@13999
   641
(*<*)oops(*>*)
kleing@13999
   642
text{* \noindent
kleing@13999
   643
Even better, \isakeyword{fixes} allows to introduce concrete syntax locally:*}
kleing@13999
   644
kleing@13999
   645
lemma comm_mono:
kleing@13999
   646
  fixes r :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix ">" 60) and
kleing@13999
   647
       f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"   (infixl "++" 70)
kleing@13999
   648
  assumes comm: "\<And>x y::'a. x ++ y = y ++ x" and
kleing@13999
   649
          mono: "\<And>x y z::'a. x > y \<Longrightarrow> x ++ z > y ++ z"
kleing@13999
   650
  shows "x > y \<Longrightarrow> z ++ x > z ++ y"
kleing@13999
   651
by(simp add: comm mono)
kleing@13999
   652
kleing@13999
   653
text{*\noindent The concrete syntax is dropped at the end of the proof and the
kleing@13999
   654
theorem becomes @{thm[display,margin=60]comm_mono}
kleing@13999
   655
\tweakskip *}
kleing@13999
   656
kleing@13999
   657
subsubsection{*\isakeyword{obtain}*}
kleing@13999
   658
kleing@13999
   659
text{* The \isakeyword{obtain} construct can introduce multiple
kleing@13999
   660
witnesses and propositions as in the following proof fragment:*}
kleing@13999
   661
kleing@13999
   662
lemma assumes A: "\<exists>x y. P x y \<and> Q x y" shows "R"
kleing@13999
   663
proof -
kleing@13999
   664
  from A obtain x y where P: "P x y" and Q: "Q x y"  by blast
kleing@13999
   665
(*<*)oops(*>*)
kleing@13999
   666
text{* Remember also that one does not even need to start with a formula
kleing@13999
   667
containing @{text"\<exists>"} as we saw in the proof of Cantor's theorem.
kleing@13999
   668
*}
kleing@13999
   669
kleing@13999
   670
subsubsection{*Combining proof styles*}
kleing@13999
   671
kleing@13999
   672
text{* Finally, whole ``scripts'' (tactic-based proofs in the style of
kleing@13999
   673
\cite{LNCS2283}) may appear in the leaves of the proof tree, although this is
kleing@13999
   674
best avoided.  Here is a contrived example: *}
kleing@13999
   675
kleing@13999
   676
lemma "A \<longrightarrow> (A \<longrightarrow> B) \<longrightarrow> B"
kleing@13999
   677
proof
kleing@13999
   678
  assume a: "A"
kleing@13999
   679
  show "(A \<longrightarrow>B) \<longrightarrow> B"
kleing@13999
   680
    apply(rule impI)
kleing@13999
   681
    apply(erule impE)
kleing@13999
   682
    apply(rule a)
kleing@13999
   683
    apply assumption
kleing@13999
   684
    done
kleing@13999
   685
qed
kleing@13999
   686
kleing@13999
   687
text{*\noindent You may need to resort to this technique if an
kleing@13999
   688
automatic step fails to prove the desired proposition.
kleing@13999
   689
kleing@13999
   690
When converting a proof from tactic-style into Isar you can proceed
kleing@13999
   691
in a top-down manner: parts of the proof can be left in script form
kleing@13999
   692
while the outer structure is already expressed in Isar. *}
kleing@13999
   693
kleing@13999
   694
(*<*)end(*>*)