src/Doc/Prog_Prove/Bool_nat_list.thy
author wenzelm
Thu, 04 Aug 2022 22:15:50 +0200
changeset 75759 0cdccd0d1699
parent 73511 2cdbb6a2f2a7
permissions -rw-r--r--
clarified context for retrieval: more explicit types, with optional close() operation;
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 Bool_nat_list
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
     3
imports Complex_Main
47269
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
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
     7
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
     8
\vspace{-4ex}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
     9
\section{\texorpdfstring{Types \<^typ>\<open>bool\<close>, \<^typ>\<open>nat\<close> and \<open>list\<close>}{Types bool, nat and list}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    10
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    11
These are the most important predefined types. We go through them one by one.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    12
Based on examples we learn how to define (possibly recursive) functions and
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    13
prove theorems about them by induction and simplification.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    14
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    15
\subsection{Type \indexed{\<^typ>\<open>bool\<close>}{bool}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    16
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    17
The type of boolean values is a predefined datatype
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    18
@{datatype[display] bool}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    19
with the two values \indexed{\<^const>\<open>True\<close>}{True} and \indexed{\<^const>\<open>False\<close>}{False} and
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    20
with many predefined functions:  \<open>\<not>\<close>, \<open>\<and>\<close>, \<open>\<or>\<close>, \<open>\<longrightarrow>\<close>, etc. Here is how conjunction could be defined by pattern matching:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
    21
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    22
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    23
fun conj :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    24
"conj True True = True" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    25
"conj _ _ = False"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    26
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
    27
text\<open>Both the datatype and function definitions roughly follow the syntax
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    28
of functional programming languages.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    29
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    30
\subsection{Type \indexed{\<^typ>\<open>nat\<close>}{nat}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    31
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    32
Natural numbers are another predefined datatype:
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    33
@{datatype[display] nat}\index{Suc@\<^const>\<open>Suc\<close>}
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    34
All values of type \<^typ>\<open>nat\<close> are generated by the constructors
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    35
\<open>0\<close> and \<^const>\<open>Suc\<close>. Thus the values of type \<^typ>\<open>nat\<close> are
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    36
\<open>0\<close>, \<^term>\<open>Suc 0\<close>, \<^term>\<open>Suc(Suc 0)\<close>, etc.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    37
There are many predefined functions: \<open>+\<close>, \<open>*\<close>, \<open>\<le>\<close>, etc. Here is how you could define your own addition:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
    38
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    39
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    40
fun add :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    41
"add 0 n = n" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    42
"add (Suc m) n = Suc(add m n)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    43
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    44
text\<open>And here is a proof of the fact that \<^prop>\<open>add m 0 = m\<close>:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    45
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    46
lemma add_02: "add m 0 = m"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    47
apply(induction m)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    48
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    49
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    50
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    51
lemma "add m 0 = m"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    52
apply(induction m)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    53
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
    54
txt\<open>The \isacom{lemma} command starts the proof and gives the lemma
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    55
a name, \<open>add_02\<close>. Properties of recursively defined functions
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    56
need to be established by induction in most cases.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    57
Command \isacom{apply}\<open>(induction m)\<close> instructs Isabelle to
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    58
start a proof by induction on \<open>m\<close>. In response, it will show the
62222
54a7b9422d3e synchronized with book
nipkow
parents: 59360
diff changeset
    59
following proof state\ifsem\footnote{See page \pageref{proof-state} for how to
62223
nipkow
parents: 62222
diff changeset
    60
display the proof state.}\fi:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    61
@{subgoals[display,indent=0]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    62
The numbered lines are known as \emph{subgoals}.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    63
The first subgoal is the base case, the second one the induction step.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    64
The prefix \<open>\<And>m.\<close> is Isabelle's way of saying ``for an arbitrary but fixed \<open>m\<close>''. The \<open>\<Longrightarrow>\<close> separates assumptions from the conclusion.
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    65
The command \isacom{apply}\<open>(auto)\<close> instructs Isabelle to try
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    66
and prove all subgoals automatically, essentially by simplifying them.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    67
Because both subgoals are easy, Isabelle can do it.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    68
The base case \<^prop>\<open>add 0 0 = 0\<close> holds by definition of \<^const>\<open>add\<close>,
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    69
and the induction step is almost as simple:
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    70
\<open>add\<^latex>\<open>~\<close>(Suc m) 0 = Suc(add m 0) = Suc m\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    71
using first the definition of \<^const>\<open>add\<close> and then the induction hypothesis.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    72
In summary, both subproofs rely on simplification with function definitions and
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    73
the induction hypothesis.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    74
As a result of that final \isacom{done}, Isabelle associates the lemma
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    75
just proved with its name. You can now inspect the lemma with the command
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
    76
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    77
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    78
thm add_02
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    79
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
    80
txt\<open>which displays @{thm[show_question_marks,display] add_02} The free
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    81
variable \<open>m\<close> has been replaced by the \concept{unknown}
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    82
\<open>?m\<close>. There is no logical difference between the two but there is an
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    83
operational one: unknowns can be instantiated, which is what you want after
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    84
some lemma has been proved.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    85
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    86
Note that there is also a proof method \<open>induct\<close>, which behaves almost
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    87
like \<open>induction\<close>; the difference is explained in \autoref{ch:Isar}.
47269
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
\begin{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    90
Terminology: We use \concept{lemma}, \concept{theorem} and \concept{rule}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    91
interchangeably for propositions that have been proved.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    92
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    93
\begin{warn}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    94
  Numerals (\<open>0\<close>, \<open>1\<close>, \<open>2\<close>, \dots) and most of the standard
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    95
  arithmetic operations (\<open>+\<close>, \<open>-\<close>, \<open>*\<close>, \<open>\<le>\<close>,
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    96
  \<open><\<close>, etc.) are overloaded: they are available
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    97
  not just for natural numbers but for other types as well.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
    98
  For example, given the goal \<open>x + 0 = x\<close>, there is nothing to indicate
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
    99
  that you are talking about natural numbers. Hence Isabelle can only infer
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   100
  that \<^term>\<open>x\<close> is of some arbitrary type where \<open>0\<close> and \<open>+\<close>
54508
nipkow
parents: 54467
diff changeset
   101
  exist. As a consequence, you will be unable to prove the goal.
nipkow
parents: 54467
diff changeset
   102
%  To alert you to such pitfalls, Isabelle flags numerals without a
nipkow
parents: 54467
diff changeset
   103
%  fixed type in its output: @ {prop"x+0 = x"}.
nipkow
parents: 54467
diff changeset
   104
  In this particular example, you need to include
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   105
  an explicit type constraint, for example \<open>x+0 = (x::nat)\<close>. If there
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   106
  is enough contextual information this may not be necessary: \<^prop>\<open>Suc x =
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   107
  x\<close> automatically implies \<open>x::nat\<close> because \<^term>\<open>Suc\<close> is not
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   108
  overloaded.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   109
\end{warn}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   110
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 48985
diff changeset
   111
\subsubsection{An Informal Proof}
47269
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
Above we gave some terse informal explanation of the proof of
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   114
\<^prop>\<open>add m 0 = m\<close>. A more detailed informal exposition of the lemma
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   115
might look like this:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   116
\bigskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   117
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   118
\noindent
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   119
\textbf{Lemma} \<^prop>\<open>add m 0 = m\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   120
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   121
\noindent
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   122
\textbf{Proof} by induction on \<open>m\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   123
\begin{itemize}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   124
\item Case \<open>0\<close> (the base case): \<^prop>\<open>add 0 0 = 0\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   125
  holds by definition of \<^const>\<open>add\<close>.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   126
\item Case \<^term>\<open>Suc m\<close> (the induction step):
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   127
  We assume \<^prop>\<open>add m 0 = m\<close>, the induction hypothesis (IH),
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   128
  and we need to show \<open>add (Suc m) 0 = Suc m\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   129
  The proof is as follows:\smallskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   130
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   131
  \begin{tabular}{@ {}rcl@ {\quad}l@ {}}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   132
  \<^term>\<open>add (Suc m) 0\<close> &\<open>=\<close>& \<^term>\<open>Suc(add m 0)\<close>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   133
  & by definition of \<open>add\<close>\\
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   134
              &\<open>=\<close>& \<^term>\<open>Suc m\<close> & by IH
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   135
  \end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   136
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   137
Throughout this book, \concept{IH} will stand for ``induction hypothesis''.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   138
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   139
We have now seen three proofs of \<^prop>\<open>add m 0 = 0\<close>: the Isabelle one, the
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   140
terse four lines explaining the base case and the induction step, and just now a
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   141
model of a traditional inductive proof. The three proofs differ in the level
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   142
of detail given and the intended reader: the Isabelle proof is for the
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   143
machine, the informal proofs are for humans. Although this book concentrates
47711
c1cca2a052e4 doc update
nipkow
parents: 47306
diff changeset
   144
on Isabelle proofs, it is important to be able to rephrase those proofs
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   145
as informal text comprehensible to a reader familiar with traditional
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   146
mathematical proofs. Later on we will introduce an Isabelle proof language
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   147
that is closer to traditional informal mathematical language and is often
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   148
directly readable.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   149
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   150
\subsection{Type \indexed{\<open>list\<close>}{list}}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   151
58521
nipkow
parents: 58504
diff changeset
   152
Although lists are already predefined, we define our own copy for
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   153
demonstration purposes:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   154
\<close>
47269
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
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   157
done 
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   158
declare [[names_short]]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   159
(*>*)
47302
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   160
datatype 'a list = Nil | Cons 'a "'a list"
59204
0cbe0a56d3fa update map definition in Prog_Prove for new datatype package
kleing
parents: 58521
diff changeset
   161
(*<*)
0cbe0a56d3fa update map definition in Prog_Prove for new datatype package
kleing
parents: 58521
diff changeset
   162
for map: map
0cbe0a56d3fa update map definition in Prog_Prove for new datatype package
kleing
parents: 58521
diff changeset
   163
(*>*)
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   164
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   165
text\<open>
47302
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   166
\begin{itemize}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   167
\item Type \<^typ>\<open>'a list\<close> is the type of lists over elements of type \<^typ>\<open>'a\<close>. Because \<^typ>\<open>'a\<close> is a type variable, lists are in fact \concept{polymorphic}: the elements of a list can be of arbitrary type (but must all be of the same type).
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   168
\item Lists have two constructors: \<^const>\<open>Nil\<close>, the empty list, and \<^const>\<open>Cons\<close>, which puts an element (of type \<^typ>\<open>'a\<close>) in front of a list (of type \<^typ>\<open>'a list\<close>).
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   169
Hence all lists are of the form \<^const>\<open>Nil\<close>, or \<^term>\<open>Cons x Nil\<close>,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   170
or \<^term>\<open>Cons x (Cons y Nil)\<close>, etc.
47302
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   171
\item \isacom{datatype} requires no quotation marks on the
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   172
left-hand side, but on the right-hand side each of the argument
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   173
types of a constructor needs to be enclosed in quotation marks, unless
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   174
it is just an identifier (e.g., \<^typ>\<open>nat\<close> or \<^typ>\<open>'a\<close>).
47302
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   175
\end{itemize}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   176
We also define two standard functions, append and reverse:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   177
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   178
fun app :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   179
"app Nil ys = ys" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   180
"app (Cons x xs) ys = Cons x (app xs ys)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   181
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   182
fun rev :: "'a list \<Rightarrow> 'a list" where
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   183
"rev Nil = Nil" |
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   184
"rev (Cons x xs) = app (rev xs) (Cons x Nil)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   185
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   186
text\<open>By default, variables \<open>xs\<close>, \<open>ys\<close> and \<open>zs\<close> are of
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   187
\<open>list\<close> type.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   188
73511
2cdbb6a2f2a7 updated to latest latex due to new mechanism for dealing with bold ccfonts
nipkow
parents: 72656
diff changeset
   189
Command \indexed{\isacom{value}}{value} evaluates a term. For example,\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   190
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   191
value "rev(Cons True (Cons False Nil))"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   192
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   193
text\<open>yields the result \<^value>\<open>rev(Cons True (Cons False Nil))\<close>. This works symbolically, too:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   194
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   195
value "rev(Cons a (Cons b Nil))"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   196
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   197
text\<open>yields \<^value>\<open>rev(Cons a (Cons b Nil))\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   198
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   199
47302
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   200
Figure~\ref{fig:MyList} shows the theory created so far.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   201
Because \<open>list\<close>, \<^const>\<open>Nil\<close>, \<^const>\<open>Cons\<close>, etc.\ are already predefined,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   202
 Isabelle prints qualified (long) names when executing this theory, for example, \<open>MyList.Nil\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   203
 instead of \<^const>\<open>Nil\<close>.
47719
8aac84627b84 the perennial doc problem of how to define lists a second time
nipkow
parents: 47711
diff changeset
   204
 To suppress the qualified names you can insert the command
8aac84627b84 the perennial doc problem of how to define lists a second time
nipkow
parents: 47711
diff changeset
   205
 \texttt{declare [[names\_short]]}.
56989
nipkow
parents: 56451
diff changeset
   206
 This is not recommended in general but is convenient for this unusual example.
47302
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   207
% Notice where the
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   208
%quotations marks are needed that we mostly sweep under the carpet.  In
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   209
%particular, notice that \isacom{datatype} requires no quotation marks on the
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   210
%left-hand side, but that on the right-hand side each of the argument
70239da25ef6 towards showing " in the tutorial
nipkow
parents: 47269
diff changeset
   211
%types of a constructor needs to be enclosed in quotation marks.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   212
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   213
\begin{figure}[htbp]
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   214
\begin{alltt}
48947
7eee8b2d2099 more standard document preparation within session context;
wenzelm
parents: 47719
diff changeset
   215
\input{MyList.thy}\end{alltt}
58521
nipkow
parents: 58504
diff changeset
   216
\caption{A theory of lists}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   217
\label{fig:MyList}
57805
nipkow
parents: 56989
diff changeset
   218
\index{comment}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   219
\end{figure}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   220
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   221
\subsubsection{Structural Induction for Lists}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   222
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   223
Just as for natural numbers, there is a proof principle of induction for
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   224
lists. Induction over a list is essentially induction over the length of
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   225
the list, although the length remains implicit. To prove that some property
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   226
\<open>P\<close> holds for all lists \<open>xs\<close>, i.e., \mbox{\<^prop>\<open>P(xs)\<close>},
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   227
you need to prove
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   228
\begin{enumerate}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   229
\item the base case \<^prop>\<open>P(Nil)\<close> and
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   230
\item the inductive case \<^prop>\<open>P(Cons x xs)\<close> under the assumption \<^prop>\<open>P(xs)\<close>, for some arbitrary but fixed \<open>x\<close> and \<open>xs\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   231
\end{enumerate}
55318
nipkow
parents: 55317
diff changeset
   232
This is often called \concept{structural induction} for lists.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   233
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   234
\subsection{The Proof Process}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   235
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   236
We will now demonstrate the typical proof process, which involves
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   237
the formulation and proof of auxiliary lemmas.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   238
Our goal is to show that reversing a list twice produces the original
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   239
list.\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   240
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   241
theorem rev_rev [simp]: "rev(rev xs) = xs"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   242
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   243
txt\<open>Commands \isacom{theorem} and \isacom{lemma} are
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   244
interchangeable and merely indicate the importance we attach to a
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   245
proposition. Via the bracketed attribute \<open>simp\<close> we also tell Isabelle
55317
834a84553e02 started index
nipkow
parents: 54508
diff changeset
   246
to make the eventual theorem a \conceptnoidx{simplification rule}: future proofs
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   247
involving simplification will replace occurrences of \<^term>\<open>rev(rev xs)\<close> by
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   248
\<^term>\<open>xs\<close>. The proof is by induction:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   249
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   250
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   251
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   252
txt\<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   253
As explained above, we obtain two subgoals, namely the base case (\<^const>\<open>Nil\<close>) and the induction step (\<^const>\<open>Cons\<close>):
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   254
@{subgoals[display,indent=0,margin=65]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   255
Let us try to solve both goals automatically:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   256
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   257
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   258
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   259
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   260
txt\<open>Subgoal~1 is proved, and disappears; the simplified version
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   261
of subgoal~2 becomes the new subgoal~1:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   262
@{subgoals[display,indent=0,margin=70]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   263
In order to simplify this subgoal further, a lemma suggests itself.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   264
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   265
\subsubsection{A First Lemma}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   266
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   267
We insert the following lemma in front of the main theorem:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   268
\<close>
47269
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
oops
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   271
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   272
lemma rev_app [simp]: "rev(app xs ys) = app (rev ys) (rev xs)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   273
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   274
txt\<open>There are two variables that we could induct on: \<open>xs\<close> and
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   275
\<open>ys\<close>. Because \<^const>\<open>app\<close> is defined by recursion on
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   276
the first argument, \<open>xs\<close> is the correct one:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   277
\<close>
47269
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
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   280
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   281
txt\<open>This time not even the base case is solved automatically:\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   282
apply(auto)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   283
txt\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   284
\vspace{-5ex}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   285
@{subgoals[display,goals_limit=1]}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   286
Again, we need to abandon this proof attempt and prove another simple lemma
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   287
first.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   288
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   289
\subsubsection{A Second Lemma}
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
We again try the canonical proof procedure:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   292
\<close>
47269
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
oops
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   295
(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   296
lemma app_Nil2 [simp]: "app xs Nil = xs"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   297
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   298
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   299
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   300
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   301
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   302
Thankfully, this worked.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   303
Now we can continue with our stuck proof attempt of the first lemma:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   304
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   305
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   306
lemma rev_app [simp]: "rev(app xs ys) = app (rev ys) (rev xs)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   307
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   308
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   309
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   310
txt\<open>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   311
We find that this time \<open>auto\<close> solves the base case, but the
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   312
induction step merely simplifies to
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   313
@{subgoals[display,indent=0,goals_limit=1]}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   314
The missing lemma is associativity of \<^const>\<open>app\<close>,
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   315
which we insert in front of the failed lemma \<open>rev_app\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   316
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   317
\subsubsection{Associativity of \<^const>\<open>app\<close>}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   318
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   319
The canonical proof procedure succeeds without further ado:
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   320
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   321
(*<*)oops(*>*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   322
lemma app_assoc [simp]: "app (app xs ys) zs = app xs (app ys zs)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   323
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   324
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   325
done
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
lemma rev_app [simp]: "rev(app xs ys) = app (rev ys)(rev xs)"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   328
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   329
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   330
done
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
theorem rev_rev [simp]: "rev(rev xs) = xs"
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   333
apply(induction xs)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   334
apply(auto)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   335
done
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   336
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   337
text\<open>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   338
Finally the proofs of @{thm[source] rev_app} and @{thm[source] rev_rev}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   339
succeed, too.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   340
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 48985
diff changeset
   341
\subsubsection{Another Informal Proof}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   342
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   343
Here is the informal proof of associativity of \<^const>\<open>app\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   344
corresponding to the Isabelle proof above.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   345
\bigskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   346
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   347
\noindent
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   348
\textbf{Lemma} \<^prop>\<open>app (app xs ys) zs = app xs (app ys zs)\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   349
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   350
\noindent
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   351
\textbf{Proof} by induction on \<open>xs\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   352
\begin{itemize}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   353
\item Case \<open>Nil\<close>: \ \<^prop>\<open>app (app Nil ys) zs = app ys zs\<close> \<open>=\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   354
  \mbox{\<^term>\<open>app Nil (app ys zs)\<close>} \ holds by definition of \<open>app\<close>.
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   355
\item Case \<open>Cons x xs\<close>: We assume
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   356
  \begin{center} \hfill \<^term>\<open>app (app xs ys) zs\<close> \<open>=\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   357
  \<^term>\<open>app xs (app ys zs)\<close> \hfill (IH) \end{center}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   358
  and we need to show
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   359
  \begin{center} \<^prop>\<open>app (app (Cons x xs) ys) zs = app (Cons x xs) (app ys zs)\<close>.\end{center}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   360
  The proof is as follows:\smallskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   361
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   362
  \begin{tabular}{@ {}l@ {\quad}l@ {}}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   363
  \<^term>\<open>app (app (Cons x xs) ys) zs\<close>\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   364
  \<open>= app (Cons x (app xs ys)) zs\<close> & by definition of \<open>app\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   365
  \<open>= Cons x (app (app xs ys) zs)\<close> & by definition of \<open>app\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   366
  \<open>= Cons x (app xs (app ys zs))\<close> & by IH\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   367
  \<open>= app (Cons x xs) (app ys zs)\<close> & by definition of \<open>app\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   368
  \end{tabular}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   369
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   370
\medskip
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   371
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   372
\noindent Didn't we say earlier that all proofs are by simplification? But
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   373
in both cases, going from left to right, the last equality step is not a
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   374
simplification at all! In the base case it is \<^prop>\<open>app ys zs = app Nil (app
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   375
ys zs)\<close>. It appears almost mysterious because we suddenly complicate the
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   376
term by appending \<open>Nil\<close> on the left. What is really going on is this:
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   377
when proving some equality \mbox{\<^prop>\<open>s = t\<close>}, both \<open>s\<close> and \<open>t\<close> are
54467
663a927fdc88 comments by Sean Seefried
nipkow
parents: 54436
diff changeset
   378
simplified until they ``meet in the middle''. This heuristic for equality proofs
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   379
works well for a functional programming context like ours. In the base case
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   380
both \<^term>\<open>app (app Nil ys) zs\<close> and \<^term>\<open>app Nil (app
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   381
ys zs)\<close> are simplified to \<^term>\<open>app ys zs\<close>, the term in the middle.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   382
52361
7d5ad23b8245 all headings in upper case
nipkow
parents: 48985
diff changeset
   383
\subsection{Predefined Lists}
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   384
\label{sec:predeflists}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   385
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   386
Isabelle's predefined lists are the same as the ones above, but with
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   387
more syntactic sugar:
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   388
\begin{itemize}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   389
\item \<open>[]\<close> is \indexed{\<^const>\<open>Nil\<close>}{Nil},
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   390
\item \<^term>\<open>x # xs\<close> is \<^term>\<open>Cons x xs\<close>\index{Cons@\<^const>\<open>Cons\<close>},
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   391
\item \<open>[x\<^sub>1, \<dots>, x\<^sub>n]\<close> is \<open>x\<^sub>1 # \<dots> # x\<^sub>n # []\<close>, and
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   392
\item \<^term>\<open>xs @ ys\<close> is \<^term>\<open>app xs ys\<close>.
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   393
\end{itemize}
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   394
There is also a large library of predefined functions.
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   395
The most important ones are the length function
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   396
\<open>length :: 'a list \<Rightarrow> nat\<close>\index{length@\<^const>\<open>length\<close>} (with the obvious definition),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   397
and the \indexed{\<^const>\<open>map\<close>}{map} function that applies a function to all elements of a list:
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   398
\begin{isabelle}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   399
\isacom{fun} \<^const>\<open>map\<close> \<open>::\<close> @{typ[source] "('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list"} \isacom{where}\\
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   400
\<open>"\<close>@{thm list.map(1) [of f]}\<open>" |\<close>\\
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   401
\<open>"\<close>@{thm list.map(2) [of f x xs]}\<open>"\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   402
\end{isabelle}
52782
b11d73dbfb76 tuned intro
nipkow
parents: 52718
diff changeset
   403
b11d73dbfb76 tuned intro
nipkow
parents: 52718
diff changeset
   404
\ifsem
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   405
Also useful are the \concept{head} of a list, its first element,
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   406
and the \concept{tail}, the rest of the list:
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   407
\begin{isabelle}\index{hd@\<^const>\<open>hd\<close>}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   408
\isacom{fun} \<open>hd :: 'a list \<Rightarrow> 'a\<close>\\
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   409
\<^prop>\<open>hd(x#xs) = x\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   410
\end{isabelle}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   411
\begin{isabelle}\index{tl@\<^const>\<open>tl\<close>}
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   412
\isacom{fun} \<open>tl :: 'a list \<Rightarrow> 'a list\<close>\\
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   413
\<^prop>\<open>tl [] = []\<close> \<open>|\<close>\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   414
\<^prop>\<open>tl(x#xs) = xs\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   415
\end{isabelle}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   416
Note that since HOL is a logic of total functions, \<^term>\<open>hd []\<close> is defined,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   417
but we do not know what the result is. That is, \<^term>\<open>hd []\<close> is not undefined
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   418
but underdefined.
52782
b11d73dbfb76 tuned intro
nipkow
parents: 52718
diff changeset
   419
\fi
47306
56d72c923281 made sure that " is shown in tutorial text
nipkow
parents: 47302
diff changeset
   420
%
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   421
52842
3682e1b7ce86 tuned exercises
nipkow
parents: 52782
diff changeset
   422
From now on lists are always the predefined lists.
3682e1b7ce86 tuned exercises
nipkow
parents: 52782
diff changeset
   423
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   424
\ifsem\else
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   425
\subsection{Types \<^typ>\<open>int\<close> and \<^typ>\<open>real\<close>}
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   426
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   427
In addition to \<^typ>\<open>nat\<close> there are also the types \<^typ>\<open>int\<close> and \<^typ>\<open>real\<close>, the mathematical integers
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   428
and real numbers. As mentioned above, numerals and most of the standard arithmetic operations are overloaded.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   429
In particular they are defined on \<^typ>\<open>int\<close> and \<^typ>\<open>real\<close>.
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   430
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   431
\begin{warn}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   432
There are two infix exponentiation operators:
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   433
\<^term>\<open>(^)\<close> for \<^typ>\<open>nat\<close> and \<^typ>\<open>int\<close> (with exponent of type \<^typ>\<open>nat\<close> in both cases)
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   434
and \<^term>\<open>(powr)\<close> for \<^typ>\<open>real\<close>.
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   435
\end{warn}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   436
\begin{warn}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   437
Type  \<^typ>\<open>int\<close> is already part of theory \<^theory>\<open>Main\<close>, but in order to use \<^typ>\<open>real\<close> as well, you have to import
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   438
theory \<^theory>\<open>Complex_Main\<close> instead of \<^theory>\<open>Main\<close>.
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   439
\end{warn}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   440
72656
a17c17ab931c tuned terminology
nipkow
parents: 69597
diff changeset
   441
There are three coercion functions that are inclusions and do not lose information:
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   442
\begin{quote}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   443
\begin{tabular}{rcl}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   444
\<^const>\<open>int\<close> &\<open>::\<close>& \<^typ>\<open>nat \<Rightarrow> int\<close>\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   445
\<^const>\<open>real\<close> &\<open>::\<close>& \<^typ>\<open>nat \<Rightarrow> real\<close>\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   446
\<^const>\<open>real_of_int\<close> &\<open>::\<close>& \<^typ>\<open>int \<Rightarrow> real\<close>\\
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   447
\end{tabular}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   448
\end{quote}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   449
72656
a17c17ab931c tuned terminology
nipkow
parents: 69597
diff changeset
   450
Isabelle inserts these inclusions automatically once you import \<open>Complex_Main\<close>.
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   451
If there are multiple type-correct completions, Isabelle chooses an arbitrary one.
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   452
For example, the input \noquotes{@{term[source] "(i::int) + (n::nat)"}} has the unique
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   453
type-correct completion \<^term>\<open>(i::int) + int(n::nat)\<close>. In contrast,
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   454
\noquotes{@{term[source] "((n::nat) + n) :: real"}} has two type-correct completions,
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   455
\noquotes{@{term[source]"real(n+n)"}} and \noquotes{@{term[source]"real n + real n"}}.
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   456
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   457
There are also the coercion functions in the other direction:
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   458
\begin{quote}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   459
\begin{tabular}{rcl}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   460
\<^const>\<open>nat\<close> &\<open>::\<close>& \<^typ>\<open>int \<Rightarrow> nat\<close>\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   461
\<^const>\<open>floor\<close> &\<open>::\<close>& \<^typ>\<open>real \<Rightarrow> int\<close>\\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   462
\<^const>\<open>ceiling\<close> &\<open>::\<close>& \<^typ>\<open>real \<Rightarrow> int\<close>\\
68919
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   463
\end{tabular}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   464
\end{quote}
027219002f32 added int and real
nipkow
parents: 67406
diff changeset
   465
\fi
52842
3682e1b7ce86 tuned exercises
nipkow
parents: 52782
diff changeset
   466
54436
nipkow
parents: 54121
diff changeset
   467
\subsection*{Exercises}
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   468
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   469
\begin{exercise}
54121
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   470
Use the \isacom{value} command to evaluate the following expressions:
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   471
@{term[source] "1 + (2::nat)"}, @{term[source] "1 + (2::int)"},
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   472
@{term[source] "1 - (2::nat)"} and @{term[source] "1 - (2::int)"}.
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   473
\end{exercise}
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   474
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   475
\begin{exercise}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   476
Start from the definition of \<^const>\<open>add\<close> given above.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   477
Prove that \<^const>\<open>add\<close> is associative and commutative.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   478
Define a recursive function \<open>double\<close> \<open>::\<close> \<^typ>\<open>nat \<Rightarrow> nat\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   479
and prove \<^prop>\<open>double m = add m m\<close>.
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   480
\end{exercise}
52718
0faf89b8d928 tuned exercises
nipkow
parents: 52593
diff changeset
   481
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   482
\begin{exercise}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   483
Define a function \<open>count ::\<close> \<^typ>\<open>'a \<Rightarrow> 'a list \<Rightarrow> nat\<close>
52842
3682e1b7ce86 tuned exercises
nipkow
parents: 52782
diff changeset
   484
that counts the number of occurrences of an element in a list. Prove
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   485
\<^prop>\<open>count x xs \<le> length xs\<close>.
52842
3682e1b7ce86 tuned exercises
nipkow
parents: 52782
diff changeset
   486
\end{exercise}
3682e1b7ce86 tuned exercises
nipkow
parents: 52782
diff changeset
   487
3682e1b7ce86 tuned exercises
nipkow
parents: 52782
diff changeset
   488
\begin{exercise}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   489
Define a recursive function \<open>snoc ::\<close> \<^typ>\<open>'a list \<Rightarrow> 'a \<Rightarrow> 'a list\<close>
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   490
that appends an element to the end of a list. With the help of \<open>snoc\<close>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   491
define a recursive function \<open>reverse ::\<close> \<^typ>\<open>'a list \<Rightarrow> 'a list\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   492
that reverses a list. Prove \<^prop>\<open>reverse(reverse xs) = xs\<close>.
54121
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   493
\end{exercise}
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   494
4e7d71037bb6 tuned exercises
nipkow
parents: 53015
diff changeset
   495
\begin{exercise}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   496
Define a recursive function \<open>sum_upto ::\<close> \<^typ>\<open>nat \<Rightarrow> nat\<close> such that
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68919
diff changeset
   497
\mbox{\<open>sum_upto n\<close>} \<open>=\<close> \<open>0 + ... + n\<close> and prove
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   498
\<^prop>\<open> sum_upto (n::nat) = n * (n+1) div 2\<close>.
52593
aedf7b01c6e4 added exercises
nipkow
parents: 52361
diff changeset
   499
\end{exercise}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 65514
diff changeset
   500
\<close>
47269
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   501
(*<*)
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   502
end
29aa0c071875 New manual Programming and Proving in Isabelle/HOL
nipkow
parents:
diff changeset
   503
(*>*)