src/Doc/ProgProve/Types_and_funs.thy
author nipkow
Tue, 04 Feb 2014 17:44:15 +0100
changeset 55318 908fd015cf2e
parent 55317 834a84553e02
child 55348 366718f2ff85
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     1
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     2
theory Types_and_funs
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     3
imports Main
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}
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
     8
\section{Type and Function Definitions}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     9
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    10
Type synonyms are abbreviations for existing types, for example
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    11
*}
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
type_synonym string = "char list"
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
text{*
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    16
Type synonyms are expanded after parsing and are not present in internal representation and output. They are mere conveniences for the reader.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    17
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    18
\subsection{Datatypes}
51465
c5c466706549 add label for referencing in semantics book
kleing
parents: 51393
diff changeset
    19
\label{sec:datatypes}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    20
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    21
The general form of a datatype definition looks like this:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    22
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    23
\begin{tabular}{@ {}rclcll}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
    24
\isacom{datatype} @{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    25
     & = & $C_1 \ @{text"\""}\tau_{1,1}@{text"\""} \dots @{text"\""}\tau_{1,n_1}@{text"\""}$ \\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    26
     & $|$ & \dots \\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    27
     & $|$ & $C_k \ @{text"\""}\tau_{k,1}@{text"\""} \dots @{text"\""}\tau_{k,n_k}@{text"\""}$
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    28
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    29
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    30
It introduces the constructors \
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
    31
$C_i :: \tau_{i,1}\Rightarrow \cdots \Rightarrow \tau_{i,n_i} \Rightarrow$~@{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"} \ and expresses that any value of this type is built from these constructors in a unique manner. Uniqueness is implied by the following
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    32
properties of the constructors:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    33
\begin{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    34
\item \emph{Distinctness:} $C_i\ \ldots \neq C_j\ \dots$ \quad if $i \neq j$
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    35
\item \emph{Injectivity:}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    36
\begin{tabular}[t]{l}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    37
 $(C_i \ x_1 \dots x_{n_i} = C_i \ y_1 \dots y_{n_i}) =$\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    38
 $(x_1 = y_1 \land \dots \land x_{n_i} = y_{n_i})$
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    39
\end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    40
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    41
The fact that any value of the datatype is built from the constructors implies
55318
nipkow
parents: 55317
diff changeset
    42
the \concept{structural induction} rule: to show
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
    43
$P~x$ for all $x$ of type @{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"},
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    44
one needs to show $P(C_i\ x_1 \dots x_{n_i})$ (for each $i$) assuming
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
    45
$P(x_j)$ for all $j$ where $\tau_{i,j} =$~@{text "('a\<^sub>1,\<dots>,'a\<^sub>n)t"}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    46
Distinctness and injectivity are applied automatically by @{text auto}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    47
and other proof methods. Induction must be applied explicitly.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    48
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    49
Datatype values can be taken apart with case-expressions, for example
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    50
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    51
\noquotes{@{term[source] "(case xs of [] \<Rightarrow> 0 | x # _ \<Rightarrow> Suc x)"}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    52
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    53
just like in functional programming languages. Case expressions
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    54
must be enclosed in parentheses.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    55
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    56
As an example, consider binary trees:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    57
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    58
47306
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47269
diff changeset
    59
datatype 'a tree = Tip | Node  "'a tree"  'a  "'a tree"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    60
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    61
text{* with a mirror function: *}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    62
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    63
fun mirror :: "'a tree \<Rightarrow> 'a tree" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    64
"mirror Tip = Tip" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    65
"mirror (Node l a r) = Node (mirror r) a (mirror l)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    66
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    67
text{* The following lemma illustrates induction: *}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    68
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    69
lemma "mirror(mirror t) = t"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    70
apply(induction t)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    71
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    72
txt{* yields
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    73
@{subgoals[display]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    74
The induction step contains two induction hypotheses, one for each subtree.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    75
An application of @{text auto} finishes the proof.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    76
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    77
A very simple but also very useful datatype is the predefined
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    78
@{datatype[display] option}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    79
Its sole purpose is to add a new element @{const None} to an existing
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    80
type @{typ 'a}. To make sure that @{const None} is distinct from all the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    81
elements of @{typ 'a}, you wrap them up in @{const Some} and call
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    82
the new type @{typ"'a option"}. A typical application is a lookup function
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    83
on a list of key-value pairs, often called an association list:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    84
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    85
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    86
apply auto
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    87
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    88
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    89
fun lookup :: "('a * 'b) list \<Rightarrow> 'a \<Rightarrow> 'b option" where
47306
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47269
diff changeset
    90
"lookup [] x = None" |
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47269
diff changeset
    91
"lookup ((a,b) # ps) x = (if a = x then Some b else lookup ps x)"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    92
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    93
text{*
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
    94
Note that @{text"\<tau>\<^sub>1 * \<tau>\<^sub>2"} is the type of pairs, also written @{text"\<tau>\<^sub>1 \<times> \<tau>\<^sub>2"}.
51393
df0f306f030f added pairs
nipkow
parents: 48985
diff changeset
    95
Pairs can be taken apart either by pattern matching (as above) or with the
df0f306f030f added pairs
nipkow
parents: 48985
diff changeset
    96
projection functions @{const fst} and @{const snd}: @{thm fst_conv} and @{thm snd_conv}. Tuples are simulated by pairs nested to the right: @{term"(a,b,c)"}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
    97
abbreviates @{text"(a, (b, c))"} and @{text "\<tau>\<^sub>1 \<times> \<tau>\<^sub>2 \<times> \<tau>\<^sub>3"} abbreviates
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
    98
@{text "\<tau>\<^sub>1 \<times> (\<tau>\<^sub>2 \<times> \<tau>\<^sub>3)"}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    99
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   100
\subsection{Definitions}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   101
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   102
Non recursive functions can be defined as in the following example:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   103
*}
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
definition sq :: "nat \<Rightarrow> nat" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   106
"sq n = n * n"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   107
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   108
text{* Such definitions do not allow pattern matching but only
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   109
@{text"f x\<^sub>1 \<dots> x\<^sub>n = t"}, where @{text f} does not occur in @{text t}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   110
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   111
\subsection{Abbreviations}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   112
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   113
Abbreviations are similar to definitions:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   114
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   115
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   116
abbreviation sq' :: "nat \<Rightarrow> nat" where
52045
nipkow
parents: 51465
diff changeset
   117
"sq' n \<equiv> n * n"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   118
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   119
text{* The key difference is that @{const sq'} is only syntactic sugar:
52045
nipkow
parents: 51465
diff changeset
   120
after parsing, @{term"sq' t"} is replaced by \mbox{@{term"t*t"}}, and
nipkow
parents: 51465
diff changeset
   121
before printing, every occurrence of @{term"u*u"} is replaced by
nipkow
parents: 51465
diff changeset
   122
\mbox{@{term"sq' u"}}.  Internally, @{const sq'} does not exist.
nipkow
parents: 51465
diff changeset
   123
This is the
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   124
advantage of abbreviations over definitions: definitions need to be expanded
52045
nipkow
parents: 51465
diff changeset
   125
explicitly (\autoref{sec:rewr-defs}) whereas abbreviations are already
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   126
expanded upon parsing. However, abbreviations should be introduced sparingly:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   127
if abused, they can lead to a confusing discrepancy between the internal and
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   128
external view of a term.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   129
52045
nipkow
parents: 51465
diff changeset
   130
The ASCII representation of @{text"\<equiv>"} is \texttt{==} or \xsymbol{equiv}.
nipkow
parents: 51465
diff changeset
   131
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
   132
\subsection{Recursive Functions}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   133
\label{sec:recursive-funs}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   134
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   135
Recursive functions are defined with \isacom{fun} by pattern matching
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   136
over datatype constructors. The order of equations matters. Just as in
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   137
functional programming languages. However, all HOL functions must be
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   138
total. This simplifies the logic---terms are always defined---but means
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   139
that recursive functions must terminate. Otherwise one could define a
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   140
function @{prop"f n = f n + (1::nat)"} and conclude \mbox{@{prop"(0::nat) = 1"}} by
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   141
subtracting @{term"f n"} on both sides.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   142
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   143
Isabelle's automatic termination checker requires that the arguments of
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   144
recursive calls on the right-hand side must be strictly smaller than the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   145
arguments on the left-hand side. In the simplest case, this means that one
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   146
fixed argument position decreases in size with each recursive call. The size
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   147
is measured as the number of constructors (excluding 0-ary ones, e.g.\ @{text
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   148
Nil}). Lexicographic combinations are also recognized. In more complicated
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   149
situations, the user may have to prove termination by hand. For details
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   150
see~\cite{Krauss}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   151
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   152
Functions defined with \isacom{fun} come with their own induction schema
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   153
that mirrors the recursion schema and is derived from the termination
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   154
order. For example,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   155
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   156
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   157
fun div2 :: "nat \<Rightarrow> nat" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   158
"div2 0 = 0" |
54136
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   159
"div2 (Suc 0) = 0" |
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   160
"div2 (Suc(Suc n)) = Suc(div2 n)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   161
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   162
text{* does not just define @{const div2} but also proves a
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   163
customized induction rule:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   164
\[
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   165
\inferrule{
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   166
\mbox{@{thm (prem 1) div2.induct}}\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   167
\mbox{@{thm (prem 2) div2.induct}}\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   168
\mbox{@{thm (prem 3) div2.induct}}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   169
{\mbox{@{thm (concl) div2.induct[of _ "m"]}}}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   170
\]
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   171
This customized induction rule can simplify inductive proofs. For example,
47269
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
54571
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   174
lemma "div2(n) = n div 2"
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   175
apply(induction n rule: div2.induct)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   176
54571
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   177
txt{* (where the infix @{text div} is the predefined division operation)
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   178
yields the 3 subgoals
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   179
@{subgoals[display,margin=65]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   180
An application of @{text auto} finishes the proof.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   181
Had we used ordinary structural induction on @{text n},
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   182
the proof would have needed an additional
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   183
case analysis in the induction step.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   184
54571
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   185
This example leads to the following induction heuristic:
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   186
\begin{quote}
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   187
\emph{Let @{text f} be a recursive function.
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   188
If the definition of @{text f} is more complicated
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   189
than having one equation for each constructor of some datatype,
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   190
then properties of @{text f} are best proved via @{text "f.induct"}.}
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   191
\end{quote}
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   192
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   193
The general case is often called \concept{computation induction},
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   194
because the induction follows the (terminating!) computation.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   195
For every defining equation
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   196
\begin{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   197
@{text "f(e) = \<dots> f(r\<^sub>1) \<dots> f(r\<^sub>k) \<dots>"}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   198
\end{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   199
where @{text"f(r\<^sub>i)"}, @{text"i=1\<dots>k"}, are all the recursive calls,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   200
the induction rule @{text"f.induct"} contains one premise of the form
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   201
\begin{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   202
@{text"P(r\<^sub>1) \<Longrightarrow> \<dots> \<Longrightarrow> P(r\<^sub>k) \<Longrightarrow> P(e)"}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   203
\end{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   204
If @{text "f :: \<tau>\<^sub>1 \<Rightarrow> \<dots> \<Rightarrow> \<tau>\<^sub>n \<Rightarrow> \<tau>"} then @{text"f.induct"} is applied like this:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   205
\begin{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   206
\isacom{apply}@{text"(induction x\<^sub>1 \<dots> x\<^sub>n rule: f.induct)"}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   207
\end{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   208
where typically there is a call @{text"f x\<^sub>1 \<dots> x\<^sub>n"} in the goal.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   209
But note that the induction rule does not mention @{text f} at all,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   210
except in its name, and is applicable independently of @{text f}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   211
54136
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   212
54436
nipkow
parents: 54195
diff changeset
   213
\subsection*{Exercises}
54136
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   214
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   215
\begin{exercise}
54194
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   216
Starting from the type @{text "'a tree"} defined in the text, define
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   217
a function @{text "contents ::"} @{typ "'a tree \<Rightarrow> 'a list"}
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   218
that collects all values in a tree in a list, in any order,
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   219
without removing duplicates.
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   220
Then define a function @{text "treesum ::"} @{typ "nat tree \<Rightarrow> nat"}
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   221
that sums up all values in a tree of natural numbers
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   222
and prove @{prop "treesum t = listsum(contents t)"}.
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   223
\end{exercise}
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   224
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   225
\begin{exercise}
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   226
Define a new type @{text "'a tree2"} of binary trees where values are also
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   227
stored in the leaves of the tree.  Also reformulate the
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   228
@{const mirror} function accordingly. Define two functions
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   229
@{text "pre_order"} and @{text "post_order"} of type @{text "'a tree2 \<Rightarrow> 'a list"}
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   230
that traverse a tree and collect all stored values in the respective order in
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   231
a list. Prove @{prop "pre_order (mirror t) = rev (post_order t)"}.
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   232
\end{exercise}
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   233
12c9ad3142be added exercises
nipkow
parents: 54136
diff changeset
   234
\begin{exercise}
54571
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   235
Define a function @{text "intersperse ::"} @{typ "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"}
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   236
such that @{text "intersperse a [x\<^sub>1, ..., x\<^sub>n] = [x\<^sub>1, a, x\<^sub>2, a, ..., a, x\<^sub>n]"}.
be1186cb03ce replaced bad example
nipkow
parents: 54467
diff changeset
   237
Now prove that @{prop "map f (intersperse a xs) = intersperse (f a) (map f xs)"}.
54136
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   238
\end{exercise}
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   239
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   240
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
   241
\section{Induction Heuristics}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   242
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   243
We have already noted that theorems about recursive functions are proved by
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   244
induction. In case the function has more than one argument, we have followed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   245
the following heuristic in the proofs about the append function:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   246
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   247
\emph{Perform induction on argument number $i$\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   248
 if the function is defined by recursion on argument number $i$.}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   249
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   250
The key heuristic, and the main point of this section, is to
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   251
\emph{generalize the goal before induction}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   252
The reason is simple: if the goal is
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   253
too specific, the induction hypothesis is too weak to allow the induction
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   254
step to go through. Let us illustrate the idea with an example.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   255
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   256
Function @{const rev} has quadratic worst-case running time
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   257
because it calls append for each element of the list and
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   258
append is linear in its first argument.  A linear time version of
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   259
@{const rev} requires an extra argument where the result is accumulated
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   260
gradually, using only~@{text"#"}:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   261
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   262
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   263
apply auto
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   264
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   265
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   266
fun itrev :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   267
"itrev []        ys = ys" |
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   268
"itrev (x#xs) ys = itrev xs (x#ys)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   270
text{* The behaviour of @{const itrev} is simple: it reverses
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   271
its first argument by stacking its elements onto the second argument,
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   272
and it returns that second argument when the first one becomes
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   273
empty. Note that @{const itrev} is tail-recursive: it can be
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   274
compiled into a loop, no stack is necessary for executing it.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   275
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   276
Naturally, we would like to show that @{const itrev} does indeed reverse
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   277
its first argument provided the second one is empty:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   278
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   279
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   280
lemma "itrev xs [] = rev xs"
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
txt{* There is no choice as to the induction variable:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   283
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   284
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   285
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   286
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   287
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   288
txt{*
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   289
Unfortunately, this attempt does not prove
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   290
the induction step:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   291
@{subgoals[display,margin=70]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   292
The induction hypothesis is too weak.  The fixed
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   293
argument,~@{term"[]"}, prevents it from rewriting the conclusion.  
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   294
This example suggests a heuristic:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   295
\begin{quote}
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   296
\emph{Generalize goals for induction by replacing constants by variables.}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   297
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   298
Of course one cannot do this na\"{\i}vely: @{prop"itrev xs ys = rev xs"} is
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   299
just not true.  The correct generalization is
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   300
*};
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   301
(*<*)oops;(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   302
lemma "itrev xs ys = rev xs @ ys"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   303
(*<*)apply(induction xs, auto)(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   304
txt{*
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   305
If @{text ys} is replaced by @{term"[]"}, the right-hand side simplifies to
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   306
@{term"rev xs"}, as required.
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   307
In this instance it was easy to guess the right generalization.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   308
Other situations can require a good deal of creativity.  
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   309
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   310
Although we now have two variables, only @{text xs} is suitable for
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   311
induction, and we repeat our proof attempt. Unfortunately, we are still
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   312
not there:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   313
@{subgoals[display,margin=65]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   314
The induction hypothesis is still too weak, but this time it takes no
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   315
intuition to generalize: the problem is that the @{text ys}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   316
in the induction hypothesis is fixed,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   317
but the induction hypothesis needs to be applied with
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   318
@{term"a # ys"} instead of @{text ys}. Hence we prove the theorem
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   319
for all @{text ys} instead of a fixed one. We can instruct induction
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   320
to perform this generalization for us by adding @{text "arbitrary: ys"}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   321
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   322
(*<*)oops
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   323
lemma "itrev xs ys = rev xs @ ys"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   324
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   325
apply(induction xs arbitrary: ys)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   326
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   327
txt{* The induction hypothesis in the induction step is now universally quantified over @{text ys}:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   328
@{subgoals[display,margin=65]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   329
Thus the proof succeeds:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   330
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   331
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   332
apply auto
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   333
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   334
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   335
text{*
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   336
This leads to another heuristic for generalization:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   337
\begin{quote}
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   338
\emph{Generalize induction by generalizing all free
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   339
variables\\ {\em(except the induction variable itself)}.}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   340
\end{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   341
Generalization is best performed with @{text"arbitrary: y\<^sub>1 \<dots> y\<^sub>k"}. 
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   342
This heuristic prevents trivial failures like the one above.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   343
However, it should not be applied blindly.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   344
It is not always required, and the additional quantifiers can complicate
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   345
matters in some cases. The variables that need to be quantified are typically
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   346
those that change in recursive calls.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   347
54136
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   348
54436
nipkow
parents: 54195
diff changeset
   349
\subsection*{Exercises}
54136
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   350
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   351
\begin{exercise}
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   352
Write a tail-recursive variant of the @{text add} function on @{typ nat}:
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   353
@{term "itadd :: nat \<Rightarrow> nat \<Rightarrow> nat"}.
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   354
Tail-recursive means that in the recursive case, @{text itadd} needs to call
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   355
itself directly: \mbox{@{term"itadd (Suc m) n"}} @{text"= itadd \<dots>"}.
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   356
Prove @{prop "itadd m n = add m n"}.
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   357
\end{exercise}
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   358
c206cb2a3afe more exercises
nipkow
parents: 53015
diff changeset
   359
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   360
\section{Simplification}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   361
55317
834a84553e02 started index
nipkow
parents: 54605
diff changeset
   362
So far we have talked a lot about simplifying terms without explaining the concept.
834a84553e02 started index
nipkow
parents: 54605
diff changeset
   363
\conceptidx{Simplification}{simplification} means
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   364
\begin{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   365
\item using equations $l = r$ from left to right (only),
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   366
\item as long as possible.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   367
\end{itemize}
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   368
To emphasize the directionality, equations that have been given the
55317
834a84553e02 started index
nipkow
parents: 54605
diff changeset
   369
@{text"simp"} attribute are called \conceptidx{simplification rules}{simplification rule}.
834a84553e02 started index
nipkow
parents: 54605
diff changeset
   370
Logically, they are still symmetric, but proofs by
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   371
simplification use them only in the left-to-right direction.  The proof tool
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   372
that performs simplifications is called the \concept{simplifier}. It is the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   373
basis of @{text auto} and other related proof methods.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   374
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   375
The idea of simplification is best explained by an example. Given the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   376
simplification rules
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   377
\[
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   378
\begin{array}{rcl@ {\quad}l}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   379
@{term"0 + n::nat"} &@{text"="}& @{text n} & (1) \\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   380
@{term"(Suc m) + n"} &@{text"="}& @{term"Suc (m + n)"} & (2) \\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   381
@{text"(Suc m \<le> Suc n)"} &@{text"="}& @{text"(m \<le> n)"} & (3)\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   382
@{text"(0 \<le> m)"} &@{text"="}& @{const True} & (4)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   383
\end{array}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   384
\]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   385
the formula @{prop"0 + Suc 0 \<le> Suc 0 + x"} is simplified to @{const True}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   386
as follows:
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
\begin{array}{r@ {}c@ {}l@ {\quad}l}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   389
@{text"(0 + Suc 0"} & \leq & @{text"Suc 0 + x)"}  & \stackrel{(1)}{=} \\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   390
@{text"(Suc 0"}     & \leq & @{text"Suc 0 + x)"}  & \stackrel{(2)}{=} \\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   391
@{text"(Suc 0"}     & \leq & @{text"Suc (0 + x)"} & \stackrel{(3)}{=} \\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   392
@{text"(0"}         & \leq & @{text"0 + x)"}      & \stackrel{(4)}{=} \\[1ex]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   393
 & @{const True}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   394
\end{array}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   395
\]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   396
Simplification is often also called \concept{rewriting}
55317
834a84553e02 started index
nipkow
parents: 54605
diff changeset
   397
and simplification rules \conceptidx{rewrite rules}{rewrite rule}.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   398
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
   399
\subsection{Simplification Rules}
47269
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
The attribute @{text"simp"} declares theorems to be simplification rules,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   402
which the simplifier will use automatically. In addition,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   403
\isacom{datatype} and \isacom{fun} commands implicitly declare some
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   404
simplification rules: \isacom{datatype} the distinctness and injectivity
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   405
rules, \isacom{fun} the defining equations. Definitions are not declared
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   406
as simplification rules automatically! Nearly any theorem can become a
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   407
simplification rule. The simplifier will try to transform it into an
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   408
equation. For example, the theorem @{prop"\<not> P"} is turned into @{prop"P = False"}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   409
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   410
Only equations that really simplify, like @{prop"rev (rev xs) = xs"} and
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   411
@{prop"xs @ [] = xs"}, should be declared as simplification
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   412
rules. Equations that may be counterproductive as simplification rules
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   413
should only be used in specific proof steps (see \S\ref{sec:simp} below).
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   414
Distributivity laws, for example, alter the structure of terms
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   415
and can produce an exponential blow-up.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   416
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
   417
\subsection{Conditional Simplification Rules}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   418
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   419
Simplification rules can be conditional.  Before applying such a rule, the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   420
simplifier will first try to prove the preconditions, again by
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   421
simplification. For example, given the simplification rules
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   422
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   423
@{prop"p(0::nat) = True"}\\
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   424
@{prop"p(x) \<Longrightarrow> f(x) = g(x)"},
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   425
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   426
the term @{term"f(0::nat)"} simplifies to @{term"g(0::nat)"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   427
but @{term"f(1::nat)"} does not simplify because @{prop"p(1::nat)"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   428
is not provable.
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
\subsection{Termination}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   431
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   432
Simplification can run forever, for example if both @{prop"f x = g x"} and
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   433
@{prop"g(x) = f(x)"} are simplification rules. It is the user's
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   434
responsibility not to include simplification rules that can lead to
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   435
nontermination, either on their own or in combination with other
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   436
simplification rules. The right-hand side of a simplification rule should
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   437
always be ``simpler'' than the left-hand side---in some sense. But since
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   438
termination is undecidable, such a check cannot be automated completely
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   439
and Isabelle makes little attempt to detect nontermination.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   440
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   441
When conditional simplification rules are applied, their preconditions are
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   442
proved first. Hence all preconditions need to be
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   443
simpler than the left-hand side of the conclusion. For example
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   444
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   445
@{prop "n < m \<Longrightarrow> (n < Suc m) = True"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   446
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   447
is suitable as a simplification rule: both @{prop"n<m"} and @{const True}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   448
are simpler than \mbox{@{prop"n < Suc m"}}. But
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   449
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   450
@{prop "Suc n < m \<Longrightarrow> (n < m) = True"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   451
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   452
leads to nontermination: when trying to rewrite @{prop"n<m"} to @{const True}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   453
one first has to prove \mbox{@{prop"Suc n < m"}}, which can be rewritten to @{const True} provided @{prop"Suc(Suc n) < m"}, \emph{ad infinitum}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   454
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
   455
\subsection{The @{text simp} Proof Method}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   456
\label{sec:simp}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   457
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   458
So far we have only used the proof method @{text auto}.  Method @{text simp}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   459
is the key component of @{text auto}, but @{text auto} can do much more. In
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   460
some cases, @{text auto} is overeager and modifies the proof state too much.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   461
In such cases the more predictable @{text simp} method should be used.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   462
Given a goal
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   463
\begin{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   464
@{text"1. \<lbrakk> P\<^sub>1; \<dots>; P\<^sub>m \<rbrakk> \<Longrightarrow> C"}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   465
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   466
the command
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   467
\begin{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   468
\isacom{apply}@{text"(simp add: th\<^sub>1 \<dots> th\<^sub>n)"}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   469
\end{quote}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   470
simplifies the assumptions @{text "P\<^sub>i"} and the conclusion @{text C} using
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   471
\begin{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   472
\item all simplification rules, including the ones coming from \isacom{datatype} and \isacom{fun},
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52718
diff changeset
   473
\item the additional lemmas @{text"th\<^sub>1 \<dots> th\<^sub>n"}, and
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   474
\item the assumptions.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   475
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   476
In addition to or instead of @{text add} there is also @{text del} for removing
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   477
simplification rules temporarily. Both are optional. Method @{text auto}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   478
can be modified similarly:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   479
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   480
\isacom{apply}@{text"(auto simp add: \<dots> simp del: \<dots>)"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   481
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   482
Here the modifiers are @{text"simp add"} and @{text"simp del"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   483
instead of just @{text add} and @{text del} because @{text auto}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   484
does not just perform simplification.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   485
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   486
Note that @{text simp} acts only on subgoal 1, @{text auto} acts on all
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   487
subgoals. There is also @{text simp_all}, which applies @{text simp} to
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   488
all subgoals.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   489
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
   490
\subsection{Rewriting With Definitions}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   491
\label{sec:rewr-defs}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   492
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   493
Definitions introduced by the command \isacom{definition}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   494
can also be used as simplification rules,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   495
but by default they are not: the simplifier does not expand them
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   496
automatically. Definitions are intended for introducing abstract concepts and
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   497
not merely as abbreviations. Of course, we need to expand the definition
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   498
initially, but once we have proved enough abstract properties of the new
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   499
constant, we can forget its original definition. This style makes proofs more
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   500
robust: if the definition has to be changed, only the proofs of the abstract
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   501
properties will be affected.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   502
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   503
The definition of a function @{text f} is a theorem named @{text f_def}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   504
and can be added to a call of @{text simp} just like any other theorem:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   505
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   506
\isacom{apply}@{text"(simp add: f_def)"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   507
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   508
In particular, let-expressions can be unfolded by
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   509
making @{thm[source] Let_def} a simplification rule.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   510
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 52045
diff changeset
   511
\subsection{Case Splitting With @{text simp}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   512
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   513
Goals containing if-expressions are automatically split into two cases by
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   514
@{text simp} using the rule
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   515
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   516
@{prop"P(if A then s else t) = ((A \<longrightarrow> P(s)) \<and> (~A \<longrightarrow> P(t)))"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   517
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   518
For example, @{text simp} can prove
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   519
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   520
@{prop"(A \<and> B) = (if A then B else False)"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   521
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   522
because both @{prop"A \<longrightarrow> (A & B) = B"} and @{prop"~A \<longrightarrow> (A & B) = False"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   523
simplify to @{const True}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   524
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   525
We can split case-expressions similarly. For @{text nat} the rule looks like this:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   526
@{prop[display,margin=65,indent=4]"P(case e of 0 \<Rightarrow> a | Suc n \<Rightarrow> b n) = ((e = 0 \<longrightarrow> P a) & (ALL n. e = Suc n \<longrightarrow> P(b n)))"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   527
Case expressions are not split automatically by @{text simp}, but @{text simp}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   528
can be instructed to do so:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   529
\begin{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   530
\isacom{apply}@{text"(simp split: nat.split)"}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   531
\end{quote}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   532
splits all case-expressions over natural numbers. For an arbitrary
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   533
datatype @{text t} it is @{text "t.split"} instead of @{thm[source] nat.split}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   534
Method @{text auto} can be modified in exactly the same way.
54605
nipkow
parents: 54571
diff changeset
   535
The modifier @{text "split:"} can be followed by multiple names.
nipkow
parents: 54571
diff changeset
   536
Splitting if or case-expressions in the assumptions requires 
nipkow
parents: 54571
diff changeset
   537
@{text "split: if_splits"} or @{text "split: t.splits"}.
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   538
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   539
54436
nipkow
parents: 54195
diff changeset
   540
\subsection*{Exercises}
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   541
54195
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   542
\exercise\label{exe:tree0}
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   543
Define a datatype @{text tree0} of binary tree skeletons which do not store
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   544
any information, neither in the inner nodes nor in the leaves.
54467
663a927fdc88 comments by Sean Seefried
nipkow
parents: 54436
diff changeset
   545
Define a function @{text "nodes :: tree0 \<Rightarrow> nat"} that counts the number of
54195
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   546
all nodes (inner nodes and leaves) in such a tree.
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   547
Consider the following recursive function:
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   548
*}
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   549
(*<*)
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   550
datatype tree0 = Tip | Node tree0 tree0
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   551
(*>*)
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   552
fun explode :: "nat \<Rightarrow> tree0 \<Rightarrow> tree0" where
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   553
"explode 0 t = t" |
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   554
"explode (Suc n) t = explode n (Node t t)"
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   555
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   556
text {*
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   557
Find an equation expressing the size of a tree after exploding it
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   558
(\noquotes{@{term [source] "nodes (explode n t)"}}) as a function
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   559
of @{term "nodes t"} and @{text n}. Prove your equation.
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   560
You may use the usual arithmetic operators including the exponentiation
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   561
operator ``@{text"^"}''. For example, \noquotes{@{prop [source] "2 ^ 2 = 4"}}.
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   562
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   563
Hint: simplifying with the list of theorems @{thm[source] algebra_simps}
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   564
takes care of common algebraic properties of the arithmetic operators.
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   565
\endexercise
1e685123926d more exercises
nipkow
parents: 54194
diff changeset
   566
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   567
\exercise
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   568
Define arithmetic expressions in one variable over integers (type @{typ int})
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   569
as a data type:
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   570
*}
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   571
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   572
datatype exp = Var | Const int | Add exp exp | Mult exp exp
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   573
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   574
text{*
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   575
Define a function \noquotes{@{term [source]"eval :: exp \<Rightarrow> int \<Rightarrow> int"}}
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   576
such that @{term"eval e x"} evaluates @{text e} at the value
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   577
@{text x}.
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   578
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   579
A polynomial can be represented as a list of coefficients, starting with
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   580
the constant. For example, @{term "[4, 2, -1, 3::int]"} represents the
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   581
polynomial $4 + 2x - x^2 + 3x^3$.
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   582
Define a function \noquotes{@{term [source] "evalp :: int list \<Rightarrow> int \<Rightarrow> int"}}
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   583
that evaluates a polynomial at the given value.
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   584
Define a function \noquotes{@{term[source] "coeffs :: exp \<Rightarrow> int list"}}
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   585
that transforms an expression into a polynomial. This may require auxiliary
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   586
functions. Prove that @{text coeffs} preserves the value of the expression:
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   587
\mbox{@{prop"evalp (coeffs e) x = eval e x"}.}
54467
663a927fdc88 comments by Sean Seefried
nipkow
parents: 54436
diff changeset
   588
Hint: consider the hint in Exercise~\ref{exe:tree0}.
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   589
\endexercise
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   590
*}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   591
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   592
end
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   593
(*>*)