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