src/Doc/Prog_Prove/Isar.thy
author wenzelm
Thu, 03 Mar 2022 17:13:24 +0100
changeset 75203 ee1bd0687c2b
parent 73511 2cdbb6a2f2a7
child 76987 4c275405faae
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     1
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     2
theory Isar
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     3
imports LaTeXsugar
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     4
begin
52059
2f970c7f722b proper option quick_and_dirty;
wenzelm
parents: 51445
diff changeset
     5
declare [[quick_and_dirty]]
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     6
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
     7
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     8
Apply-scripts are unreadable and hard to maintain. The language of choice
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     9
for larger proofs is \concept{Isar}. The two key features of Isar are:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    10
\begin{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    11
\item It is structured, not linear.
56989
nipkow
parents: 56451
diff changeset
    12
\item It is readable without its being run because
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    13
you need to state what you are proving at any given point.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    14
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    15
Whereas apply-scripts are like assembly language programs, Isar proofs
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    16
are like structured programs with comments. A typical Isar proof looks like this:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
    17
\<close>text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    18
\begin{tabular}{@ {}l}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    19
\isacom{proof}\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    20
\quad\isacom{assume} \<open>"\<close>$\mathit{formula}_0$\<open>"\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    21
\quad\isacom{have} \<open>"\<close>$\mathit{formula}_1$\<open>"\<close> \quad\isacom{by} \<open>simp\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    22
\quad\vdots\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    23
\quad\isacom{have} \<open>"\<close>$\mathit{formula}_n$\<open>"\<close> \quad\isacom{by} \<open>blast\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    24
\quad\isacom{show} \<open>"\<close>$\mathit{formula}_{n+1}$\<open>"\<close> \quad\isacom{by} \<open>\<dots>\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    25
\isacom{qed}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    26
\end{tabular}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
    27
\<close>text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    28
It proves $\mathit{formula}_0 \Longrightarrow \mathit{formula}_{n+1}$
47704
8b4cd98f944e doc update
nipkow
parents: 47306
diff changeset
    29
(provided each proof step succeeds).
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    30
The intermediate \isacom{have} statements are merely stepping stones
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    31
on the way towards the \isacom{show} statement that proves the actual
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    32
goal. In more detail, this is the Isar core syntax:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    33
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    34
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    35
\begin{tabular}{@ {}lcl@ {}}
55359
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
    36
\textit{proof} &=& \indexed{\isacom{by}}{by} \textit{method}\\
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
    37
      &$\mid$& \indexed{\isacom{proof}}{proof} [\textit{method}] \ \textit{step}$^*$ \ \indexed{\isacom{qed}}{qed}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    38
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    39
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    40
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    41
\begin{tabular}{@ {}lcl@ {}}
55359
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
    42
\textit{step} &=& \indexed{\isacom{fix}}{fix} \textit{variables} \\
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
    43
      &$\mid$& \indexed{\isacom{assume}}{assume} \textit{proposition} \\
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
    44
      &$\mid$& [\indexed{\isacom{from}}{from} \textit{fact}$^+$] (\indexed{\isacom{have}}{have} $\mid$ \indexed{\isacom{show}}{show}) \ \textit{proposition} \ \textit{proof}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    45
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    46
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    47
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    48
\begin{tabular}{@ {}lcl@ {}}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    49
\textit{proposition} &=& [\textit{name}:] \<open>"\<close>\textit{formula}\<open>"\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    50
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    51
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    52
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    53
\begin{tabular}{@ {}lcl@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    54
\textit{fact} &=& \textit{name} \ $\mid$ \ \dots
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    55
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    56
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    57
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    58
\noindent A proof can either be an atomic \isacom{by} with a single proof
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    59
method which must finish off the statement being proved, for example \<open>auto\<close>,  or it can be a \isacom{proof}--\isacom{qed} block of multiple
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    60
steps. Such a block can optionally begin with a proof method that indicates
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    61
how to start off the proof, e.g., \mbox{\<open>(induction xs)\<close>}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    62
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    63
A step either assumes a proposition or states a proposition
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    64
together with its proof. The optional \isacom{from} clause
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    65
indicates which facts are to be used in the proof.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    66
Intermediate propositions are stated with \isacom{have}, the overall goal
56989
nipkow
parents: 56451
diff changeset
    67
is stated with \isacom{show}. A step can also introduce new local variables with
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    68
\isacom{fix}. Logically, \isacom{fix} introduces \<open>\<And>\<close>-quantified
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    69
variables, \isacom{assume} introduces the assumption of an implication
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    70
(\<open>\<Longrightarrow>\<close>) and \isacom{have}/\isacom{show} introduce the conclusion.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    71
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    72
Propositions are optionally named formulas. These names can be referred to in
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    73
later \isacom{from} clauses. In the simplest case, a fact is such a name.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    74
But facts can also be composed with \<open>OF\<close> and \<open>of\<close> as shown in
58602
ab56811d76c6 tuned spaces
nipkow
parents: 58522
diff changeset
    75
\autoref{sec:forward-proof} --- hence the \dots\ in the above grammar.  Note
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    76
that assumptions, intermediate \isacom{have} statements and global lemmas all
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    77
have the same status and are thus collectively referred to as
55317
834a84553e02 started index
nipkow
parents: 54839
diff changeset
    78
\conceptidx{facts}{fact}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    79
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    80
Fact names can stand for whole lists of facts. For example, if \<open>f\<close> is
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    81
defined by command \isacom{fun}, \<open>f.simps\<close> refers to the whole list of
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    82
recursion equations defining \<open>f\<close>. Individual facts can be selected by
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
    83
writing \<open>f.simps(2)\<close>, whole sublists by writing \<open>f.simps(2-4)\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    84
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    85
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
    86
\section{Isar by Example}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    87
47704
8b4cd98f944e doc update
nipkow
parents: 47306
diff changeset
    88
We show a number of proofs of Cantor's theorem that a function from a set to
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    89
its powerset cannot be surjective, illustrating various features of Isar. The
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    90
constant \<^const>\<open>surj\<close> is predefined.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
    91
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    92
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    93
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    94
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    95
  assume 0: "surj f"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    96
  from 0 have 1: "\<forall>A. \<exists>a. A = f a" by(simp add: surj_def)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    97
  from 1 have 2: "\<exists>a. {x. x \<notin> f x} = f a" by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    98
  from 2 show "False" by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    99
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   100
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   101
text\<open>
56989
nipkow
parents: 56451
diff changeset
   102
The \isacom{proof} command lacks an explicit method by which to perform
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   103
the proof. In such cases Isabelle tries to use some standard introduction
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   104
rule, in the above case for \<open>\<not>\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   105
\[
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   106
\inferrule{
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   107
\mbox{@{thm (prem 1) notI}}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   108
{\mbox{@{thm (concl) notI}}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   109
\]
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   110
In order to prove \<^prop>\<open>~ P\<close>, assume \<open>P\<close> and show \<open>False\<close>.
64852
f3504bc69ea3 fix problems because of "surj" input abbreviation; tuned
nipkow
parents: 61517
diff changeset
   111
Thus we may assume \mbox{\noquotes{@{prop [source] "surj f"}}}. The proof shows that names of propositions
58602
ab56811d76c6 tuned spaces
nipkow
parents: 58522
diff changeset
   112
may be (single!) digits --- meaningful names are hard to invent and are often
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   113
not necessary. Both \isacom{have} steps are obvious. The second one introduces
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   114
the diagonal set \<^term>\<open>{x. x \<notin> f x}\<close>, the key idea in the proof.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   115
If you wonder why \<open>2\<close> directly implies \<open>False\<close>: from \<open>2\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   116
it follows that \<^prop>\<open>a \<notin> f a \<longleftrightarrow> a \<in> f a\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   117
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   118
\subsection{\indexed{\<open>this\<close>}{this}, \indexed{\isacom{then}}{then}, \indexed{\isacom{hence}}{hence} and \indexed{\isacom{thus}}{thus}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   119
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   120
Labels should be avoided. They interrupt the flow of the reader who has to
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   121
scan the context for the point where the label was introduced. Ideally, the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   122
proof is a linear flow, where the output of one step becomes the input of the
58605
nipkow
parents: 58602
diff changeset
   123
next step, piping the previously proved fact into the next proof, like
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   124
in a UNIX pipe. In such cases the predefined name \<open>this\<close> can be used
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   125
to refer to the proposition proved in the previous step. This allows us to
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   126
eliminate all labels from our proof (we suppress the \isacom{lemma} statement):
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   127
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   128
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   129
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   130
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   131
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   132
  assume "surj f"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   133
  from this have "\<exists>a. {x. x \<notin> f x} = f a" by(auto simp: surj_def)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   134
  from this show "False" by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   135
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   136
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   137
text\<open>We have also taken the opportunity to compress the two \isacom{have}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   138
steps into one.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   139
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   140
To compact the text further, Isar has a few convenient abbreviations:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   141
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   142
54839
nipkow
parents: 54577
diff changeset
   143
\begin{tabular}{r@ {\quad=\quad}l}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   144
\isacom{then} & \isacom{from} \<open>this\<close>\\
54839
nipkow
parents: 54577
diff changeset
   145
\isacom{thus} & \isacom{then} \isacom{show}\\
nipkow
parents: 54577
diff changeset
   146
\isacom{hence} & \isacom{then} \isacom{have}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   147
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   148
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   149
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   150
\noindent
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   151
With the help of these abbreviations the proof becomes
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   152
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   153
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   154
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   155
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   156
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   157
  assume "surj f"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   158
  hence "\<exists>a. {x. x \<notin> f x} = f a" by(auto simp: surj_def)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   159
  thus "False" by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   160
qed
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   161
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   162
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   163
There are two further linguistic variations:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   164
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   165
54839
nipkow
parents: 54577
diff changeset
   166
\begin{tabular}{r@ {\quad=\quad}l}
55359
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
   167
(\isacom{have}$\mid$\isacom{show}) \ \textit{prop} \ \indexed{\isacom{using}}{using} \ \textit{facts}
54839
nipkow
parents: 54577
diff changeset
   168
&
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   169
\isacom{from} \ \textit{facts} \ (\isacom{have}$\mid$\isacom{show}) \ \textit{prop}\\
55359
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
   170
\indexed{\isacom{with}}{with} \ \textit{facts} & \isacom{from} \ \textit{facts} \isa{this}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   171
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   172
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   173
47711
c1cca2a052e4 doc update
nipkow
parents: 47704
diff changeset
   174
\noindent The \isacom{using} idiom de-emphasizes the used facts by moving them
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   175
behind the proposition.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   176
55359
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
   177
\subsection{Structured Lemma Statements: \indexed{\isacom{fixes}}{fixes}, \indexed{\isacom{assumes}}{assumes}, \indexed{\isacom{shows}}{shows}}
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
   178
\index{lemma@\isacom{lemma}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   179
Lemmas can also be stated in a more structured fashion. To demonstrate this
64852
f3504bc69ea3 fix problems because of "surj" input abbreviation; tuned
nipkow
parents: 61517
diff changeset
   180
feature with Cantor's theorem, we rephrase \noquotes{@{prop[source]"\<not> surj f"}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   181
a little:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   182
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   183
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   184
lemma
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   185
  fixes f :: "'a \<Rightarrow> 'a set"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   186
  assumes s: "surj f"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   187
  shows "False"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   188
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   189
txt\<open>The optional \isacom{fixes} part allows you to state the types of
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   190
variables up front rather than by decorating one of their occurrences in the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   191
formula with a type constraint. The key advantage of the structured format is
47306
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47269
diff changeset
   192
the \isacom{assumes} part that allows you to name each assumption; multiple
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47269
diff changeset
   193
assumptions can be separated by \isacom{and}. The
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   194
\isacom{shows} part gives the goal. The actual theorem that will come out of
64852
f3504bc69ea3 fix problems because of "surj" input abbreviation; tuned
nipkow
parents: 61517
diff changeset
   195
the proof is \noquotes{@{prop[source]"surj f \<Longrightarrow> False"}}, but during the proof the assumption
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   196
\noquotes{@{prop[source]"surj f"}} is available under the name \<open>s\<close> like any other fact.
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   197
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   198
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   199
proof -
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   200
  have "\<exists> a. {x. x \<notin> f x} = f a" using s
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   201
    by(auto simp: surj_def)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   202
  thus "False" by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   203
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   204
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   205
text\<open>
56312
nipkow
parents: 55417
diff changeset
   206
\begin{warn}
nipkow
parents: 55417
diff changeset
   207
Note the hyphen after the \isacom{proof} command.
56989
nipkow
parents: 56451
diff changeset
   208
It is the null method that does nothing to the goal. Leaving it out would be asking
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   209
Isabelle to try some suitable introduction rule on the goal \<^const>\<open>False\<close> --- but
56312
nipkow
parents: 55417
diff changeset
   210
there is no such rule and \isacom{proof} would fail.
nipkow
parents: 55417
diff changeset
   211
\end{warn}
64852
f3504bc69ea3 fix problems because of "surj" input abbreviation; tuned
nipkow
parents: 61517
diff changeset
   212
In the \isacom{have} step the assumption \noquotes{@{prop[source]"surj f"}} is now
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   213
referenced by its name \<open>s\<close>. The duplication of \noquotes{@{prop[source]"surj f"}} in the
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   214
above proofs (once in the statement of the lemma, once in its proof) has been
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   215
eliminated.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   216
47704
8b4cd98f944e doc update
nipkow
parents: 47306
diff changeset
   217
Stating a lemma with \isacom{assumes}-\isacom{shows} implicitly introduces the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   218
name \indexed{\<open>assms\<close>}{assms} that stands for the list of all assumptions. You can refer
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   219
to individual assumptions by \<open>assms(1)\<close>, \<open>assms(2)\<close>, etc.,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   220
thus obviating the need to name them individually.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   221
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
   222
\section{Proof Patterns}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   223
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   224
We show a number of important basic proof patterns. Many of them arise from
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   225
the rules of natural deduction that are applied by \isacom{proof} by
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   226
default. The patterns are phrased in terms of \isacom{show} but work for
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   227
\isacom{have} and \isacom{lemma}, too.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   228
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   229
\ifsem\else
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   230
\subsection{Logic}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   231
\fi
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   232
47711
c1cca2a052e4 doc update
nipkow
parents: 47704
diff changeset
   233
We start with two forms of \concept{case analysis}:
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   234
starting from a formula \<open>P\<close> we have the two cases \<open>P\<close> and
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   235
\<^prop>\<open>~P\<close>, and starting from a fact \<^prop>\<open>P \<or> Q\<close>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   236
we have the two cases \<open>P\<close> and \<open>Q\<close>:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   237
\<close>text_raw\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   238
\begin{tabular}{@ {}ll@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   239
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   240
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   241
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   242
(*<*)lemma "R" proof-(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   243
show "R"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   244
proof cases
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   245
  assume "P"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   246
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   247
  show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   248
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   249
  assume "\<not> P"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   250
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   251
  show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   252
qed(*<*)oops(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   253
text_raw \<open>}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   254
\end{minipage}\index{cases@\<open>cases\<close>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   255
&
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   256
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   257
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   258
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   259
(*<*)lemma "R" proof-(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   260
have "P \<or> Q" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   261
then show "R"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   262
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   263
  assume "P"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   264
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   265
  show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   266
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   267
  assume "Q"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   268
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   269
  show "R" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   270
qed(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   271
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   272
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   273
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   274
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   275
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   276
\begin{isamarkuptext}%
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   277
How to prove a logical equivalence:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   278
\end{isamarkuptext}%
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   279
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   280
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   281
(*<*)lemma "P\<longleftrightarrow>Q" proof-(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   282
show "P \<longleftrightarrow> Q"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   283
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   284
  assume "P"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   285
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   286
  show "Q" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   287
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   288
  assume "Q"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   289
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   290
  show "P" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   291
qed(*<*)qed(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   292
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   293
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   294
\begin{isamarkuptext}%
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   295
Proofs by contradiction (@{thm[source] ccontr} stands for ``classical contradiction''):
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   296
\end{isamarkuptext}%
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   297
\begin{tabular}{@ {}ll@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   298
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   299
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   300
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   301
(*<*)lemma "\<not> P" proof-(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   302
show "\<not> P"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   303
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   304
  assume "P"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   305
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   306
  show "False" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   307
qed(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   308
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   309
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   310
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   311
&
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   312
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   313
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   314
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   315
(*<*)lemma "P" proof-(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   316
show "P"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   317
proof (rule ccontr)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   318
  assume "\<not>P"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   319
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   320
  show "False" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   321
qed(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   322
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   323
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   324
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   325
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   326
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   327
\begin{isamarkuptext}%
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   328
How to prove quantified formulas:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   329
\end{isamarkuptext}%
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   330
\begin{tabular}{@ {}ll@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   331
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   332
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   333
\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   334
(*<*)lemma "\<forall>x. P x" proof-(*>*)
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   335
show "\<forall>x. P(x)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   336
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   337
  fix x
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   338
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   339
  show "P(x)" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   340
qed(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   341
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   342
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   343
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   344
&
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   345
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   346
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   347
\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   348
(*<*)lemma "\<exists>x. P(x)" proof-(*>*)
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   349
show "\<exists>x. P(x)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   350
proof
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   351
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   352
  show "P(witness)" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   353
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   354
(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   355
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   356
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   357
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   358
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   359
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   360
\begin{isamarkuptext}%
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   361
In the proof of \noquotes{@{prop[source]"\<forall>x. P(x)"}},
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   362
the step \indexed{\isacom{fix}}{fix}~\<open>x\<close> introduces a locally fixed variable \<open>x\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   363
into the subproof, the proverbial ``arbitrary but fixed value''.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   364
Instead of \<open>x\<close> we could have chosen any name in the subproof.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   365
In the proof of \noquotes{@{prop[source]"\<exists>x. P(x)"}},
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   366
\<open>witness\<close> is some arbitrary
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   367
term for which we can prove that it satisfies \<open>P\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   368
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   369
How to reason forward from \noquotes{@{prop[source] "\<exists>x. P(x)"}}:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   370
\end{isamarkuptext}%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   371
\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   372
(*<*)lemma True proof- assume 1: "\<exists>x. P x"(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   373
have "\<exists>x. P(x)" (*<*)by(rule 1)(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   374
then obtain x where p: "P(x)" by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   375
(*<*)oops(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   376
text\<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   377
After the \indexed{\isacom{obtain}}{obtain} step, \<open>x\<close> (we could have chosen any name)
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   378
is a fixed local
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   379
variable, and \<open>p\<close> is the name of the fact
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   380
\noquotes{@{prop[source] "P(x)"}}.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   381
This pattern works for one or more \<open>x\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   382
As an example of the \isacom{obtain} command, here is the proof of
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   383
Cantor's theorem in more detail:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   384
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   385
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   386
lemma "\<not> surj(f :: 'a \<Rightarrow> 'a set)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   387
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   388
  assume "surj f"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   389
  hence  "\<exists>a. {x. x \<notin> f x} = f a" by(auto simp: surj_def)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   390
  then obtain a where  "{x. x \<notin> f x} = f a"  by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   391
  hence  "a \<notin> f a \<longleftrightarrow> a \<in> f a"  by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   392
  thus "False" by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   393
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   394
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   395
text_raw\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   396
\begin{isamarkuptext}%
47306
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47269
diff changeset
   397
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47269
diff changeset
   398
Finally, how to prove set equality and subset relationship:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   399
\end{isamarkuptext}%
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   400
\begin{tabular}{@ {}ll@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   401
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   402
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   403
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   404
(*<*)lemma "A = (B::'a set)" proof-(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   405
show "A = B"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   406
proof
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   407
  show "A \<subseteq> B" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   408
next
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   409
  show "B \<subseteq> A" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   410
qed(*<*)qed(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   411
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   412
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   413
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   414
&
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   415
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   416
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   417
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   418
(*<*)lemma "A <= (B::'a set)" proof-(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   419
show "A \<subseteq> B"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   420
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   421
  fix x
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   422
  assume "x \<in> A"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   423
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   424
  show "x \<in> B" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   425
qed(*<*)qed(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   426
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   427
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   428
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   429
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   430
\begin{isamarkuptext}%
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   431
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   432
\ifsem\else
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   433
\subsection{Chains of (In)Equations}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   434
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   435
In textbooks, chains of equations (and inequations) are often displayed like this:
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   436
\begin{quote}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   437
\begin{tabular}{@ {}l@ {\qquad}l@ {}}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   438
$t_1 = t_2$ & \isamath{\,\langle\mathit{justification}\rangle}\\
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   439
$\phantom{t_1} = t_3$ & \isamath{\,\langle\mathit{justification}\rangle}\\
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   440
\quad $\vdots$\\
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   441
$\phantom{t_1} = t_n$ & \isamath{\,\langle\mathit{justification}\rangle}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   442
\end{tabular}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   443
\end{quote}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   444
The Isar equivalent is this:
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   445
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   446
\begin{samepage}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   447
\begin{quote}
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   448
\isacom{have} \<open>"t\<^sub>1 = t\<^sub>2"\<close> \isasymproof\\
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   449
\isacom{also have} \<open>"... = t\<^sub>3"\<close> \isasymproof\\
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   450
\quad $\vdots$\\
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   451
\isacom{also have} \<open>"... = t\<^sub>n"\<close> \isasymproof \\
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   452
\isacom{finally show} \<open>"t\<^sub>1 = t\<^sub>n"\<close>\ \texttt{.}
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   453
\end{quote}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   454
\end{samepage}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   455
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   456
\noindent
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   457
The ``\<open>...\<close>'' and ``\<open>.\<close>'' deserve some explanation:
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   458
\begin{description}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   459
\item[``\<open>...\<close>''] is literally three dots. It is the name of an unknown that Isar
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   460
automatically instantiates with the right-hand side of the previous equation.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   461
In general, if \<open>this\<close> is the theorem \<^term>\<open>p t\<^sub>1 t\<^sub>2\<close> then ``\<open>...\<close>''
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   462
stands for \<open>t\<^sub>2\<close>.
65349
nipkow
parents: 65348
diff changeset
   463
\item[``\<open>.\<close>''] (a single dot) is a proof method that solves a goal by one of the
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   464
assumptions. This works here because the result of \isacom{finally}
65349
nipkow
parents: 65348
diff changeset
   465
is the theorem \mbox{\<open>t\<^sub>1 = t\<^sub>n\<close>},
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   466
\isacom{show} \<open>"t\<^sub>1 = t\<^sub>n"\<close> states the theorem explicitly,
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   467
and ``\<open>.\<close>'' proves the theorem with the result of \isacom{finally}.
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   468
\end{description}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   469
The above proof template also works for arbitrary mixtures of \<open>=\<close>, \<open>\<le>\<close> and \<open><\<close>,
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   470
for example:
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   471
\begin{quote}
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   472
\isacom{have} \<open>"t\<^sub>1 < t\<^sub>2"\<close> \isasymproof\\
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   473
\isacom{also have} \<open>"... = t\<^sub>3"\<close> \isasymproof\\
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   474
\quad $\vdots$\\
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   475
\isacom{also have} \<open>"... \<le> t\<^sub>n"\<close> \isasymproof \\
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   476
\isacom{finally show} \<open>"t\<^sub>1 < t\<^sub>n"\<close>\ \texttt{.}
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   477
\end{quote}
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   478
The relation symbol in the \isacom{finally} step needs to be the most precise one
65349
nipkow
parents: 65348
diff changeset
   479
possible. In the example above, you must not write \<open>t\<^sub>1 \<le> t\<^sub>n\<close> instead of \mbox{\<open>t\<^sub>1 < t\<^sub>n\<close>}.
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   480
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   481
\begin{warn}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   482
Isabelle only supports \<open>=\<close>, \<open>\<le>\<close> and \<open><\<close> but not \<open>\<ge>\<close> and \<open>>\<close>
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   483
in (in)equation chains (by default).
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   484
\end{warn}
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   485
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   486
If you want to go beyond merely using the above proof patterns and want to
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   487
understand what \isacom{also} and \isacom{finally} mean, read on.
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   488
There is an Isar theorem variable called \<open>calculation\<close>, similar to \<open>this\<close>.
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   489
When the first \isacom{also} in a chain is encountered, Isabelle sets
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   490
\<open>calculation := this\<close>. In each subsequent \isacom{also} step,
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   491
Isabelle composes the theorems \<open>calculation\<close> and \<open>this\<close> (i.e.\ the two previous
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   492
(in)equalities) using some predefined set of rules including transitivity
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   493
of \<open>=\<close>, \<open>\<le>\<close> and \<open><\<close> but also mixed rules like \<^prop>\<open>\<lbrakk> x \<le> y; y < z \<rbrakk> \<Longrightarrow> x < z\<close>.
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   494
The result of this composition is assigned to \<open>calculation\<close>. Consider
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   495
\begin{quote}
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   496
\isacom{have} \<open>"t\<^sub>1 \<le> t\<^sub>2"\<close> \isasymproof\\
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   497
\isacom{also} \isacom{have} \<open>"... < t\<^sub>3"\<close> \isasymproof\\
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   498
\isacom{also} \isacom{have} \<open>"... = t\<^sub>4"\<close> \isasymproof\\
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   499
\isacom{finally show} \<open>"t\<^sub>1 < t\<^sub>4"\<close>\ \texttt{.}
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   500
\end{quote}
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   501
After the first \isacom{also}, \<open>calculation\<close> is \<open>"t\<^sub>1 \<le> t\<^sub>2"\<close>,
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   502
and after the second \isacom{also}, \<open>calculation\<close> is \<open>"t\<^sub>1 < t\<^sub>3"\<close>.
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   503
The command \isacom{finally} is short for \isacom{also from} \<open>calculation\<close>.
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   504
Therefore the \isacom{also} hidden in \isacom{finally} sets \<open>calculation\<close>
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   505
to \<open>t\<^sub>1 < t\<^sub>4\<close> and the final ``\texttt{.}'' succeeds.
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   506
67299
ba52a058942f prefer formal citations;
wenzelm
parents: 67039
diff changeset
   507
For more information on this style of proof see @{cite "BauerW-TPHOLs01"}.
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   508
\fi
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   509
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
   510
\section{Streamlining Proofs}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   511
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
   512
\subsection{Pattern Matching and Quotations}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   513
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   514
In the proof patterns shown above, formulas are often duplicated.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   515
This can make the text harder to read, write and maintain. Pattern matching
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   516
is an abbreviation mechanism to avoid such duplication. Writing
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   517
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   518
\isacom{show} \ \textit{formula} \<open>(\<close>\indexed{\isacom{is}}{is} \textit{pattern}\<open>)\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   519
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   520
matches the pattern against the formula, thus instantiating the unknowns in
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   521
the pattern for later use. As an example, consider the proof pattern for
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   522
\<open>\<longleftrightarrow>\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   523
\end{isamarkuptext}%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   524
\<close>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   525
(*<*)lemma "formula\<^sub>1 \<longleftrightarrow> formula\<^sub>2" proof-(*>*)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   526
show "formula\<^sub>1 \<longleftrightarrow> formula\<^sub>2" (is "?L \<longleftrightarrow> ?R")
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   527
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   528
  assume "?L"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   529
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   530
  show "?R" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   531
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   532
  assume "?R"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   533
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   534
  show "?L" (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   535
qed(*<*)qed(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   536
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   537
text\<open>Instead of duplicating \<open>formula\<^sub>i\<close> in the text, we introduce
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   538
the two abbreviations \<open>?L\<close> and \<open>?R\<close> by pattern matching.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   539
Pattern matching works wherever a formula is stated, in particular
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   540
with \isacom{have} and \isacom{lemma}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   541
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   542
The unknown \indexed{\<open>?thesis\<close>}{thesis} is implicitly matched against any goal stated by
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   543
\isacom{lemma} or \isacom{show}. Here is a typical example:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   544
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   545
lemma "formula"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   546
proof -
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   547
  text_raw\<open>\\\mbox{}\quad$\vdots$\\\mbox{}\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   548
  show ?thesis (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   549
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   550
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   551
text\<open>
55359
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
   552
Unknowns can also be instantiated with \indexed{\isacom{let}}{let} commands
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   553
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   554
\isacom{let} \<open>?t\<close> = \<open>"\<close>\textit{some-big-term}\<open>"\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   555
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   556
Later proof steps can refer to \<open>?t\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   557
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   558
\isacom{have} \<open>"\<close>\dots \<open>?t\<close> \dots\<open>"\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   559
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   560
\begin{warn}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   561
Names of facts are introduced with \<open>name:\<close> and refer to proved
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   562
theorems. Unknowns \<open>?X\<close> refer to terms or formulas.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   563
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   564
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   565
Although abbreviations shorten the text, the reader needs to remember what
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   566
they stand for. Similarly for names of facts. Names like \<open>1\<close>, \<open>2\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   567
and \<open>3\<close> are not helpful and should only be used in short proofs. For
47704
8b4cd98f944e doc update
nipkow
parents: 47306
diff changeset
   568
longer proofs, descriptive names are better. But look at this example:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   569
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   570
\isacom{have} \ \<open>x_gr_0: "x > 0"\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   571
$\vdots$\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   572
\isacom{from} \<open>x_gr_0\<close> \dots
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   573
\end{quote}
56989
nipkow
parents: 56451
diff changeset
   574
The name is longer than the fact it stands for! Short facts do not need names;
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   575
one can refer to them easily by quoting them:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   576
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   577
\isacom{have} \ \<open>"x > 0"\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   578
$\vdots$\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   579
\isacom{from} \<open>`x>0`\<close> \dots\index{$IMP053@\<open>`...`\<close>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   580
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   581
Note that the quotes around \<open>x>0\<close> are \conceptnoidx{back quotes}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   582
They refer to the fact not by name but by value.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   583
55359
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
   584
\subsection{\indexed{\isacom{moreover}}{moreover}}
2d8222c76020 indexed document
nipkow
parents: 55317
diff changeset
   585
\index{ultimately@\isacom{ultimately}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   586
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   587
Sometimes one needs a number of facts to enable some deduction. Of course
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   588
one can name these facts individually, as shown on the right,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   589
but one can also combine them with \isacom{moreover}, as shown on the left:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   590
\<close>text_raw\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   591
\begin{tabular}{@ {}ll@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   592
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   593
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   594
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   595
(*<*)lemma "P" proof-(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   596
have "P\<^sub>1" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   597
moreover have "P\<^sub>2" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   598
moreover
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   599
text_raw\<open>\\$\vdots$\\\hspace{-1.4ex}\<close>(*<*)have "True" ..(*>*)
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   600
moreover have "P\<^sub>n" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   601
ultimately have "P"  (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   602
(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   603
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   604
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   605
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   606
&
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   607
\qquad
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   608
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   609
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   610
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   611
(*<*)lemma "P" proof-(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   612
have lab\<^sub>1: "P\<^sub>1" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   613
have lab\<^sub>2: "P\<^sub>2" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   614
text_raw\<open>\\$\vdots$\\\hspace{-1.4ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   615
have lab\<^sub>n: "P\<^sub>n" (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   616
from lab\<^sub>1 lab\<^sub>2 text_raw\<open>\ $\dots$\\\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   617
have "P"  (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   618
(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   619
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   620
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   621
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   622
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   623
\begin{isamarkuptext}%
65348
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   624
The \isacom{moreover} version is no shorter but expresses the structure
b5ce7100ddc8 added chains of (in)equations
nipkow
parents: 64852
diff changeset
   625
a bit more clearly and avoids new names.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   626
67039
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   627
\subsection{Local Lemmas}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   628
56989
nipkow
parents: 56451
diff changeset
   629
Sometimes one would like to prove some lemma locally within a proof,
nipkow
parents: 56451
diff changeset
   630
a lemma that shares the current context of assumptions but that
58502
nipkow
parents: 58486
diff changeset
   631
has its own assumptions and is generalized over its locally fixed
67039
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   632
variables at the end. This is simply an extension of the basic
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   633
\indexed{\isacom{have}}{have} construct:
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   634
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   635
\indexed{\isacom{have}}{have} \<open>B\<close>\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   636
 \indexed{\isacom{if}}{if} \<open>name:\<close> \<open>A\<^sub>1 \<dots> A\<^sub>m\<close>\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   637
 \indexed{\isacom{for}}{for} \<open>x\<^sub>1 \<dots> x\<^sub>n\<close>\\
67039
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   638
\isasymproof
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   639
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   640
proves \<open>\<lbrakk> A\<^sub>1; \<dots> ; A\<^sub>m \<rbrakk> \<Longrightarrow> B\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   641
where all \<open>x\<^sub>i\<close> have been replaced by unknowns \<open>?x\<^sub>i\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   642
As an example we prove a simple fact about divisibility on integers.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   643
The definition of \<open>dvd\<close> is @{thm dvd_def}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   644
\end{isamarkuptext}%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   645
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   646
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   647
lemma fixes a b :: int assumes "b dvd (a+b)" shows "b dvd a"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   648
proof -
67039
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   649
  have "\<exists>k'. a = b*k'" if asm: "a+b = b*k" for k
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   650
  proof
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   651
    show "a = b*(k - 1)" using asm by(simp add: algebra_simps)
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   652
  qed
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   653
  then show ?thesis using assms by(auto simp add: dvd_def)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   654
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   655
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   656
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   657
54436
nipkow
parents: 54292
diff changeset
   658
\subsection*{Exercises}
52706
5f86b30badd9 added exerciese
nipkow
parents: 52661
diff changeset
   659
52661
a3b04f0ab6a4 added exercise
nipkow
parents: 52593
diff changeset
   660
\exercise
a3b04f0ab6a4 added exercise
nipkow
parents: 52593
diff changeset
   661
Give a readable, structured proof of the following lemma:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   662
\<close>
54218
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
   663
lemma assumes T: "\<forall>x y. T x y \<or> T y x"
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
   664
  and A: "\<forall>x y. A x y \<and> A y x \<longrightarrow> x = y"
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
   665
  and TA: "\<forall>x y. T x y \<longrightarrow> A x y" and "A x y"
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
   666
  shows "T x y"
52661
a3b04f0ab6a4 added exercise
nipkow
parents: 52593
diff changeset
   667
(*<*)oops(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   668
text\<open>
52661
a3b04f0ab6a4 added exercise
nipkow
parents: 52593
diff changeset
   669
\endexercise
a3b04f0ab6a4 added exercise
nipkow
parents: 52593
diff changeset
   670
52706
5f86b30badd9 added exerciese
nipkow
parents: 52661
diff changeset
   671
\exercise
5f86b30badd9 added exerciese
nipkow
parents: 52661
diff changeset
   672
Give a readable, structured proof of the following lemma:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   673
\<close>
67039
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   674
lemma "\<exists>ys zs. xs = ys @ zs \<and>
690b4b334889 Replaced Raw Proof Blocks by Local Lemmas
nipkow
parents: 65437
diff changeset
   675
            (length ys = length zs \<or> length ys = length zs + 1)"
52706
5f86b30badd9 added exerciese
nipkow
parents: 52661
diff changeset
   676
(*<*)oops(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   677
text\<open>
52706
5f86b30badd9 added exerciese
nipkow
parents: 52661
diff changeset
   678
Hint: There are predefined functions @{const_typ take} and @{const_typ drop}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   679
such that \<open>take k [x\<^sub>1,\<dots>] = [x\<^sub>1,\<dots>,x\<^sub>k]\<close> and
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   680
\<open>drop k [x\<^sub>1,\<dots>] = [x\<^bsub>k+1\<^esub>,\<dots>]\<close>. Let sledgehammer find and apply
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   681
the relevant \<^const>\<open>take\<close> and \<^const>\<open>drop\<close> lemmas for you.
52706
5f86b30badd9 added exerciese
nipkow
parents: 52661
diff changeset
   682
\endexercise
5f86b30badd9 added exerciese
nipkow
parents: 52661
diff changeset
   683
54218
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
   684
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
   685
\section{Case Analysis and Induction}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   686
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
   687
\subsection{Datatype Case Analysis}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   688
\index{case analysis|(}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   689
47711
c1cca2a052e4 doc update
nipkow
parents: 47704
diff changeset
   690
We have seen case analysis on formulas. Now we want to distinguish
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   691
which form some term takes: is it \<open>0\<close> or of the form \<^term>\<open>Suc n\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   692
is it \<^term>\<open>[]\<close> or of the form \<^term>\<open>x#xs\<close>, etc. Here is a typical example
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   693
proof by case analysis on the form of \<open>xs\<close>:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   694
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   695
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   696
lemma "length(tl xs) = length xs - 1"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   697
proof (cases xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   698
  assume "xs = []"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   699
  thus ?thesis by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   700
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   701
  fix y ys assume "xs = y#ys"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   702
  thus ?thesis by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   703
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   704
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   705
text\<open>\index{cases@\<open>cases\<close>|(}Function \<open>tl\<close> (''tail'') is defined by @{thm list.sel(2)} and
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   706
@{thm list.sel(3)}. Note that the result type of \<^const>\<open>length\<close> is \<^typ>\<open>nat\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   707
and \<^prop>\<open>0 - 1 = (0::nat)\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   708
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   709
This proof pattern works for any term \<open>t\<close> whose type is a datatype.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   710
The goal has to be proved for each constructor \<open>C\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   711
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   712
\isacom{fix} \ \<open>x\<^sub>1 \<dots> x\<^sub>n\<close> \isacom{assume} \<open>"t = C x\<^sub>1 \<dots> x\<^sub>n"\<close>
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   713
\end{quote}\index{case@\isacom{case}|(}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   714
Each case can be written in a more compact form by means of the \isacom{case}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   715
command:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   716
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   717
\isacom{case} \<open>(C x\<^sub>1 \<dots> x\<^sub>n)\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   718
\end{quote}
47704
8b4cd98f944e doc update
nipkow
parents: 47306
diff changeset
   719
This is equivalent to the explicit \isacom{fix}-\isacom{assume} line
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   720
but also gives the assumption \<open>"t = C x\<^sub>1 \<dots> x\<^sub>n"\<close> a name: \<open>C\<close>,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   721
like the constructor.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   722
Here is the \isacom{case} version of the proof above:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   723
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   724
(*<*)lemma "length(tl xs) = length xs - 1"(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   725
proof (cases xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   726
  case Nil
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   727
  thus ?thesis by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   728
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   729
  case (Cons y ys)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   730
  thus ?thesis by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   731
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   732
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   733
text\<open>Remember that \<open>Nil\<close> and \<open>Cons\<close> are the alphanumeric names
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   734
for \<open>[]\<close> and \<open>#\<close>. The names of the assumptions
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   735
are not used because they are directly piped (via \isacom{thus})
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   736
into the proof of the claim.
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   737
\index{case analysis|)}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   738
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
   739
\subsection{Structural Induction}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   740
\index{induction|(}
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   741
\index{structural induction|(}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   742
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   743
We illustrate structural induction with an example based on natural numbers:
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   744
the sum (\<open>\<Sum>\<close>) of the first \<open>n\<close> natural numbers
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   745
(\<open>{0..n::nat}\<close>) is equal to \mbox{\<^term>\<open>n*(n+1) div 2::nat\<close>}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   746
Never mind the details, just focus on the pattern:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   747
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   748
47711
c1cca2a052e4 doc update
nipkow
parents: 47704
diff changeset
   749
lemma "\<Sum>{0..n::nat} = n*(n+1) div 2"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   750
proof (induction n)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   751
  show "\<Sum>{0..0::nat} = 0*(0+1) div 2" by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   752
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   753
  fix n assume "\<Sum>{0..n::nat} = n*(n+1) div 2"
47711
c1cca2a052e4 doc update
nipkow
parents: 47704
diff changeset
   754
  thus "\<Sum>{0..Suc n} = Suc n*(Suc n+1) div 2" by simp
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   755
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   756
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   757
text\<open>Except for the rewrite steps, everything is explicitly given. This
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   758
makes the proof easily readable, but the duplication means it is tedious to
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   759
write and maintain. Here is how pattern
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   760
matching can completely avoid any duplication:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   761
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   762
lemma "\<Sum>{0..n::nat} = n*(n+1) div 2" (is "?P n")
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   763
proof (induction n)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   764
  show "?P 0" by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   765
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   766
  fix n assume "?P n"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   767
  thus "?P(Suc n)" by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   768
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   769
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   770
text\<open>The first line introduces an abbreviation \<open>?P n\<close> for the goal.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   771
Pattern matching \<open>?P n\<close> with the goal instantiates \<open>?P\<close> to the
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   772
function \<^term>\<open>\<lambda>n. \<Sum>{0..n::nat} = n*(n+1) div 2\<close>.  Now the proposition to
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   773
be proved in the base case can be written as \<open>?P 0\<close>, the induction
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   774
hypothesis as \<open>?P n\<close>, and the conclusion of the induction step as
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   775
\<open>?P(Suc n)\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   776
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   777
Induction also provides the \isacom{case} idiom that abbreviates
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   778
the \isacom{fix}-\isacom{assume} step. The above proof becomes
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   779
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   780
(*<*)lemma "\<Sum>{0..n::nat} = n*(n+1) div 2"(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   781
proof (induction n)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   782
  case 0
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   783
  show ?case by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   784
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   785
  case (Suc n)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   786
  thus ?case by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   787
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   788
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   789
text\<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   790
The unknown \<open>?case\<close>\index{case?@\<open>?case\<close>|(} is set in each case to the required
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   791
claim, i.e., \<open>?P 0\<close> and \mbox{\<open>?P(Suc n)\<close>} in the above proof,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   792
without requiring the user to define a \<open>?P\<close>. The general
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   793
pattern for induction over \<^typ>\<open>nat\<close> is shown on the left-hand side:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   794
\<close>text_raw\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   795
\begin{tabular}{@ {}ll@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   796
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   797
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   798
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   799
(*<*)lemma "P(n::nat)" proof -(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   800
show "P(n)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   801
proof (induction n)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   802
  case 0
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   803
  text_raw\<open>\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   804
  show ?case (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   805
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   806
  case (Suc n)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   807
  text_raw\<open>\\\mbox{}\ \ $\vdots$\\\mbox{}\hspace{-1ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   808
  show ?case (*<*)sorry(*>*) text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   809
qed(*<*)qed(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   810
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   811
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   812
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   813
&
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   814
\begin{minipage}[t]{.4\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   815
~\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   816
~\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   817
\isacom{let} \<open>?case = "P(0)"\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   818
~\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   819
~\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   820
~\\[1ex]
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   821
\isacom{fix} \<open>n\<close> \isacom{assume} \<open>Suc: "P(n)"\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   822
\isacom{let} \<open>?case = "P(Suc n)"\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   823
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   824
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   825
\medskip
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   826
\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   827
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   828
On the right side you can see what the \isacom{case} command
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   829
on the left stands for.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   830
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   831
In case the goal is an implication, induction does one more thing: the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   832
proposition to be proved in each case is not the whole implication but only
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   833
its conclusion; the premises of the implication are immediately made
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   834
assumptions of that case. That is, if in the above proof we replace
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   835
\isacom{show}~\<open>"P(n)"\<close> by
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   836
\mbox{\isacom{show}~\<open>"A(n) \<Longrightarrow> P(n)"\<close>}
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   837
then \isacom{case}~\<open>0\<close> stands for
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   838
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   839
\isacom{assume} \ \<open>0: "A(0)"\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   840
\isacom{let} \<open>?case = "P(0)"\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   841
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   842
and \isacom{case}~\<open>(Suc n)\<close> stands for
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   843
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   844
\isacom{fix} \<open>n\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   845
\isacom{assume} \<open>Suc:\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   846
  \begin{tabular}[t]{l}\<open>"A(n) \<Longrightarrow> P(n)"\<close>\\\<open>"A(Suc n)"\<close>\end{tabular}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   847
\isacom{let} \<open>?case = "P(Suc n)"\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   848
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   849
The list of assumptions \<open>Suc\<close> is actually subdivided
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   850
into \<open>Suc.IH\<close>, the induction hypotheses (here \<open>A(n) \<Longrightarrow> P(n)\<close>),
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   851
and \<open>Suc.prems\<close>, the premises of the goal being proved
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   852
(here \<open>A(Suc n)\<close>).
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   853
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   854
Induction works for any datatype.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   855
Proving a goal \<open>\<lbrakk> A\<^sub>1(x); \<dots>; A\<^sub>k(x) \<rbrakk> \<Longrightarrow> P(x)\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   856
by induction on \<open>x\<close> generates a proof obligation for each constructor
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   857
\<open>C\<close> of the datatype. The command \isacom{case}~\<open>(C x\<^sub>1 \<dots> x\<^sub>n)\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   858
performs the following steps:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   859
\begin{enumerate}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   860
\item \isacom{fix} \<open>x\<^sub>1 \<dots> x\<^sub>n\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   861
\item \isacom{assume} the induction hypotheses (calling them \<open>C.IH\<close>\index{IH@\<open>.IH\<close>})
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   862
 and the premises \mbox{\<open>A\<^sub>i(C x\<^sub>1 \<dots> x\<^sub>n)\<close>} (calling them \<open>C.prems\<close>\index{prems@\<open>.prems\<close>})
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   863
 and calling the whole list \<open>C\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   864
\item \isacom{let} \<open>?case = "P(C x\<^sub>1 \<dots> x\<^sub>n)"\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   865
\end{enumerate}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   866
\index{structural induction|)}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   867
65437
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   868
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   869
\ifsem\else
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   870
\subsection{Computation Induction}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   871
\index{rule induction}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   872
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   873
In \autoref{sec:recursive-funs} we introduced computation induction and
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   874
its realization in Isabelle: the definition
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   875
of a recursive function \<open>f\<close> via \isacom{fun} proves the corresponding computation
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   876
induction rule called \<open>f.induct\<close>. Induction with this rule looks like in
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   877
\autoref{sec:recursive-funs}, but now with \isacom{proof} instead of \isacom{apply}:
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   878
\begin{quote}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   879
\isacom{proof} (\<open>induction x\<^sub>1 \<dots> x\<^sub>k rule: f.induct\<close>)
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   880
\end{quote}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   881
Just as for structural induction, this creates several cases, one for each
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   882
defining equation for \<open>f\<close>. By default (if the equations have not been named
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   883
by the user), the cases are numbered. That is, they are started by
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   884
\begin{quote}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   885
\isacom{case} (\<open>i x y ...\<close>)
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   886
\end{quote}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   887
where \<open>i = 1,...,n\<close>, \<open>n\<close> is the number of equations defining \<open>f\<close>,
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   888
and \<open>x y ...\<close> are the variables in equation \<open>i\<close>. Note the following:
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   889
\begin{itemize}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   890
\item
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   891
Although \<open>i\<close> is an Isar name, \<open>i.IH\<close> (or similar) is not. You need
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   892
double quotes: "\<open>i.IH\<close>". When indexing the name, write "\<open>i.IH\<close>"(1),
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   893
not "\<open>i.IH\<close>(1)".
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   894
\item
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   895
If defining equations for \<open>f\<close> overlap, \isacom{fun} instantiates them to make
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   896
them nonoverlapping. This means that one user-provided equation may lead to
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   897
several equations and thus to several cases in the induction rule.
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   898
These have names of the form "\<open>i_j\<close>", where \<open>i\<close> is the number of the original
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   899
equation and the system-generated \<open>j\<close> indicates the subcase.
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   900
\end{itemize}
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   901
In Isabelle/jEdit, the \<open>induction\<close> proof method displays a proof skeleton
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   902
with all \isacom{case}s. This is particularly useful for computation induction
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   903
and the following rule induction.
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   904
\fi
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   905
b8fc7e2e1b35 doc for comp. ind.
nipkow
parents: 65352
diff changeset
   906
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
   907
\subsection{Rule Induction}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
   908
\index{rule induction|(}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   909
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   910
Recall the inductive and recursive definitions of even numbers in
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   911
\autoref{sec:inductive-defs}:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   912
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   913
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   914
inductive ev :: "nat \<Rightarrow> bool" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   915
ev0: "ev 0" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   916
evSS: "ev n \<Longrightarrow> ev(Suc(Suc n))"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   917
61429
63fb7a68a12c even -> evn to avoid clash with global even
nipkow
parents: 61022
diff changeset
   918
fun evn :: "nat \<Rightarrow> bool" where
63fb7a68a12c even -> evn to avoid clash with global even
nipkow
parents: 61022
diff changeset
   919
"evn 0 = True" |
63fb7a68a12c even -> evn to avoid clash with global even
nipkow
parents: 61022
diff changeset
   920
"evn (Suc 0) = False" |
63fb7a68a12c even -> evn to avoid clash with global even
nipkow
parents: 61022
diff changeset
   921
"evn (Suc(Suc n)) = evn n"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   922
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   923
text\<open>We recast the proof of \<^prop>\<open>ev n \<Longrightarrow> evn n\<close> in Isar. The
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   924
left column shows the actual proof text, the right column shows
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   925
the implicit effect of the two \isacom{case} commands:\<close>text_raw\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   926
\begin{tabular}{@ {}l@ {\qquad}l@ {}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   927
\begin{minipage}[t]{.5\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   928
\isa{%
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   929
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   930
61429
63fb7a68a12c even -> evn to avoid clash with global even
nipkow
parents: 61022
diff changeset
   931
lemma "ev n \<Longrightarrow> evn n"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   932
proof(induction rule: ev.induct)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   933
  case ev0
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   934
  show ?case by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   935
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   936
  case evSS
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   937
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   938
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   939
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   940
  thus ?case by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   941
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   942
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   943
text_raw \<open>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   944
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   945
&
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   946
\begin{minipage}[t]{.5\textwidth}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   947
~\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   948
~\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   949
\isacom{let} \<open>?case = "evn 0"\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   950
~\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   951
~\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   952
\isacom{fix} \<open>n\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   953
\isacom{assume} \<open>evSS:\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   954
  \begin{tabular}[t]{l} \<open>"ev n"\<close>\\\<open>"evn n"\<close>\end{tabular}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   955
\isacom{let} \<open>?case = "evn(Suc(Suc n))"\<close>\\
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   956
\end{minipage}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   957
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   958
\medskip
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   959
\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   960
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   961
The proof resembles structural induction, but the induction rule is given
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   962
explicitly and the names of the cases are the names of the rules in the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   963
inductive definition.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   964
Let us examine the two assumptions named @{thm[source]evSS}:
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   965
\<^prop>\<open>ev n\<close> is the premise of rule @{thm[source]evSS}, which we may assume
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   966
because we are in the case where that rule was used; \<^prop>\<open>evn n\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   967
is the induction hypothesis.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   968
\begin{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   969
Because each \isacom{case} command introduces a list of assumptions
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   970
named like the case name, which is the name of a rule of the inductive
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   971
definition, those rules now need to be accessed with a qualified name, here
58522
nipkow
parents: 58521
diff changeset
   972
@{thm[source] ev.ev0} and @{thm[source] ev.evSS}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   973
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   974
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   975
In the case @{thm[source]evSS} of the proof above we have pretended that the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   976
system fixes a variable \<open>n\<close>.  But unless the user provides the name
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   977
\<open>n\<close>, the system will just invent its own name that cannot be referred
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   978
to.  In the above proof, we do not need to refer to it, hence we do not give
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   979
it a specific name. In case one needs to refer to it one writes
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   980
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   981
\isacom{case} \<open>(evSS m)\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   982
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   983
like \isacom{case}~\<open>(Suc n)\<close> in earlier structural inductions.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   984
The name \<open>m\<close> is an arbitrary choice. As a result,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   985
case @{thm[source] evSS} is derived from a renamed version of
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   986
rule @{thm[source] evSS}: \<open>ev m \<Longrightarrow> ev(Suc(Suc m))\<close>.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
   987
Here is an example with a (contrived) intermediate step that refers to \<open>m\<close>:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   988
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   989
61429
63fb7a68a12c even -> evn to avoid clash with global even
nipkow
parents: 61022
diff changeset
   990
lemma "ev n \<Longrightarrow> evn n"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   991
proof(induction rule: ev.induct)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   992
  case ev0 show ?case by simp
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   993
next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   994
  case (evSS m)
61429
63fb7a68a12c even -> evn to avoid clash with global even
nipkow
parents: 61022
diff changeset
   995
  have "evn(Suc(Suc m)) = evn m" by simp
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 69597
diff changeset
   996
  thus ?case using `evn m` by blast
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   997
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   998
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
   999
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1000
\indent
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1001
In general, let \<open>I\<close> be a (for simplicity unary) inductively defined
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1002
predicate and let the rules in the definition of \<open>I\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1003
be called \<open>rule\<^sub>1\<close>, \dots, \<open>rule\<^sub>n\<close>. A proof by rule
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1004
induction follows this pattern:\index{inductionrule@\<open>induction ... rule:\<close>}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1005
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1006
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1007
(*<*)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
  1008
inductive I where rule\<^sub>1: "I()" |  rule\<^sub>2: "I()" |  rule\<^sub>n: "I()"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1009
lemma "I x \<Longrightarrow> P x" proof-(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1010
show "I x \<Longrightarrow> P x"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1011
proof(induction rule: I.induct)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
  1012
  case rule\<^sub>1
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1013
  text_raw\<open>\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1014
  show ?case (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1015
next
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1016
  text_raw\<open>\\[-.4ex]$\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1017
(*<*)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
  1018
  case rule\<^sub>2
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1019
  show ?case sorry
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1020
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1021
next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
  1022
  case rule\<^sub>n
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1023
  text_raw\<open>\\[-.4ex]\mbox{}\ \ $\vdots$\\[-.4ex]\mbox{}\hspace{-1ex}\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1024
  show ?case (*<*)sorry(*>*)text_raw\<open>\ \isasymproof\\\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1025
qed(*<*)qed(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1026
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1027
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1028
One can provide explicit variable names by writing
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1029
\isacom{case}~\<open>(rule\<^sub>i x\<^sub>1 \<dots> x\<^sub>k)\<close>, thus renaming the first \<open>k\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1030
free variables in rule \<open>i\<close> to \<open>x\<^sub>1 \<dots> x\<^sub>k\<close>,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1031
going through rule \<open>i\<close> from left to right.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1032
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
  1033
\subsection{Assumption Naming}
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1034
\label{sec:assm-naming}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1035
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1036
In any induction, \isacom{case}~\<open>name\<close> sets up a list of assumptions
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1037
also called \<open>name\<close>, which is subdivided into three parts:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1038
\begin{description}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1039
\item[\<open>name.IH\<close>]\index{IH@\<open>.IH\<close>} contains the induction hypotheses.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1040
\item[\<open>name.hyps\<close>]\index{hyps@\<open>.hyps\<close>} contains all the other hypotheses of this case in the
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1041
induction rule. For rule inductions these are the hypotheses of rule
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1042
\<open>name\<close>, for structural inductions these are empty.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1043
\item[\<open>name.prems\<close>]\index{prems@\<open>.prems\<close>} contains the (suitably instantiated) premises
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1044
of the statement being proved, i.e., the \<open>A\<^sub>i\<close> when
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1045
proving \<open>\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1046
\end{description}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1047
\begin{warn}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1048
Proof method \<open>induct\<close> differs from \<open>induction\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1049
only in this naming policy: \<open>induct\<close> does not distinguish
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1050
\<open>IH\<close> from \<open>hyps\<close> but subsumes \<open>IH\<close> under \<open>hyps\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1051
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1052
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1053
More complicated inductive proofs than the ones we have seen so far
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1054
often need to refer to specific assumptions --- just \<open>name\<close> or even
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1055
\<open>name.prems\<close> and \<open>name.IH\<close> can be too unspecific.
58504
nipkow
parents: 58502
diff changeset
  1056
This is where the indexing of fact lists comes in handy, e.g.,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1057
\<open>name.IH(2)\<close> or \<open>name.prems(1-2)\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1058
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
  1059
\subsection{Rule Inversion}
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
  1060
\label{sec:rule-inversion}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
  1061
\index{rule inversion|(}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1062
47711
c1cca2a052e4 doc update
nipkow
parents: 47704
diff changeset
  1063
Rule inversion is case analysis of which rule could have been used to
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
  1064
derive some fact. The name \conceptnoidx{rule inversion} emphasizes that we are
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1065
reasoning backwards: by which rules could some given fact have been proved?
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1066
For the inductive definition of \<^const>\<open>ev\<close>, rule inversion can be summarized
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1067
like this:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
  1068
@{prop[display]"ev n \<Longrightarrow> n = 0 \<or> (\<exists>k. n = Suc(Suc k) \<and> ev k)"}
47711
c1cca2a052e4 doc update
nipkow
parents: 47704
diff changeset
  1069
The realisation in Isabelle is a case analysis.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1070
A simple example is the proof that \<^prop>\<open>ev n \<Longrightarrow> ev (n - 2)\<close>. We
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1071
already went through the details informally in \autoref{sec:Logic:even}. This
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1072
is the Isar proof:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1073
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1074
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1075
notepad
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1076
begin fix n
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1077
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1078
  assume "ev n"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1079
  from this have "ev(n - 2)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1080
  proof cases
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1081
    case ev0 thus "ev(n - 2)" by (simp add: ev.ev0)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1082
  next
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1083
    case (evSS k) thus "ev(n - 2)" by (simp add: ev.evSS)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1084
  qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1085
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1086
end
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1087
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1088
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1089
text\<open>The key point here is that a case analysis over some inductively
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1090
defined predicate is triggered by piping the given fact
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1091
(here: \isacom{from}~\<open>this\<close>) into a proof by \<open>cases\<close>.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1092
Let us examine the assumptions available in each case. In case \<open>ev0\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1093
we have \<open>n = 0\<close> and in case \<open>evSS\<close> we have \<^prop>\<open>n = Suc(Suc k)\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1094
and \<^prop>\<open>ev k\<close>. In each case the assumptions are available under the name
56989
nipkow
parents: 56451
diff changeset
  1095
of the case; there is no fine-grained naming schema like there is for induction.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1096
47704
8b4cd98f944e doc update
nipkow
parents: 47306
diff changeset
  1097
Sometimes some rules could not have been used to derive the given fact
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1098
because constructors clash. As an extreme example consider
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1099
rule inversion applied to \<^prop>\<open>ev(Suc 0)\<close>: neither rule \<open>ev0\<close> nor
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1100
rule \<open>evSS\<close> can yield \<^prop>\<open>ev(Suc 0)\<close> because \<open>Suc 0\<close> unifies
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1101
neither with \<open>0\<close> nor with \<^term>\<open>Suc(Suc n)\<close>. Impossible cases do not
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1102
have to be proved. Hence we can prove anything from \<^prop>\<open>ev(Suc 0)\<close>:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1103
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1104
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1105
notepad begin fix P
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1106
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1107
  assume "ev(Suc 0)" then have P by cases
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1108
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1109
end
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1110
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1111
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1112
text\<open>That is, \<^prop>\<open>ev(Suc 0)\<close> is simply not provable:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1113
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1114
lemma "\<not> ev(Suc 0)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1115
proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1116
  assume "ev(Suc 0)" then show False by cases
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1117
qed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1118
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1119
text\<open>Normally not all cases will be impossible. As a simple exercise,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1120
prove that \mbox{\<^prop>\<open>\<not> ev(Suc(Suc(Suc 0)))\<close>.}
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1121
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52059
diff changeset
  1122
\subsection{Advanced Rule Induction}
51445
nipkow
parents: 51443
diff changeset
  1123
\label{sec:advanced-rule-induction}
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1124
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1125
So far, rule induction was always applied to goals of the form \<open>I x y z \<Longrightarrow> \<dots>\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1126
where \<open>I\<close> is some inductively defined predicate and \<open>x\<close>, \<open>y\<close>, \<open>z\<close>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1127
are variables. In some rare situations one needs to deal with an assumption where
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1128
not all arguments \<open>r\<close>, \<open>s\<close>, \<open>t\<close> are variables:
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1129
\begin{isabelle}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1130
\isacom{lemma} \<open>"I r s t \<Longrightarrow> \<dots>"\<close>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1131
\end{isabelle}
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1132
Applying the standard form of
54577
nipkow
parents: 54436
diff changeset
  1133
rule induction in such a situation will lead to strange and typically unprovable goals.
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1134
We can easily reduce this situation to the standard one by introducing
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1135
new variables \<open>x\<close>, \<open>y\<close>, \<open>z\<close> and reformulating the goal like this:
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1136
\begin{isabelle}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1137
\isacom{lemma} \<open>"I x y z \<Longrightarrow> x = r \<Longrightarrow> y = s \<Longrightarrow> z = t \<Longrightarrow> \<dots>"\<close>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1138
\end{isabelle}
56989
nipkow
parents: 56451
diff changeset
  1139
Standard rule induction will work fine now, provided the free variables in
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1140
\<open>r\<close>, \<open>s\<close>, \<open>t\<close> are generalized via \<open>arbitrary\<close>.
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1141
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1142
However, induction can do the above transformation for us, behind the curtains, so we never
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1143
need to see the expanded version of the lemma. This is what we need to write:
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1144
\begin{isabelle}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1145
\isacom{lemma} \<open>"I r s t \<Longrightarrow> \<dots>"\<close>\isanewline
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1146
\isacom{proof}\<open>(induction "r" "s" "t" arbitrary: \<dots> rule: I.induct)\<close>\index{inductionrule@\<open>induction ... rule:\<close>}\index{arbitrary@\<open>arbitrary:\<close>}
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1147
\end{isabelle}
58605
nipkow
parents: 58602
diff changeset
  1148
Like for rule inversion, cases that are impossible because of constructor clashes
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1149
will not show up at all. Here is a concrete example:\<close>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1150
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1151
lemma "ev (Suc m) \<Longrightarrow> \<not> ev m"
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1152
proof(induction "Suc m" arbitrary: m rule: ev.induct)
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1153
  fix n assume IH: "\<And>m. n = Suc m \<Longrightarrow> \<not> ev m"
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1154
  show "\<not> ev (Suc n)"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67406
diff changeset
  1155
  proof \<comment> \<open>contradiction\<close>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1156
    assume "ev(Suc n)"
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1157
    thus False
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67406
diff changeset
  1158
    proof cases \<comment> \<open>rule inversion\<close>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1159
      fix k assume "n = Suc k" "ev k"
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1160
      thus False using IH by auto
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1161
    qed
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1162
  qed
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1163
qed
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1164
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1165
text\<open>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1166
Remarks:
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1167
\begin{itemize}
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1168
\item 
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1169
Instead of the \isacom{case} and \<open>?case\<close> magic we have spelled all formulas out.
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1170
This is merely for greater clarity.
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1171
\item
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1172
We only need to deal with one case because the @{thm[source] ev0} case is impossible.
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1173
\item
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1174
The form of the \<open>IH\<close> shows us that internally the lemma was expanded as explained
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1175
above: \noquotes{@{prop[source]"ev x \<Longrightarrow> x = Suc m \<Longrightarrow> \<not> ev m"}}.
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1176
\item
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1177
The goal \<^prop>\<open>\<not> ev (Suc n)\<close> may surprise. The expanded version of the lemma
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1178
would suggest that we have a \isacom{fix} \<open>m\<close> \isacom{assume} \<^prop>\<open>Suc(Suc n) = Suc m\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1179
and need to show \<^prop>\<open>\<not> ev m\<close>. What happened is that Isabelle immediately
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1180
simplified \<^prop>\<open>Suc(Suc n) = Suc m\<close> to \<^prop>\<open>Suc n = m\<close> and could then eliminate
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1181
\<open>m\<close>. Beware of such nice surprises with this advanced form of induction.
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1182
\end{itemize}
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1183
\begin{warn}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1184
This advanced form of induction does not support the \<open>IH\<close>
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1185
naming schema explained in \autoref{sec:assm-naming}:
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1186
the induction hypotheses are instead found under the name \<open>hyps\<close>,
56989
nipkow
parents: 56451
diff changeset
  1187
as they are for the simpler
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1188
\<open>induct\<close> method.
51443
4edb82207c5c added advanced rule induction subsection
nipkow
parents: 49837
diff changeset
  1189
\end{warn}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
  1190
\index{induction|)}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1191
\index{cases@\<open>cases\<close>|)}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
  1192
\index{case@\isacom{case}|)}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1193
\index{case?@\<open>?case\<close>|)}
55361
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
  1194
\index{rule induction|)}
d459a63ca42f more indexing
nipkow
parents: 55359
diff changeset
  1195
\index{rule inversion|)}
54218
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
  1196
54436
nipkow
parents: 54292
diff changeset
  1197
\subsection*{Exercises}
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
  1198
54232
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1199
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1200
\exercise
54292
ce4a17b2e373 more exercises
nipkow
parents: 54232
diff changeset
  1201
Give a structured proof by rule inversion:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1202
\<close>
54232
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1203
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1204
lemma assumes a: "ev(Suc(Suc n))" shows "ev n"
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1205
(*<*)oops(*>*)
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1206
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1207
text\<open>
54232
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1208
\endexercise
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1209
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
  1210
\begin{exercise}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1211
Give a structured proof of \<^prop>\<open>\<not> ev(Suc(Suc(Suc 0)))\<close>
54232
e039a9b9700d more exercises
nipkow
parents: 54218
diff changeset
  1212
by rule inversions. If there are no cases to be proved you can close
61013
nipkow
parents: 61012
diff changeset
  1213
a proof immediately with \isacom{qed}.
54218
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
  1214
\end{exercise}
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
  1215
07c0c121a8dc more exercises
nipkow
parents: 53015
diff changeset
  1216
\begin{exercise}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1217
Recall predicate \<open>star\<close> from \autoref{sec:star} and \<open>iter\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1218
from Exercise~\ref{exe:iter}. Prove \<^prop>\<open>iter r n x y \<Longrightarrow> star r x y\<close>
56989
nipkow
parents: 56451
diff changeset
  1219
in a structured style; do not just sledgehammer each case of the
54292
ce4a17b2e373 more exercises
nipkow
parents: 54232
diff changeset
  1220
required induction.
ce4a17b2e373 more exercises
nipkow
parents: 54232
diff changeset
  1221
\end{exercise}
ce4a17b2e373 more exercises
nipkow
parents: 54232
diff changeset
  1222
ce4a17b2e373 more exercises
nipkow
parents: 54232
diff changeset
  1223
\begin{exercise}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1224
Define a recursive function \<open>elems ::\<close> \<^typ>\<open>'a list \<Rightarrow> 'a set\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1225
and prove \<^prop>\<open>x \<in> elems xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> elems ys\<close>.
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
  1226
\end{exercise}
61012
40a0a4077126 nex exercise
nipkow
parents: 58999
diff changeset
  1227
40a0a4077126 nex exercise
nipkow
parents: 58999
diff changeset
  1228
\begin{exercise}
40a0a4077126 nex exercise
nipkow
parents: 58999
diff changeset
  1229
Extend Exercise~\ref{exe:cfg} with a function that checks if some
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1230
\mbox{\<open>alpha list\<close>} is a balanced
61022
1c4ae64636bb tuned spacing
nipkow
parents: 61021
diff changeset
  1231
string of parentheses. More precisely, define a \mbox{recursive} function
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1232
\<open>balanced :: nat \<Rightarrow> alpha list \<Rightarrow> bool\<close> such that \<^term>\<open>balanced n w\<close>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
  1233
is true iff (informally) \<open>S (a\<^sup>n @ w)\<close>. Formally, prove that
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1234
\<^prop>\<open>balanced n w \<longleftrightarrow> S (replicate n a @ w)\<close> where
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1235
\<^const>\<open>replicate\<close> \<open>::\<close> \<^typ>\<open>nat \<Rightarrow> 'a \<Rightarrow> 'a list\<close> is predefined
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
  1236
and \<^term>\<open>replicate n x\<close> yields the list \<open>[x, \<dots>, x]\<close> of length \<open>n\<close>.
61012
40a0a4077126 nex exercise
nipkow
parents: 58999
diff changeset
  1237
\end{exercise}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67299
diff changeset
  1238
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1239
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1240
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1241
end
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
  1242
(*>*)