src/Doc/Prog_Prove/Logic.thy
author wenzelm
Sun, 02 May 2021 14:07:19 +0200
changeset 73613 c1d8cd6d1a49
parent 72315 8162ca81ea8a
child 74763 dbac0ebb4a85
permissions -rw-r--r--
early definition of ML antiquotations;
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 Logic
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
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     5
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
     6
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     7
\vspace{-5ex}
51436
790310525e97 tuned (in particular bold fonts)
nipkow
parents: 51433
diff changeset
     8
\section{Formulas}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     9
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
    10
The core syntax of formulas (\textit{form} below)
47720
nipkow
parents: 47711
diff changeset
    11
provides the standard logical constructs, in decreasing order of precedence:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    12
\[
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    13
\begin{array}{rcl}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    14
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    15
\mathit{form} & ::= &
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    16
  \<open>(form)\<close> ~\mid~
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    17
  \<^const>\<open>True\<close> ~\mid~
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    18
  \<^const>\<open>False\<close> ~\mid~
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    19
  \<^prop>\<open>term = term\<close>\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    20
 &\mid& \<^prop>\<open>\<not> form\<close>\index{$HOL4@\isasymnot} ~\mid~
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    21
  \<^prop>\<open>form \<and> form\<close>\index{$HOL0@\isasymand} ~\mid~
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    22
  \<^prop>\<open>form \<or> form\<close>\index{$HOL1@\isasymor} ~\mid~
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    23
  \<^prop>\<open>form \<longrightarrow> form\<close>\index{$HOL2@\isasymlongrightarrow}\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    24
 &\mid& \<^prop>\<open>\<forall>x. form\<close>\index{$HOL6@\isasymforall} ~\mid~  \<^prop>\<open>\<exists>x. form\<close>\index{$HOL7@\isasymexists}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    25
\end{array}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    26
\]
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
    27
