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