Terms are the ones we have seen all along, built from constants, variables,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    28
function application and \<open>\<lambda>\<close>-abstraction, including all the syntactic
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    29
sugar like infix symbols, \<open>if\<close>, \<open>case\<close>, etc.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    30
\begin{warn}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    31
Remember that formulas are simply terms of type \<open>bool\<close>. Hence
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    32
\<open>=\<close> also works for formulas. Beware that \<open>=\<close> has a higher
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    33
precedence than the other logical operators. Hence \<^prop>\<open>s = t \<and> A\<close> means
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    34
\<open>(s = t) \<and> A\<close>, and \<^prop>\<open>A\<and>B = B\<and>A\<close> means \<open>A \<and> (B = B) \<and> A\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    35
Logical equivalence can also be written with
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    36
\<open>\<longleftrightarrow>\<close> instead of \<open>=\<close>, where \<open>\<longleftrightarrow>\<close> has the same low
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    37
precedence as \<open>\<longrightarrow>\<close>. Hence \<open>A \<and> B \<longleftrightarrow> B \<and> A\<close> really means
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    38
\<open>(A \<and> B) \<longleftrightarrow> (B \<and> A)\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    39
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    40
\begin{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    41
Quantifiers need to be enclosed in parentheses if they are nested within
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    42
other constructs (just like \<open>if\<close>, \<open>case\<close> and \<open>let\<close>).
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    43
\end{warn}
52404
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    44
The most frequent logical symbols and their ASCII representations are shown
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    45
in Fig.~\ref{fig:log-symbols}.
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    46
\begin{figure}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    47
\begin{center}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    48
\begin{tabular}{l@ {\qquad}l@ {\qquad}l}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    49
\<open>\<forall>\<close> & \xsymbol{forall} & \texttt{ALL}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    50
\<open>\<exists>\<close> & \xsymbol{exists} & \texttt{EX}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    51
\<open>\<lambda>\<close> & \xsymbol{lambda} & \texttt{\%}\\
72315
8162ca81ea8a suppress ligatures more robustly, notably for lualatex;
wenzelm
parents: 69597
diff changeset
    52
\<open>\<longrightarrow>\<close> & \texttt{-{\kern0pt}->}\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    53
\<open>\<longleftrightarrow>\<close> & \texttt{<->}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    54
\<open>\<and>\<close> & \texttt{/\char`\\} & \texttt{\&}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    55
\<open>\<or>\<close> & \texttt{\char`\\/} & \texttt{|}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    56
\<open>\<not>\<close> & \xsymbol{not} & \texttt{\char`~}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    57
\<open>\<noteq>\<close> & \xsymbol{noteq} & \texttt{\char`~=}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    58
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    59
\end{center}
52404
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    60
\caption{Logical symbols and their ASCII forms}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    61
\label{fig:log-symbols}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    62
\end{figure}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    63
The first column shows the symbols, the other columns ASCII representations.
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    64
The \texttt{\char`\\}\texttt{<...>} form is always converted into the symbolic form
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    65
by the Isabelle interfaces, the treatment of the other ASCII forms
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    66
depends on the interface. The ASCII forms \texttt{/\char`\\} and
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    67
\texttt{\char`\\/}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    68
are special in that they are merely keyboard shortcuts for the interface and
33524382335b more set theory
nipkow
parents: 52361
diff changeset
    69
not logical symbols by themselves.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    70
\begin{warn}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    71
The implication \<open>\<Longrightarrow>\<close> is part of the Isabelle framework. It structures
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
    72
theorems and proof states, separating assumptions from conclusions.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    73
The implication \<open>\<longrightarrow>\<close> is part of the logic HOL and can occur inside the
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    74
formulas that make up the assumptions and conclusion.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    75
Theorems should be of the form \<open>\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A\<close>,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    76
not \<open>A\<^sub>1 \<and> \<dots> \<and> A\<^sub>n \<longrightarrow> A\<close>. Both are logically equivalent
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    77
but the first one works better when using the theorem in further proofs.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    78
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    79
51436
790310525e97 tuned (in particular bold fonts)
nipkow
parents: 51433
diff changeset
    80
\section{Sets}
51038
nipkow
parents: 49615
diff changeset
    81
\label{sec:Sets}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    82
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    83
Sets of elements of type \<^typ>\<open>'a\<close> have type \<^typ>\<open>'a set\<close>\index{set@\<open>set\<close>}.
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
    84
They can be finite or infinite. Sets come with the usual notation:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    85
\begin{itemize}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    86
\item \indexed{\<^term>\<open>{}\<close>}{$IMP042},\quad \<open>{e\<^sub>1,\<dots>,e\<^sub>n}\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    87
\item \<^prop>\<open>e \<in> A\<close>\index{$HOLSet0@\isasymin},\quad \<^prop>\<open>A \<subseteq> B\<close>\index{$HOLSet2@\isasymsubseteq}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    88
\item \<^term>\<open>A \<union> B\<close>\index{$HOLSet4@\isasymunion},\quad \<^term>\<open>A \<inter> B\<close>\index{$HOLSet5@\isasyminter},\quad \<^term>\<open>A - B\<close>,\quad \<^term>\<open>-A\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    89
\end{itemize}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    90
(where \<^term>\<open>A-B\<close> and \<open>-A\<close> are set difference and complement)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    91
and much more. \<^const>\<open>UNIV\<close> is the set of all elements of some type.
55348
366718f2ff85 indexed document
nipkow
parents: 55317
diff changeset
    92
Set comprehension\index{set comprehension} is written
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    93
\<^term>\<open>{x. P}\<close>\index{$IMP042@\<^term>\<open>{x. P}\<close>} rather than \<open>{x | P}\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    94
\begin{warn}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    95
In \<^term>\<open>{x. P}\<close> the \<open>x\<close> must be a variable. Set comprehension
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    96
involving a proper term \<open>t\<close> must be written
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    97
\noquotes{@{term[source] "{t | x y. P}"}}\index{$IMP042@\<open>{t |x. P}\<close>},
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    98
where \<open>x y\<close> are those free variables in \<open>t\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
    99
that occur in \<open>P\<close>.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   100
This is just a shorthand for \<^term>\<open>{v. \<exists>x y. v = t \<and> P}\<close>, where
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   101
\<open>v\<close> is a new variable. For example, \<^term>\<open>{x+y|x. x \<in> A}\<close>
49615
nipkow
parents: 48985
diff changeset
   102
is short for \noquotes{@{term[source]"{v. \<exists>x. v = x+y \<and> x \<in> A}"}}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   103
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   104
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   105
Here are the ASCII representations of the mathematical symbols:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   106
\begin{center}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   107
\begin{tabular}{l@ {\quad}l@ {\quad}l}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   108
\<open>\<in>\<close> & \texttt{\char`\\\char`\<in>} & \texttt{:}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   109
\<open>\<subseteq>\<close> & \texttt{\char`\\\char`\<subseteq>} & \texttt{<=}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   110
\<open>\<union>\<close> & \texttt{\char`\\\char`\<union>} & \texttt{Un}\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   111
\<open>\<inter>\<close> & \texttt{\char`\\\char`\<inter>} & \texttt{Int}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   112
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   113
\end{center}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   114
Sets also allow bounded quantifications \<^prop>\<open>\<forall>x \<in> A. P\<close> and
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   115
\<^prop>\<open>\<exists>x \<in> A. P\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   116
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   117
For the more ambitious, there are also \<open>\<Union>\<close>\index{$HOLSet6@\isasymUnion}
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   118
and \<open>\<Inter>\<close>\index{$HOLSet7@\isasymInter}:
52404
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   119
\begin{center}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   120
@{thm Union_eq} \qquad @{thm Inter_eq}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   121
\end{center}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   122
The ASCII forms of \<open>\<Union>\<close> are \texttt{\char`\\\char`\<Union>} and \texttt{Union},
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   123
those of \<open>\<Inter>\<close> are \texttt{\char`\\\char`\<Inter>} and \texttt{Inter}.
52404
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   124
There are also indexed unions and intersections:
51784
89fb9f4abf84 more funs
nipkow
parents: 51436
diff changeset
   125
\begin{center}
68800
nipkow
parents: 67613
diff changeset
   126
@{thm[eta_contract=false] UNION_eq} \\ @{thm[eta_contract=false] INTER_eq}
52404
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   127
\end{center}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   128
The ASCII forms are \ \texttt{UN x:A.~B} \ and \ \texttt{INT x:A. B} \
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   129
where \texttt{x} may occur in \texttt{B}.
58605
nipkow
parents: 58504
diff changeset
   130
If \texttt{A} is \texttt{UNIV} you can write \ \texttt{UN x.~B} \ and \ \texttt{INT x. B}.
52404
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   131
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   132
Some other frequently useful functions on sets are the following:
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   133
\begin{center}
33524382335b more set theory
nipkow
parents: 52361
diff changeset
   134
\begin{tabular}{l@ {\quad}l}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   135
@{const_typ set}\index{set@\<^const>\<open>set\<close>} & converts a list to the set of its elements\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   136
@{const_typ finite}\index{finite@\<^const>\<open>finite\<close>} & is true iff its argument is finite\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   137
\noquotes{@{term[source] "card :: 'a set \<Rightarrow> nat"}}\index{card@\<^const>\<open>card\<close>} & is the cardinality of a finite set\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   138
 & and is \<open>0\<close> for all infinite sets\\
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   139
@{thm image_def}\index{$IMP042@\<^term>\<open>f ` A\<close>} & is the image of a function over a set
51784
89fb9f4abf84 more funs
nipkow
parents: 51436
diff changeset
   140
\end{tabular}
89fb9f4abf84 more funs
nipkow
parents: 51436
diff changeset
   141
\end{center}
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 58605
diff changeset
   142
See @{cite "Nipkow-Main"} for the wealth of further predefined functions in theory
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   143
\<^theory>\<open>Main\<close>.
51784
89fb9f4abf84 more funs
nipkow
parents: 51436
diff changeset
   144
54214
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   145
54436
nipkow
parents: 54292
diff changeset
   146
\subsection*{Exercises}
54214
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   147
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   148
\exercise
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   149
Start from the data type of binary trees defined earlier:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   150
\<close>
54214
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   151
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   152
datatype 'a tree = Tip | Node "'a tree" 'a "'a tree"
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   153
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   154
text\<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   155
Define a function \<open>set ::\<close> \<^typ>\<open>'a tree \<Rightarrow> 'a set\<close>
54214
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   156
that returns the elements in a tree and a function
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   157
\<open>ord ::\<close> \<^typ>\<open>int tree \<Rightarrow> bool\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   158
that tests if an \<^typ>\<open>int tree\<close> is ordered.
54214
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   159
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   160
Define a function \<open>ins\<close> that inserts an element into an ordered \<^typ>\<open>int tree\<close>
54214
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   161
while maintaining the order of the tree. If the element is already in the tree, the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   162
same tree should be returned. Prove correctness of \<open>ins\<close>:
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   163
\<^prop>\<open>set(ins x t) = {x} \<union> set t\<close> and \<^prop>\<open>ord t \<Longrightarrow> ord(ins i t)\<close>.
54214
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   164
\endexercise
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   165
6d0688ce4ee2 more exercises
nipkow
parents: 54212
diff changeset
   166
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   167
\section{Proof Automation}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   168
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   169
So far we have only seen \<open>simp\<close> and \indexed{\<open>auto\<close>}{auto}: Both perform
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   170
rewriting, both can also prove linear arithmetic facts (no multiplication),
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   171
and \<open>auto\<close> is also able to prove simple logical or set-theoretic goals:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   172
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   173
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   174
lemma "\<forall>x. \<exists>y. x = y"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   175
by auto
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   176
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   177
lemma "A \<subseteq> B \<inter> C \<Longrightarrow> A \<subseteq> B \<union> C"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   178
by auto
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   179
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   180
text\<open>where
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   181
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   182
\isacom{by} \textit{proof-method}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   183
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   184
is short for
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   185
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   186
\isacom{apply} \textit{proof-method}\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   187
\isacom{done}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   188
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   189
The key characteristics of both \<open>simp\<close> and \<open>auto\<close> are
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   190
\begin{itemize}
59568
nipkow
parents: 58962
diff changeset
   191
\item They show you where they got stuck, giving you an idea how to continue.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   192
\item They perform the obvious steps but are highly incomplete.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   193
\end{itemize}
55317
834a84553e02 started index
nipkow
parents: 54839
diff changeset
   194
A proof method is \conceptnoidx{complete} if it can prove all true formulas.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   195
There is no complete proof method for HOL, not even in theory.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   196
Hence all our proof methods only differ in how incomplete they are.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   197
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   198
A proof method that is still incomplete but tries harder than \<open>auto\<close> is
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   199
\indexed{\<open>fastforce\<close>}{fastforce}.  It either succeeds or fails, it acts on the first
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   200
subgoal only, and it can be modified like \<open>auto\<close>, e.g.,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   201
with \<open>simp add\<close>. Here is a typical example of what \<open>fastforce\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   202
can do:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   203
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   204
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   205
lemma "\<lbrakk> \<forall>xs \<in> A. \<exists>ys. xs = ys @ ys;  us \<in> A \<rbrakk>
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   206
   \<Longrightarrow> \<exists>n. length us = n+n"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   207
by fastforce
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   208
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   209
text\<open>This lemma is out of reach for \<open>auto\<close> because of the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   210
quantifiers.  Even \<open>fastforce\<close> fails when the quantifier structure
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   211
becomes more complicated. In a few cases, its slow version \<open>force\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   212
succeeds where \<open>fastforce\<close> fails.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   213
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   214
The method of choice for complex logical goals is \indexed{\<open>blast\<close>}{blast}. In the
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   215
following example, \<open>T\<close> and \<open>A\<close> are two binary predicates. It
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   216
is shown that if \<open>T\<close> is total, \<open>A\<close> is antisymmetric and \<open>T\<close> is
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   217
a subset of \<open>A\<close>, then \<open>A\<close> is a subset of \<open>T\<close>:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   218
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   219
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   220
lemma
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   221
  "\<lbrakk> \<forall>x y. T x y \<or> T y x;
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   222
     \<forall>x y. A x y \<and> A y x \<longrightarrow> x = y;
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   223
     \<forall>x y. T x y \<longrightarrow> A x y \<rbrakk>
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   224
   \<Longrightarrow> \<forall>x y. A x y \<longrightarrow> T x y"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   225
by blast
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   226
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   227
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   228
We leave it to the reader to figure out why this lemma is true.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   229
Method \<open>blast\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   230
\begin{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   231
\item is (in principle) a complete proof procedure for first-order formulas,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   232
  a fragment of HOL. In practice there is a search bound.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   233
\item does no rewriting and knows very little about equality.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   234
\item covers logic, sets and relations.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   235
\item either succeeds or fails.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   236
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   237
Because of its strength in logic and sets and its weakness in equality reasoning, it complements the earlier proof methods.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   238
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   239
55348
366718f2ff85 indexed document
nipkow
parents: 55317
diff changeset
   240
\subsection{\concept{Sledgehammer}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   241
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   242
Command \isacom{sledgehammer} calls a number of external automatic
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   243
theorem provers (ATPs) that run for up to 30 seconds searching for a
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   244
proof. Some of these ATPs are part of the Isabelle installation, others are
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   245
queried over the internet. If successful, a proof command is generated and can
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   246
be inserted into your proof.  The biggest win of \isacom{sledgehammer} is
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   247
that it will take into account the whole lemma library and you do not need to
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   248
feed in any lemma explicitly. For example,\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   249
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   250
lemma "\<lbrakk> xs @ ys = ys @ xs;  length xs = length ys \<rbrakk> \<Longrightarrow> xs = ys"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   251
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   252
txt\<open>cannot be solved by any of the standard proof methods, but
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   253
\isacom{sledgehammer} finds the following proof:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   254
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   255
by (metis append_eq_conv_conj)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   256
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   257
text\<open>We do not explain how the proof was found but what this command
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   258
means. For a start, Isabelle does not trust external tools (and in particular
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   259
not the translations from Isabelle's logic to those tools!)
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   260
and insists on a proof that it can check. This is what \indexed{\<open>metis\<close>}{metis} does.
58605
nipkow
parents: 58504
diff changeset
   261
It is given a list of lemmas and tries to find a proof using just those lemmas
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   262
(and pure logic). In contrast to using \<open>simp\<close> and friends who know a lot of
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   263
lemmas already, using \<open>metis\<close> manually is tedious because one has
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   264
to find all the relevant lemmas first. But that is precisely what
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   265
\isacom{sledgehammer} does for us.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   266
In this case lemma @{thm[source]append_eq_conv_conj} alone suffices:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   267
@{thm[display] append_eq_conv_conj}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   268
We leave it to the reader to figure out why this lemma suffices to prove
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   269
the above lemma, even without any knowledge of what the functions \<^const>\<open>take\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   270
and \<^const>\<open>drop\<close> do. Keep in mind that the variables in the two lemmas
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   271
are independent of each other, despite the same names, and that you can
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   272
substitute arbitrary values for the free variables in a lemma.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   273
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   274
Just as for the other proof methods we have seen, there is no guarantee that
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   275
\isacom{sledgehammer} will find a proof if it exists. Nor is
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   276
\isacom{sledgehammer} superior to the other proof methods.  They are
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   277
incomparable. Therefore it is recommended to apply \<open>simp\<close> or \<open>auto\<close> before invoking \isacom{sledgehammer} on what is left.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   278
51436
790310525e97 tuned (in particular bold fonts)
nipkow
parents: 51433
diff changeset
   279
\subsection{Arithmetic}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   280
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   281
By arithmetic formulas we mean formulas involving variables, numbers, \<open>+\<close>, \<open>-\<close>, \<open>=\<close>, \<open><\<close>, \<open>\<le>\<close> and the usual logical
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   282
connectives \<open>\<not>\<close>, \<open>\<and>\<close>, \<open>\<or>\<close>, \<open>\<longrightarrow>\<close>,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   283
\<open>\<longleftrightarrow>\<close>. Strictly speaking, this is known as \concept{linear arithmetic}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   284
because it does not involve multiplication, although multiplication with
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   285
numbers, e.g., \<open>2*n\<close>, is allowed. Such formulas can be proved by
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   286
\indexed{\<open>arith\<close>}{arith}:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   287
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   288
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   289
lemma "\<lbrakk> (a::nat) \<le> x + b; 2*x < c \<rbrakk> \<Longrightarrow> 2*a + 1 \<le> 2*b + c"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   290
by arith
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   291
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   292
text\<open>In fact, \<open>auto\<close> and \<open>simp\<close> can prove many linear
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   293
arithmetic formulas already, like the one above, by calling a weak but fast
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   294
version of \<open>arith\<close>. Hence it is usually not necessary to invoke
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   295
\<open>arith\<close> explicitly.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   296
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   297
The above example involves natural numbers, but integers (type \<^typ>\<open>int\<close>)
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   298
and real numbers (type \<open>real\<close>) are supported as well. As are a number
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   299
of further operators like \<^const>\<open>min\<close> and \<^const>\<open>max\<close>. On \<^typ>\<open>nat\<close> and
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   300
\<^typ>\<open>int\<close>, \<open>arith\<close> can even prove theorems with quantifiers in them,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   301
but we will not enlarge on that here.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   302
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   303
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   304
\subsection{Trying Them All}
47727
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   305
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   306
If you want to try all of the above automatic proof methods you simply type
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   307
\begin{isabelle}
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   308
\isacom{try}
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   309
\end{isabelle}
63414
beb987127d0f new style dummy_pats
nipkow
parents: 62223
diff changeset
   310
There is also a lightweight variant \isacom{try0} that does not call
beb987127d0f new style dummy_pats
nipkow
parents: 62223
diff changeset
   311
sledgehammer. If desired, specific simplification and introduction rules
beb987127d0f new style dummy_pats
nipkow
parents: 62223
diff changeset
   312
can be added:
47727
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   313
\begin{isabelle}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   314
\isacom{try0} \<open>simp: \<dots> intro: \<dots>\<close>
47727
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   315
\end{isabelle}
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   316
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   317
\section{Single Step Proofs}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   318
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   319
Although automation is nice, it often fails, at least initially, and you need
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   320
to find out why. When \<open>fastforce\<close> or \<open>blast\<close> simply fail, you have
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   321
no clue why. At this point, the stepwise
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   322
application of proof rules may be necessary. For example, if \<open>blast\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   323
fails on \<^prop>\<open>A \<and> B\<close>, you want to attack the two
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   324
conjuncts \<open>A\<close> and \<open>B\<close> separately. This can
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   325
be achieved by applying \emph{conjunction introduction}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   326
\[ @{thm[mode=Rule,show_question_marks]conjI}\ \<open>conjI\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   327
\]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   328
to the proof state. We will now examine the details of this process.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   329
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   330
\subsection{Instantiating Unknowns}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   331
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   332
We had briefly mentioned earlier that after proving some theorem,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   333
Isabelle replaces all free variables \<open>x\<close> by so called \conceptidx{unknowns}{unknown}
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   334
\<open>?x\<close>. We can see this clearly in rule @{thm[source] conjI}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   335
These unknowns can later be instantiated explicitly or implicitly:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   336
\begin{itemize}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   337
\item By hand, using \indexed{\<open>of\<close>}{of}.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   338
The expression \<open>conjI[of "a=b" "False"]\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   339
instantiates the unknowns in @{thm[source] conjI} from left to right with the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   340
two formulas \<open>a=b\<close> and \<open>False\<close>, yielding the rule
63935
aa1fe1103ab8 raw control symbols are superseded by Latex.embed_raw;
wenzelm
parents: 63414
diff changeset
   341
@{thm[display,mode=Rule,margin=100]conjI[of "a=b" False]}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   342
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   343
In general, \<open>th[of string\<^sub>1 \<dots> string\<^sub>n]\<close> instantiates
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   344
the unknowns in the theorem \<open>th\<close> from left to right with the terms
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   345
\<open>string\<^sub>1\<close> to \<open>string\<^sub>n\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   346
55317
834a84553e02 started index
nipkow
parents: 54839
diff changeset
   347
\item By unification. \conceptidx{Unification}{unification} is the process of making two
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   348
terms syntactically equal by suitable instantiations of unknowns. For example,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   349
unifying \<open>?P \<and> ?Q\<close> with \mbox{\<^prop>\<open>a=b \<and> False\<close>} instantiates
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   350
\<open>?P\<close> with \<^prop>\<open>a=b\<close> and \<open>?Q\<close> with \<^prop>\<open>False\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   351
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   352
We need not instantiate all unknowns. If we want to skip a particular one we
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   353
can write \<open>_\<close> instead, for example \<open>conjI[of _ "False"]\<close>.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   354
Unknowns can also be instantiated by name using \indexed{\<open>where\<close>}{where}, for example
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   355
\<open>conjI[where ?P = "a=b"\<close> \isacom{and} \<open>?Q = "False"]\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   356
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   357
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   358
\subsection{Rule Application}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   359
55317
834a84553e02 started index
nipkow
parents: 54839
diff changeset
   360
\conceptidx{Rule application}{rule application} means applying a rule backwards to a proof state.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   361
For example, applying rule @{thm[source]conjI} to a proof state
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   362
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   363
\<open>1.  \<dots>  \<Longrightarrow> A \<and> B\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   364
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   365
results in two subgoals, one for each premise of @{thm[source]conjI}:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   366
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   367
\<open>1.  \<dots>  \<Longrightarrow> A\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   368
\<open>2.  \<dots>  \<Longrightarrow> B\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   369
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   370
In general, the application of a rule \<open>\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   371
to a subgoal \mbox{\<open>\<dots> \<Longrightarrow> C\<close>} proceeds in two steps:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   372
\begin{enumerate}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   373
\item
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   374
Unify \<open>A\<close> and \<open>C\<close>, thus instantiating the unknowns in the rule.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   375
\item
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   376
Replace the subgoal \<open>C\<close> with \<open>n\<close> new subgoals \<open>A\<^sub>1\<close> to \<open>A\<^sub>n\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   377
\end{enumerate}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   378
This is the command to apply rule \<open>xyz\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   379
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   380
\isacom{apply}\<open>(rule xyz)\<close>\index{rule@\<open>rule\<close>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   381
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   382
This is also called \concept{backchaining} with rule \<open>xyz\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   383
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   384
\subsection{Introduction Rules}
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
Conjunction introduction (@{thm[source] conjI}) is one example of a whole
55317
834a84553e02 started index
nipkow
parents: 54839
diff changeset
   387
class of rules known as \conceptidx{introduction rules}{introduction rule}. They explain under which
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   388
premises some logical construct can be introduced. Here are some further
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   389
useful introduction rules:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   390
\[
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   391
\inferrule*[right=\mbox{\<open>impI\<close>}]{\mbox{\<open>?P \<Longrightarrow> ?Q\<close>}}{\mbox{\<open>?P \<longrightarrow> ?Q\<close>}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   392
\qquad
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   393
\inferrule*[right=\mbox{\<open>allI\<close>}]{\mbox{\<open>\<And>x. ?P x\<close>}}{\mbox{\<open>\<forall>x. ?P x\<close>}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   394
\]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   395
\[
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   396
\inferrule*[right=\mbox{\<open>iffI\<close>}]{\mbox{\<open>?P \<Longrightarrow> ?Q\<close>} \\ \mbox{\<open>?Q \<Longrightarrow> ?P\<close>}}
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   397
  {\mbox{\<open>?P = ?Q\<close>}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   398
\]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   399
These rules are part of the logical system of \concept{natural deduction}
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 58605
diff changeset
   400
(e.g., @{cite HuthRyan}). Although we intentionally de-emphasize the basic rules
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   401
of logic in favour of automatic proof methods that allow you to take bigger
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   402
steps, these rules are helpful in locating where and why automation fails.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   403
When applied backwards, these rules decompose the goal:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   404
\begin{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   405
\item @{thm[source] conjI} and @{thm[source]iffI} split the goal into two subgoals,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   406
\item @{thm[source] impI} moves the left-hand side of a HOL implication into the list of assumptions,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   407
\item and @{thm[source] allI} removes a \<open>\<forall>\<close> by turning the quantified variable into a fixed local variable of the subgoal.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   408
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   409
Isabelle knows about these and a number of other introduction rules.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   410
The command
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   411
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   412
\isacom{apply} \<open>rule\<close>\index{rule@\<open>rule\<close>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   413
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   414
automatically selects the appropriate rule for the current subgoal.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   415
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   416
You can also turn your own theorems into introduction rules by giving them
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   417
the \indexed{\<open>intro\<close>}{intro} attribute, analogous to the \<open>simp\<close> attribute.  In
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   418
that case \<open>blast\<close>, \<open>fastforce\<close> and (to a limited extent) \<open>auto\<close> will automatically backchain with those theorems. The \<open>intro\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   419
attribute should be used with care because it increases the search space and
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   420
can lead to nontermination.  Sometimes it is better to use it only in
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   421
specific calls of \<open>blast\<close> and friends. For example,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   422
@{thm[source] le_trans}, transitivity of \<open>\<le>\<close> on type \<^typ>\<open>nat\<close>,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   423
is not an introduction rule by default because of the disastrous effect
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   424
on the search space, but can be useful in specific situations:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   425
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   426
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   427
lemma "\<lbrakk> (a::nat) \<le> b; b \<le> c; c \<le> d; d \<le> e \<rbrakk> \<Longrightarrow> a \<le> e"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   428
by(blast intro: le_trans)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   429
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   430
text\<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   431
Of course this is just an example and could be proved by \<open>arith\<close>, too.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   432
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   433
\subsection{Forward Proof}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   434
\label{sec:forward-proof}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   435
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   436
Forward proof means deriving new theorems from old theorems. We have already
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   437
seen a very simple form of forward proof: the \<open>of\<close> operator for
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   438
instantiating unknowns in a theorem. The big brother of \<open>of\<close> is
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   439
\indexed{\<open>OF\<close>}{OF} for applying one theorem to others. Given a theorem \<^prop>\<open>A \<Longrightarrow> B\<close> called
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   440
\<open>r\<close> and a theorem \<open>A'\<close> called \<open>r'\<close>, the theorem \<open>r[OF r']\<close> is the result of applying \<open>r\<close> to \<open>r'\<close>, where \<open>r\<close> should be viewed as a function taking a theorem \<open>A\<close> and returning
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   441
\<open>B\<close>.  More precisely, \<open>A\<close> and \<open>A'\<close> are unified, thus
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   442
instantiating the unknowns in \<open>B\<close>, and the result is the instantiated
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   443
\<open>B\<close>. Of course, unification may also fail.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   444
\begin{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   445
Application of rules to other rules operates in the forward direction: from
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   446
the premises to the conclusion of the rule; application of rules to proof
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   447
states operates in the backward direction, from the conclusion to the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   448
premises.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   449
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   450
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   451
In general \<open>r\<close> can be of the form \<open>\<lbrakk> A\<^sub>1; \<dots>; A\<^sub>n \<rbrakk> \<Longrightarrow> A\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   452
and there can be multiple argument theorems \<open>r\<^sub>1\<close> to \<open>r\<^sub>m\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   453
(with \<open>m \<le> n\<close>), in which case \<open>r[OF r\<^sub>1 \<dots> r\<^sub>m]\<close> is obtained
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   454
by unifying and thus proving \<open>A\<^sub>i\<close> with \<open>r\<^sub>i\<close>, \<open>i = 1\<dots>m\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   455
Here is an example, where @{thm[source]refl} is the theorem
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   456
@{thm[show_question_marks] refl}:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   457
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   458
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   459
thm conjI[OF refl[of "a"] refl[of "b"]]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   460
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   461
text\<open>yields the theorem @{thm conjI[OF refl[of "a"] refl[of "b"]]}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   462
The command \isacom{thm} merely displays the result.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   463
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   464
Forward reasoning also makes sense in connection with proof states.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   465
Therefore \<open>blast\<close>, \<open>fastforce\<close> and \<open>auto\<close> support a modifier
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   466
\<open>dest\<close> which instructs the proof method to use certain rules in a
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   467
forward fashion. If \<open>r\<close> is of the form \mbox{\<open>A \<Longrightarrow> B\<close>}, the modifier
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   468
\mbox{\<open>dest: r\<close>}\index{dest@\<open>dest:\<close>}
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   469
allows proof search to reason forward with \<open>r\<close>, i.e.,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   470
to replace an assumption \<open>A'\<close>, where \<open>A'\<close> unifies with \<open>A\<close>,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   471
with the correspondingly instantiated \<open>B\<close>. For example, @{thm[source,show_question_marks] Suc_leD} is the theorem \mbox{@{thm Suc_leD}}, which works well for forward reasoning:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   472
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   473
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   474
lemma "Suc(Suc(Suc a)) \<le> b \<Longrightarrow> a \<le> b"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   475
by(blast dest: Suc_leD)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   476
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   477
text\<open>In this particular example we could have backchained with
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   478
@{thm[source] Suc_leD}, too, but because the premise is more complicated than the conclusion this can easily lead to nontermination.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   479
54212
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   480
%\subsection{Finding Theorems}
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   481
%
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   482
%Command \isacom{find{\isacharunderscorekeyword}theorems} searches for specific theorems in the current
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   483
%theory. Search criteria include pattern matching on terms and on names.
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 58605
diff changeset
   484
%For details see the Isabelle/Isar Reference Manual~@{cite IsarRef}.
54212
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   485
%\bigskip
47727
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   486
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   487
\begin{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   488
To ease readability we will drop the question marks
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   489
in front of unknowns from now on.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   490
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   491
47727
027c7f8cef22 doc update
nipkow
parents: 47720
diff changeset
   492
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   493
\section{Inductive Definitions}
55361
d459a63ca42f more indexing
nipkow
parents: 55348
diff changeset
   494
\label{sec:inductive-defs}\index{inductive definition|(}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   495
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   496
Inductive definitions are the third important definition facility, after
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   497
datatypes and recursive function.
52782
b11d73dbfb76 tuned intro
nipkow
parents: 52404
diff changeset
   498
\ifsem
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   499
In fact, they are the key construct in the
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   500
definition of operational semantics in the second part of the book.
52782
b11d73dbfb76 tuned intro
nipkow
parents: 52404
diff changeset
   501
\fi
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   502
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   503
\subsection{An Example: Even Numbers}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   504
\label{sec:Logic:even}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   505
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   506
Here is a simple example of an inductively defined predicate:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   507
\begin{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   508
\item 0 is even
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   509
\item If $n$ is even, so is $n+2$.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   510
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   511
The operative word ``inductive'' means that these are the only even numbers.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   512
In Isabelle we give the two rules the names \<open>ev0\<close> and \<open>evSS\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   513
and write
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   514
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   515
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   516
inductive ev :: "nat \<Rightarrow> bool" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   517
ev0:    "ev 0" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   518
evSS:  (*<*)"ev n \<Longrightarrow> ev (Suc(Suc n))"(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   519
text_raw\<open>@{prop[source]"ev n \<Longrightarrow> ev (n + 2)"}\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   520
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   521
text\<open>To get used to inductive definitions, we will first prove a few
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   522
properties of \<^const>\<open>ev\<close> informally before we descend to the Isabelle level.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   523
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   524
How do we prove that some number is even, e.g., \<^prop>\<open>ev 4\<close>? Simply by combining the defining rules for \<^const>\<open>ev\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   525
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   526
\<open>ev 0 \<Longrightarrow> ev (0 + 2) \<Longrightarrow> ev((0 + 2) + 2) = ev 4\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   527
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   528
55361
d459a63ca42f more indexing
nipkow
parents: 55348
diff changeset
   529
\subsubsection{Rule Induction}\index{rule induction|(}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   530
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   531
Showing that all even numbers have some property is more complicated.  For
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   532
example, let us prove that the inductive definition of even numbers agrees
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   533
with the following recursive one:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   534
58962
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   535
fun evn :: "nat \<Rightarrow> bool" where
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   536
"evn 0 = True" |
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   537
"evn (Suc 0) = False" |
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   538
"evn (Suc(Suc n)) = evn n"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   539
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   540
text\<open>We prove \<^prop>\<open>ev m \<Longrightarrow> evn m\<close>.  That is, we
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   541
assume \<^prop>\<open>ev m\<close> and by induction on the form of its derivation
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   542
prove \<^prop>\<open>evn m\<close>. There are two cases corresponding to the two rules
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   543
for \<^const>\<open>ev\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   544
\begin{description}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   545
\item[Case @{thm[source]ev0}:]
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   546
 \<^prop>\<open>ev m\<close> was derived by rule \<^prop>\<open>ev 0\<close>: \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   547
 \<open>\<Longrightarrow>\<close> \<^prop>\<open>m=(0::nat)\<close> \<open>\<Longrightarrow>\<close> \<open>evn m = evn 0 = True\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   548
\item[Case @{thm[source]evSS}:]
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   549
 \<^prop>\<open>ev m\<close> was derived by rule \<^prop>\<open>ev n \<Longrightarrow> ev(n+2)\<close>: \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   550
\<open>\<Longrightarrow>\<close> \<^prop>\<open>m=n+(2::nat)\<close> and by induction hypothesis \<^prop>\<open>evn n\<close>\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   551
\<open>\<Longrightarrow>\<close> \<open>evn m = evn(n + 2) = evn n = True\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   552
\end{description}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   553
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   554
What we have just seen is a special case of \concept{rule induction}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   555
Rule induction applies to propositions of this form
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   556
\begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   557
\<^prop>\<open>ev n \<Longrightarrow> P n\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   558
\end{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   559
That is, we want to prove a property \<^prop>\<open>P n\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   560
for all even \<open>n\<close>. But if we assume \<^prop>\<open>ev n\<close>, then there must be
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   561
some derivation of this assumption using the two defining rules for
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   562
\<^const>\<open>ev\<close>. That is, we must prove
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   563
\begin{description}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   564
\item[Case @{thm[source]ev0}:] \<^prop>\<open>P(0::nat)\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   565
\item[Case @{thm[source]evSS}:] \<^prop>\<open>\<lbrakk> ev n; P n \<rbrakk> \<Longrightarrow> P(n + 2::nat)\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   566
\end{description}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   567
The corresponding rule is called @{thm[source] ev.induct} and looks like this:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   568
\[
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   569
\inferrule{
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   570
\mbox{@{thm (prem 1) ev.induct[of "n"]}}\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   571
\mbox{@{thm (prem 2) ev.induct}}\\
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   572
\mbox{\<^prop>\<open>!!n. \<lbrakk> ev n; P n \<rbrakk> \<Longrightarrow> P(n+2)\<close>}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   573
{\mbox{@{thm (concl) ev.induct[of "n"]}}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   574
\]
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   575
The first premise \<^prop>\<open>ev n\<close> enforces that this rule can only be applied
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   576
in situations where we know that \<open>n\<close> is even.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   577
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   578
Note that in the induction step we may not just assume \<^prop>\<open>P n\<close> but also
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   579
\mbox{\<^prop>\<open>ev n\<close>}, which is simply the premise of rule @{thm[source]
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   580
evSS}.  Here is an example where the local assumption \<^prop>\<open>ev n\<close> comes in
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   581
handy: we prove \<^prop>\<open>ev m \<Longrightarrow> ev(m - 2)\<close> by induction on \<^prop>\<open>ev m\<close>.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   582
Case @{thm[source]ev0} requires us to prove \<^prop>\<open>ev(0 - 2)\<close>, which follows
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   583
from \<^prop>\<open>ev 0\<close> because \<^prop>\<open>0 - 2 = (0::nat)\<close> on type \<^typ>\<open>nat\<close>. In
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   584
case @{thm[source]evSS} we have \mbox{\<^prop>\<open>m = n+(2::nat)\<close>} and may assume
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   585
\<^prop>\<open>ev n\<close>, which implies \<^prop>\<open>ev (m - 2)\<close> because \<open>m - 2 = (n +
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   586
2) - 2 = n\<close>. We did not need the induction hypothesis at all for this proof (it
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   587
is just a case analysis of which rule was used) but having \<^prop>\<open>ev n\<close>
56989
nipkow
parents: 56451
diff changeset
   588
at our disposal in case @{thm[source]evSS} was essential.
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   589
This case analysis of rules is also called ``rule inversion''
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   590
and is discussed in more detail in \autoref{ch:Isar}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   591
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   592
\subsubsection{In Isabelle}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   593
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   594
Let us now recast the above informal proofs in Isabelle. For a start,
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   595
we use \<^const>\<open>Suc\<close> terms instead of numerals in rule @{thm[source]evSS}:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   596
@{thm[display] evSS}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   597
This avoids the difficulty of unifying \<open>n+2\<close> with some numeral,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   598
which is not automatic.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   599
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   600
The simplest way to prove \<^prop>\<open>ev(Suc(Suc(Suc(Suc 0))))\<close> is in a forward
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   601
direction: \<open>evSS[OF evSS[OF ev0]]\<close> yields the theorem @{thm evSS[OF
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   602
evSS[OF ev0]]}. Alternatively, you can also prove it as a lemma in backwards
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   603
fashion. Although this is more verbose, it allows us to demonstrate how each
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   604
rule application changes the proof state:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   605
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   606
lemma "ev(Suc(Suc(Suc(Suc 0))))"
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   607
txt\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   608
@{subgoals[display,indent=0,goals_limit=1]}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   609
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   610
apply(rule evSS)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   611
txt\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   612
@{subgoals[display,indent=0,goals_limit=1]}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   613
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   614
apply(rule evSS)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   615
txt\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   616
@{subgoals[display,indent=0,goals_limit=1]}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   617
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   618
apply(rule ev0)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   619
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   620
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   621
text\<open>\indent
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   622
Rule induction is applied by giving the induction rule explicitly via the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   623
\<open>rule:\<close> modifier:\index{inductionrule@\<open>induction ... rule:\<close>}\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   624
58962
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   625
lemma "ev m \<Longrightarrow> evn m"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   626
apply(induction rule: ev.induct)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   627
by(simp_all)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   628
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   629
text\<open>Both cases are automatic. Note that if there are multiple assumptions
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   630
of the form \<^prop>\<open>ev t\<close>, method \<open>induction\<close> will induct on the leftmost
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   631
one.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   632
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   633
As a bonus, we also prove the remaining direction of the equivalence of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   634
\<^const>\<open>ev\<close> and \<^const>\<open>evn\<close>:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   635
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   636
58962
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   637
lemma "evn n \<Longrightarrow> ev n"
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   638
apply(induction n rule: evn.induct)
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   639
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   640
txt\<open>This is a proof by computation induction on \<open>n\<close> (see
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   641
\autoref{sec:recursive-funs}) that sets up three subgoals corresponding to
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   642
the three equations for \<^const>\<open>evn\<close>:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   643
@{subgoals[display,indent=0]}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   644
The first and third subgoals follow with @{thm[source]ev0} and @{thm[source]evSS}, and the second subgoal is trivially true because \<^prop>\<open>evn(Suc 0)\<close> is \<^const>\<open>False\<close>:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
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
by (simp_all add: ev0 evSS)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   648
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   649
text\<open>The rules for \<^const>\<open>ev\<close> make perfect simplification and introduction
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   650
rules because their premises are always smaller than the conclusion. It
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   651
makes sense to turn them into simplification and introduction rules
55348
366718f2ff85 indexed document
nipkow
parents: 55317
diff changeset
   652
permanently, to enhance proof automation. They are named @{thm[source] ev.intros}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   653
\index{intros@\<open>.intros\<close>} by Isabelle:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   654
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   655
declare ev.intros[simp,intro]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   656
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   657
text\<open>The rules of an inductive definition are not simplification rules by
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   658
default because, in contrast to recursive functions, there is no termination
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   659
requirement for inductive definitions.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   660
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   661
\subsubsection{Inductive Versus Recursive}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   662
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   663
We have seen two definitions of the notion of evenness, an inductive and a
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   664
recursive one. Which one is better? Much of the time, the recursive one is
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   665
more convenient: it allows us to do rewriting in the middle of terms, and it
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   666
expresses both the positive information (which numbers are even) and the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   667
negative information (which numbers are not even) directly. An inductive
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   668
definition only expresses the positive information directly. The negative
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   669
information, for example, that \<open>1\<close> is not even, has to be proved from
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   670
it (by induction or rule inversion). On the other hand, rule induction is
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   671
tailor-made for proving \mbox{\<^prop>\<open>ev n \<Longrightarrow> P n\<close>} because it only asks you
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   672
to prove the positive cases. In the proof of \<^prop>\<open>evn n \<Longrightarrow> P n\<close> by
58962
e3491acee50f even -> evn because even is now in Main
nipkow
parents: 58620
diff changeset
   673
computation induction via @{thm[source]evn.induct}, we are also presented
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   674
with the trivial negative cases. If you want the convenience of both
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   675
rewriting and rule induction, you can make two definitions and show their
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   676
equivalence (as above) or make one definition and prove additional properties
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   677
from it, for example rule induction from computation induction.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   678
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   679
But many concepts do not admit a recursive definition at all because there is
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   680
no datatype for the recursion (for example, the transitive closure of a
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   681
relation), or the recursion would not terminate (for example,
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   682
an interpreter for a programming language). Even if there is a recursive
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   683
definition, if we are only interested in the positive information, the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   684
inductive definition may be much simpler.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   685
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   686
\subsection{The Reflexive Transitive Closure}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   687
\label{sec:star}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   688
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   689
Evenness is really more conveniently expressed recursively than inductively.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   690
As a second and very typical example of an inductive definition we define the
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   691
reflexive transitive closure.
52782
b11d73dbfb76 tuned intro
nipkow
parents: 52404
diff changeset
   692
\ifsem
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   693
It will also be an important building block for
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   694
some of the semantics considered in the second part of the book.
52782
b11d73dbfb76 tuned intro
nipkow
parents: 52404
diff changeset
   695
\fi
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   696
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   697
The reflexive transitive closure, called \<open>star\<close> below, is a function
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   698
that maps a binary predicate to another binary predicate: if \<open>r\<close> is of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   699
type \<open>\<tau> \<Rightarrow> \<tau> \<Rightarrow> bool\<close> then \<^term>\<open>star r\<close> is again of type \<open>\<tau> \<Rightarrow>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   700
\<tau> \<Rightarrow> bool\<close>, and \<^prop>\<open>star r x y\<close> means that \<open>x\<close> and \<open>y\<close> are in
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   701
the relation \<^term>\<open>star r\<close>. Think \<^term>\<open>r\<^sup>*\<close> when you see \<^term>\<open>star
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   702
r\<close>, because \<open>star r\<close> is meant to be the reflexive transitive closure.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   703
That is, \<^prop>\<open>star r x y\<close> is meant to be true if from \<open>x\<close> we can
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   704
reach \<open>y\<close> in finitely many \<open>r\<close> steps. This concept is naturally
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   705
defined inductively:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   706
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   707
inductive star :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"  for r where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   708
refl:  "star r x x" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   709
step:  "r x y \<Longrightarrow> star r y z \<Longrightarrow> star r x z"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   710
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   711
text\<open>The base case @{thm[source] refl} is reflexivity: \<^term>\<open>x=y\<close>. The
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   712
step case @{thm[source]step} combines an \<open>r\<close> step (from \<open>x\<close> to
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   713
\<open>y\<close>) and a \<^term>\<open>star r\<close> step (from \<open>y\<close> to \<open>z\<close>) into a
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   714
\<^term>\<open>star r\<close> step (from \<open>x\<close> to \<open>z\<close>).
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   715
The ``\isacom{for}~\<open>r\<close>'' in the header is merely a hint to Isabelle
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   716
that \<open>r\<close> is a fixed parameter of \<^const>\<open>star\<close>, in contrast to the
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   717
further parameters of \<^const>\<open>star\<close>, which change. As a result, Isabelle
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   718
generates a simpler induction rule.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   719
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   720
By definition \<^term>\<open>star r\<close> is reflexive. It is also transitive, but we
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   721
need rule induction to prove that:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   722
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   723
lemma star_trans: "star r x y \<Longrightarrow> star r y z \<Longrightarrow> star r x z"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   724
apply(induction rule: star.induct)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   725
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   726
defer
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   727
apply(rename_tac u x y)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   728
defer
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   729
(*>*)
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   730
txt\<open>The induction is over \<^prop>\<open>star r x y\<close> (the first matching assumption)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   731
and we try to prove \mbox{\<^prop>\<open>star r y z \<Longrightarrow> star r x z\<close>},
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   732
which we abbreviate by \<^prop>\<open>P x y\<close>. These are our two subgoals:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   733
@{subgoals[display,indent=0]}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   734
The first one is \<^prop>\<open>P x x\<close>, the result of case @{thm[source]refl},
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   735
and it is trivial:\index{assumption@\<open>assumption\<close>}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   736
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   737
apply(assumption)
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   738
txt\<open>Let us examine subgoal \<open>2\<close>, case @{thm[source] step}.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   739
Assumptions \<^prop>\<open>r u x\<close> and \mbox{\<^prop>\<open>star r x y\<close>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   740
are the premises of rule @{thm[source]step}.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   741
Assumption \<^prop>\<open>star r y z \<Longrightarrow> star r x z\<close> is \mbox{\<^prop>\<open>P x y\<close>},
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   742
the IH coming from \<^prop>\<open>star r x y\<close>. We have to prove \<^prop>\<open>P u y\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   743
which we do by assuming \<^prop>\<open>star r y z\<close> and proving \<^prop>\<open>star r u z\<close>.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   744
The proof itself is straightforward: from \mbox{\<^prop>\<open>star r y z\<close>} the IH
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   745
leads to \<^prop>\<open>star r x z\<close> which, together with \<^prop>\<open>r u x\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   746
leads to \mbox{\<^prop>\<open>star r u z\<close>} via rule @{thm[source]step}:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   747
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   748
apply(metis step)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   749
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   750
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   751
text\<open>\index{rule induction|)}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   752
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 51784
diff changeset
   753
\subsection{The General Case}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   754
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   755
Inductive definitions have approximately the following general form:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   756
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   757
\isacom{inductive} \<open>I :: "\<tau> \<Rightarrow> bool"\<close> \isacom{where}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   758
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   759
followed by a sequence of (possibly named) rules of the form
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   760
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   761
\<open>\<lbrakk> I a\<^sub>1; \<dots>; I a\<^sub>n \<rbrakk> \<Longrightarrow> I a\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   762
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   763
separated by \<open>|\<close>. As usual, \<open>n\<close> can be 0.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   764
The corresponding rule induction principle
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   765
\<open>I.induct\<close> applies to propositions of the form
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   766
\begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   767
\<^prop>\<open>I x \<Longrightarrow> P x\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   768
\end{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   769
where \<open>P\<close> may itself be a chain of implications.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   770
\begin{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   771
Rule induction is always on the leftmost premise of the goal.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   772
Hence \<open>I x\<close> must be the first premise.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   773
\end{warn}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   774
Proving \<^prop>\<open>I x \<Longrightarrow> P x\<close> by rule induction means proving
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   775
for every rule of \<open>I\<close> that \<open>P\<close> is invariant:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   776
\begin{quote}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   777
\<open>\<lbrakk> I a\<^sub>1; P a\<^sub>1; \<dots>; I a\<^sub>n; P a\<^sub>n \<rbrakk> \<Longrightarrow> P a\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   778
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   779
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   780
The above format for inductive definitions is simplified in a number of
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   781
respects. \<open>I\<close> can have any number of arguments and each rule can have
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   782
additional premises not involving \<open>I\<close>, so-called \conceptidx{side
55317
834a84553e02 started index
nipkow
parents: 54839
diff changeset
   783
conditions}{side condition}. In rule inductions, these side conditions appear as additional
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   784
assumptions. The \isacom{for} clause seen in the definition of the reflexive
55361
d459a63ca42f more indexing
nipkow
parents: 55348
diff changeset
   785
transitive closure simplifies the induction rule.
d459a63ca42f more indexing
nipkow
parents: 55348
diff changeset
   786
\index{inductive definition|)}
54231
2975658d49cd more exercises
nipkow
parents: 54218
diff changeset
   787
54436
nipkow
parents: 54292
diff changeset
   788
\subsection*{Exercises}
54186
ea92cce67f09 added exercise
nipkow
parents: 53015
diff changeset
   789
54231
2975658d49cd more exercises
nipkow
parents: 54218
diff changeset
   790
\begin{exercise}
58502
nipkow
parents: 56989
diff changeset
   791
Formalize the following definition of palindromes
54231
2975658d49cd more exercises
nipkow
parents: 54218
diff changeset
   792
\begin{itemize}
2975658d49cd more exercises
nipkow
parents: 54218
diff changeset
   793
\item The empty list and a singleton list are palindromes.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   794
\item If \<open>xs\<close> is a palindrome, so is \<^term>\<open>a # xs @ [a]\<close>.
54231
2975658d49cd more exercises
nipkow
parents: 54218
diff changeset
   795
\end{itemize}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   796
as an inductive predicate \<open>palindrome ::\<close> \<^typ>\<open>'a list \<Rightarrow> bool\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   797
and prove that \<^prop>\<open>rev xs = xs\<close> if \<open>xs\<close> is a palindrome.
54231
2975658d49cd more exercises
nipkow
parents: 54218
diff changeset
   798
\end{exercise}
2975658d49cd more exercises
nipkow
parents: 54218
diff changeset
   799
54212
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   800
\exercise
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   801
We could also have defined \<^const>\<open>star\<close> as follows:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   802
\<close>
54212
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   803
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   804
inductive star' :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" for r where
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   805
refl': "star' r x x" |
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   806
step': "star' r x y \<Longrightarrow> r y z \<Longrightarrow> star' r x z"
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   807
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   808
text\<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   809
The single \<open>r\<close> step is performed after rather than before the \<open>star'\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   810
steps. Prove \<^prop>\<open>star' r x y \<Longrightarrow> star r x y\<close> and
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   811
\<^prop>\<open>star r x y \<Longrightarrow> star' r x y\<close>. You may need lemmas.
54217
2fa338fd0a28 tuned text
nipkow
parents: 54214
diff changeset
   812
Note that rule induction fails
2fa338fd0a28 tuned text
nipkow
parents: 54214
diff changeset
   813
if the assumption about the inductive predicate is not the first assumption.
54212
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   814
\endexercise
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   815
54292
ce4a17b2e373 more exercises
nipkow
parents: 54290
diff changeset
   816
\begin{exercise}\label{exe:iter}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   817
Analogous to \<^const>\<open>star\<close>, give an inductive definition of the \<open>n\<close>-fold iteration
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   818
of a relation \<open>r\<close>: \<^term>\<open>iter r n x y\<close> should hold if there are \<open>x\<^sub>0\<close>, \dots, \<open>x\<^sub>n\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   819
such that \<^prop>\<open>x = x\<^sub>0\<close>, \<^prop>\<open>x\<^sub>n = y\<close> and \<open>r x\<^bsub>i\<^esub> x\<^bsub>i+1\<^esub>\<close> for
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   820
all \<^prop>\<open>i < n\<close>. Correct and prove the following claim:
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   821
\<^prop>\<open>star r x y \<Longrightarrow> iter r n x y\<close>.
54290
fee1276d47f7 added exercise
nipkow
parents: 54231
diff changeset
   822
\end{exercise}
fee1276d47f7 added exercise
nipkow
parents: 54231
diff changeset
   823
61012
40a0a4077126 nex exercise
nipkow
parents: 60605
diff changeset
   824
\begin{exercise}\label{exe:cfg}
54218
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   825
A context-free grammar can be seen as an inductive definition where each
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   826
nonterminal $A$ is an inductively defined predicate on lists of terminal
56116
nipkow
parents: 55361
diff changeset
   827
symbols: $A(w)$ means that $w$ is in the language generated by $A$.
54218
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   828
For example, the production $S \to a S b$ can be viewed as the implication
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   829
\<^prop>\<open>S w \<Longrightarrow> S (a # w @ [b])\<close> where \<open>a\<close> and \<open>b\<close> are terminal symbols,
54218
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   830
i.e., elements of some alphabet. The alphabet can be defined like this:
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   831
\isacom{datatype} \<open>alpha = a | b | \<dots>\<close>
54218
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   832
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   833
Define the two grammars (where $\varepsilon$ is the empty word)
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   834
\[
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   835
\begin{array}{r@ {\quad}c@ {\quad}l}
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   836
S &\to& \varepsilon \quad\mid\quad aSb \quad\mid\quad SS \\
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   837
T &\to& \varepsilon \quad\mid\quad TaTb
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   838
\end{array}
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   839
\]
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   840
as two inductive predicates.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   841
If you think of \<open>a\<close> and \<open>b\<close> as ``\<open>(\<close>'' and  ``\<open>)\<close>'',
56989
nipkow
parents: 56451
diff changeset
   842
the grammar defines strings of balanced parentheses.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   843
Prove \<^prop>\<open>T w \<Longrightarrow> S w\<close> and \mbox{\<^prop>\<open>S w \<Longrightarrow> T w\<close>} separately and conclude
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   844
\<^prop>\<open>S w = T w\<close>.
54218
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   845
\end{exercise}
07c0c121a8dc more exercises
nipkow
parents: 54217
diff changeset
   846
54212
5f1e2017eeea more exercises
nipkow
parents: 54210
diff changeset
   847
\ifsem
54186
ea92cce67f09 added exercise
nipkow
parents: 53015
diff changeset
   848
\begin{exercise}
54203
4d3a481fc48e more exercises
nipkow
parents: 54186
diff changeset
   849
In \autoref{sec:AExp} we defined a recursive evaluation function
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   850
\<open>aval :: aexp \<Rightarrow> state \<Rightarrow> val\<close>.
54203
4d3a481fc48e more exercises
nipkow
parents: 54186
diff changeset
   851
Define an inductive evaluation predicate
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   852
\<open>aval_rel :: aexp \<Rightarrow> state \<Rightarrow> val \<Rightarrow> bool\<close>
54203
4d3a481fc48e more exercises
nipkow
parents: 54186
diff changeset
   853
and prove that it agrees with the recursive function:
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   854
\<^prop>\<open>aval_rel a s v \<Longrightarrow> aval a s = v\<close>, 
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   855
\<^prop>\<open>aval a s = v \<Longrightarrow> aval_rel a s v\<close> and thus
54203
4d3a481fc48e more exercises
nipkow
parents: 54186
diff changeset
   856
\noquotes{@{prop [source] "aval_rel a s v \<longleftrightarrow> aval a s = v"}}.
4d3a481fc48e more exercises
nipkow
parents: 54186
diff changeset
   857
\end{exercise}
4d3a481fc48e more exercises
nipkow
parents: 54186
diff changeset
   858
4d3a481fc48e more exercises
nipkow
parents: 54186
diff changeset
   859
\begin{exercise}
54210
9d239afc1a90 more exercises
nipkow
parents: 54203
diff changeset
   860
Consider the stack machine from Chapter~3
9d239afc1a90 more exercises
nipkow
parents: 54203
diff changeset
   861
and recall the concept of \concept{stack underflow}
9d239afc1a90 more exercises
nipkow
parents: 54203
diff changeset
   862
from Exercise~\ref{exe:stack-underflow}.
9d239afc1a90 more exercises
nipkow
parents: 54203
diff changeset
   863
Define an inductive predicate
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   864
\<open>ok :: nat \<Rightarrow> instr list \<Rightarrow> nat \<Rightarrow> bool\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   865
such that \<open>ok n is n'\<close> means that with any initial stack of length
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   866
\<open>n\<close> the instructions \<open>is\<close> can be executed
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   867
without stack underflow and that the final stack has length \<open>n'\<close>.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   868
Prove that \<open>ok\<close> correctly computes the final stack size
54186
ea92cce67f09 added exercise
nipkow
parents: 53015
diff changeset
   869
@{prop[display] "\<lbrakk>ok n is n'; length stk = n\<rbrakk> \<Longrightarrow> length (exec is s stk) = n'"}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   870
and that instruction sequences generated by \<open>comp\<close>
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   871
cannot cause stack underflow: \ \<open>ok n (comp a) ?\<close> \ for
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68800
diff changeset
   872
some suitable value of \<open>?\<close>.
54186
ea92cce67f09 added exercise
nipkow
parents: 53015
diff changeset
   873
\end{exercise}
ea92cce67f09 added exercise
nipkow
parents: 53015
diff changeset
   874
\fi
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 63935
diff changeset
   875
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   876
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   877
end
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   878
(*>*)