doc-src/TutorialI/Advanced/Partial.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 12 Aug 2010 15:03:34 +0200
branchisac-from-Isabelle2009-2
changeset 37913 20e3616b2d9c
parent 35413 d8d7d1b785af
permissions -rw-r--r--
prepare reactivation of isac-update-Isa09-2
wenzelm@17914
     1
(*<*)theory Partial imports While_Combinator begin(*>*)
nipkow@10654
     2
paulson@11494
     3
text{*\noindent Throughout this tutorial, we have emphasized
paulson@11494
     4
that all functions in HOL are total.  We cannot hope to define
paulson@11310
     5
truly partial functions, but must make them total.  A straightforward
paulson@11310
     6
method is to lift the result type of the function from $\tau$ to
nipkow@11277
     7
$\tau$~@{text option} (see \ref{sec:option}), where @{term None} is
nipkow@11277
     8
returned if the function is applied to an argument not in its
nipkow@11277
     9
domain. Function @{term assoc} in \S\ref{sec:Trie} is a simple example.
nipkow@11277
    10
We do not pursue this schema further because it should be clear
nipkow@11277
    11
how it works. Its main drawback is that the result of such a lifted
nipkow@11277
    12
function has to be unpacked first before it can be processed
nipkow@11277
    13
further. Its main advantage is that you can distinguish if the
nipkow@11277
    14
function was applied to an argument in its domain or not. If you do
nipkow@11277
    15
not need to make this distinction, for example because the function is
nipkow@11277
    16
never used outside its domain, it is easier to work with
paulson@11428
    17
\emph{underdefined}\index{functions!underdefined} functions: for
nipkow@11277
    18
certain arguments we only know that a result exists, but we do not
nipkow@11277
    19
know what it is. When defining functions that are normally considered
nipkow@11277
    20
partial, underdefinedness turns out to be a very reasonable
nipkow@11277
    21
alternative.
nipkow@10654
    22
nipkow@10654
    23
We have already seen an instance of underdefinedness by means of
nipkow@10654
    24
non-exhaustive pattern matching: the definition of @{term last} in
nipkow@25258
    25
\S\ref{sec:fun}. The same is allowed for \isacommand{primrec}
nipkow@10654
    26
*}
nipkow@10654
    27
nipkow@10654
    28
consts hd :: "'a list \<Rightarrow> 'a"
nipkow@10654
    29
primrec "hd (x#xs) = x"
nipkow@10654
    30
nipkow@10654
    31
text{*\noindent
nipkow@10654
    32
although it generates a warning.
nipkow@10654
    33
Even ordinary definitions allow underdefinedness, this time by means of
nipkow@10654
    34
preconditions:
nipkow@10654
    35
*}
nipkow@10654
    36
haftmann@35413
    37
definition subtract :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
wenzelm@19248
    38
"n \<le> m \<Longrightarrow> subtract m n \<equiv> m - n"
nipkow@10654
    39
nipkow@10654
    40
text{*
nipkow@10654
    41
The rest of this section is devoted to the question of how to define
nipkow@11256
    42
partial recursive functions by other means than non-exhaustive pattern
nipkow@10654
    43
matching.
nipkow@10654
    44
*}
nipkow@10654
    45
paulson@10885
    46
subsubsection{*Guarded Recursion*}
nipkow@10654
    47
paulson@11494
    48
text{* 
paulson@11494
    49
\index{recursion!guarded}%
paulson@11494
    50
Neither \isacommand{primrec} nor \isacommand{recdef} allow to
nipkow@10654
    51
prefix an equation with a condition in the way ordinary definitions do
wenzelm@19248
    52
(see @{const subtract} above). Instead we have to move the condition over
nipkow@10654
    53
to the right-hand side of the equation. Given a partial function $f$
nipkow@10654
    54
that should satisfy the recursion equation $f(x) = t$ over its domain
nipkow@10654
    55
$dom(f)$, we turn this into the \isacommand{recdef}
nipkow@10654
    56
@{prop[display]"f(x) = (if x \<in> dom(f) then t else arbitrary)"}
nipkow@10654
    57
where @{term arbitrary} is a predeclared constant of type @{typ 'a}
nipkow@10654
    58
which has no definition. Thus we know nothing about its value,
nipkow@10654
    59
which is ideal for specifying underdefined functions on top of it.
nipkow@10654
    60
nipkow@10654
    61
As a simple example we define division on @{typ nat}:
nipkow@10654
    62
*}
nipkow@10654
    63
nipkow@10654
    64
consts divi :: "nat \<times> nat \<Rightarrow> nat"
nipkow@12334
    65
recdef divi "measure(\<lambda>(m,n). m)"
nipkow@12334
    66
  "divi(m,0) = arbitrary"
nipkow@12334
    67
  "divi(m,n) = (if m < n then 0 else divi(m-n,n)+1)"
nipkow@10654
    68
nipkow@10654
    69
text{*\noindent Of course we could also have defined
nipkow@10654
    70
@{term"divi(m,0)"} to be some specific number, for example 0. The
nipkow@10654
    71
latter option is chosen for the predefined @{text div} function, which
paulson@10885
    72
simplifies proofs at the expense of deviating from the
paulson@10885
    73
standard mathematical division function.
nipkow@10654
    74
nipkow@10654
    75
As a more substantial example we consider the problem of searching a graph.
nipkow@11277
    76
For simplicity our graph is given by a function @{term f} of
nipkow@10654
    77
type @{typ"'a \<Rightarrow> 'a"} which
nipkow@12334
    78
maps each node to its successor; the graph has out-degree 1.
nipkow@11196
    79
The task is to find the end of a chain, modelled by a node pointing to
nipkow@11196
    80
itself. Here is a first attempt:
nipkow@10654
    81
@{prop[display]"find(f,x) = (if f x = x then x else find(f, f x))"}
nipkow@12334
    82
This may be viewed as a fixed point finder or as the second half of the well
nipkow@12334
    83
known \emph{Union-Find} algorithm.
nipkow@11149
    84
The snag is that it may not terminate if @{term f} has non-trivial cycles.
nipkow@10654
    85
Phrased differently, the relation
nipkow@10654
    86
*}
nipkow@10654
    87
haftmann@35413
    88
definition step1 :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<times> 'a)set" where
nipkow@10654
    89
  "step1 f \<equiv> {(y,x). y = f x \<and> y \<noteq> x}"
nipkow@10654
    90
nipkow@10654
    91
text{*\noindent
nipkow@10654
    92
must be well-founded. Thus we make the following definition:
nipkow@10654
    93
*}
nipkow@10654
    94
nipkow@10654
    95
consts find :: "('a \<Rightarrow> 'a) \<times> 'a \<Rightarrow> 'a"
nipkow@10654
    96
recdef find "same_fst (\<lambda>f. wf(step1 f)) step1"
nipkow@10654
    97
  "find(f,x) = (if wf(step1 f)
nipkow@10654
    98
                then if f x = x then x else find(f, f x)
nipkow@10654
    99
                else arbitrary)"
nipkow@11285
   100
(hints recdef_simp: step1_def)
nipkow@10654
   101
nipkow@10654
   102
text{*\noindent
nipkow@10654
   103
The recursion equation itself should be clear enough: it is our aborted
nipkow@10654
   104
first attempt augmented with a check that there are no non-trivial loops.
nipkow@11277
   105
To express the required well-founded relation we employ the
nipkow@11196
   106
predefined combinator @{term same_fst} of type
nipkow@10654
   107
@{text[display]"('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> ('b\<times>'b)set) \<Rightarrow> (('a\<times>'b) \<times> ('a\<times>'b))set"}
nipkow@10654
   108
defined as
nipkow@10654
   109
@{thm[display]same_fst_def[no_vars]}
nipkow@11196
   110
This combinator is designed for
nipkow@11196
   111
recursive functions on pairs where the first component of the argument is
nipkow@11196
   112
passed unchanged to all recursive calls. Given a constraint on the first
nipkow@11196
   113
component and a relation on the second component, @{term same_fst} builds the
nipkow@11196
   114
required relation on pairs.  The theorem
nipkow@11196
   115
@{thm[display]wf_same_fst[no_vars]}
nipkow@11196
   116
is known to the well-foundedness prover of \isacommand{recdef}.  Thus
nipkow@11196
   117
well-foundedness of the relation given to \isacommand{recdef} is immediate.
nipkow@11196
   118
Furthermore, each recursive call descends along that relation: the first
nipkow@11196
   119
argument stays unchanged and the second one descends along @{term"step1
haftmann@15904
   120
f"}. The proof requires unfolding the definition of @{const step1},
nipkow@11285
   121
as specified in the \isacommand{hints} above.
nipkow@10654
   122
paulson@11494
   123
Normally you will then derive the following conditional variant from
paulson@11494
   124
the recursion equation:
nipkow@10654
   125
*}
nipkow@10654
   126
nipkow@10654
   127
lemma [simp]:
nipkow@10654
   128
  "wf(step1 f) \<Longrightarrow> find(f,x) = (if f x = x then x else find(f, f x))"
nipkow@10654
   129
by simp
nipkow@10654
   130
paulson@11494
   131
text{*\noindent Then you should disable the original recursion equation:*}
nipkow@10654
   132
nipkow@10654
   133
declare find.simps[simp del]
nipkow@10654
   134
nipkow@10654
   135
text{*
paulson@11494
   136
Reasoning about such underdefined functions is like that for other
paulson@11494
   137
recursive functions.  Here is a simple example of recursion induction:
nipkow@10654
   138
*}
nipkow@10654
   139
nipkow@10654
   140
lemma "wf(step1 f) \<longrightarrow> f(find(f,x)) = find(f,x)"
wenzelm@12815
   141
apply(induct_tac f x rule: find.induct);
nipkow@10654
   142
apply simp
nipkow@10654
   143
done
nipkow@10654
   144
paulson@10885
   145
subsubsection{*The {\tt\slshape while} Combinator*}
nipkow@10654
   146
nipkow@10654
   147
text{*If the recursive function happens to be tail recursive, its
paulson@11428
   148
definition becomes a triviality if based on the predefined \cdx{while}
nipkow@12332
   149
combinator.  The latter lives in the Library theory \thydx{While_Combinator}.
nipkow@12332
   150
% which is not part of {text Main} but needs to
nipkow@12332
   151
% be included explicitly among the ancestor theories.
nipkow@10654
   152
nipkow@10654
   153
Constant @{term while} is of type @{text"('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a"}
nipkow@10654
   154
and satisfies the recursion equation @{thm[display]while_unfold[no_vars]}
nipkow@10654
   155
That is, @{term"while b c s"} is equivalent to the imperative program
nipkow@10654
   156
\begin{verbatim}
nipkow@10654
   157
     x := s; while b(x) do x := c(x); return x
nipkow@10654
   158
\end{verbatim}
paulson@11494
   159
In general, @{term s} will be a tuple or record.  As an example
haftmann@15904
   160
consider the following definition of function @{const find}:
nipkow@10654
   161
*}
nipkow@10654
   162
haftmann@35413
   163
definition find2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
nipkow@10654
   164
  "find2 f x \<equiv>
nipkow@10654
   165
   fst(while (\<lambda>(x,x'). x' \<noteq> x) (\<lambda>(x,x'). (x',f x')) (x,f x))"
nipkow@10654
   166
nipkow@10654
   167
text{*\noindent
nipkow@10654
   168
The loop operates on two ``local variables'' @{term x} and @{term x'}
nipkow@10654
   169
containing the ``current'' and the ``next'' value of function @{term f}.
paulson@11310
   170
They are initialized with the global @{term x} and @{term"f x"}. At the
nipkow@10654
   171
end @{term fst} selects the local @{term x}.
nipkow@10654
   172
nipkow@11158
   173
Although the definition of tail recursive functions via @{term while} avoids
nipkow@11158
   174
termination proofs, there is no free lunch. When proving properties of
nipkow@11158
   175
functions defined by @{term while}, termination rears its ugly head
paulson@11494
   176
again. Here is \tdx{while_rule}, the well known proof rule for total
nipkow@10654
   177
correctness of loops expressed with @{term while}:
nipkow@11158
   178
@{thm[display,margin=50]while_rule[no_vars]} @{term P} needs to be true of
nipkow@11158
   179
the initial state @{term s} and invariant under @{term c} (premises 1
nipkow@11158
   180
and~2). The post-condition @{term Q} must become true when leaving the loop
nipkow@11158
   181
(premise~3). And each loop iteration must descend along a well-founded
nipkow@11158
   182
relation @{term r} (premises 4 and~5).
nipkow@10654
   183
haftmann@15904
   184
Let us now prove that @{const find2} does indeed find a fixed point. Instead
nipkow@10654
   185
of induction we apply the above while rule, suitably instantiated.
nipkow@10654
   186
Only the final premise of @{thm[source]while_rule} is left unproved
nipkow@10654
   187
by @{text auto} but falls to @{text simp}:
nipkow@10654
   188
*}
nipkow@10654
   189
nipkow@11277
   190
lemma lem: "wf(step1 f) \<Longrightarrow>
nipkow@11277
   191
  \<exists>y. while (\<lambda>(x,x'). x' \<noteq> x) (\<lambda>(x,x'). (x',f x')) (x,f x) = (y,y) \<and>
paulson@10885
   192
       f y = y"
nipkow@10654
   193
apply(rule_tac P = "\<lambda>(x,x'). x' = f x" and
nipkow@10654
   194
               r = "inv_image (step1 f) fst" in while_rule);
nipkow@10654
   195
apply auto
wenzelm@12815
   196
apply(simp add: inv_image_def step1_def)
nipkow@10654
   197
done
nipkow@10654
   198
nipkow@10654
   199
text{*
nipkow@10654
   200
The theorem itself is a simple consequence of this lemma:
nipkow@10654
   201
*}
nipkow@10654
   202
nipkow@10654
   203
theorem "wf(step1 f) \<Longrightarrow> f(find2 f x) = find2 f x"
nipkow@10654
   204
apply(drule_tac x = x in lem)
wenzelm@12815
   205
apply(auto simp add: find2_def)
nipkow@10654
   206
done
nipkow@10654
   207
nipkow@10654
   208
text{* Let us conclude this section on partial functions by a
nipkow@10654
   209
discussion of the merits of the @{term while} combinator. We have
paulson@11494
   210
already seen that the advantage of not having to
paulson@11310
   211
provide a termination argument when defining a function via @{term
nipkow@10654
   212
while} merely puts off the evil hour. On top of that, tail recursive
nipkow@10654
   213
functions tend to be more complicated to reason about. So why use
nipkow@10654
   214
@{term while} at all? The only reason is executability: the recursion
nipkow@10654
   215
equation for @{term while} is a directly executable functional
nipkow@10654
   216
program. This is in stark contrast to guarded recursion as introduced
nipkow@10654
   217
above which requires an explicit test @{prop"x \<in> dom f"} in the
nipkow@10654
   218
function body.  Unless @{term dom} is trivial, this leads to a
nipkow@11196
   219
definition that is impossible to execute or prohibitively slow.
paulson@10885
   220
Thus, if you are aiming for an efficiently executable definition
nipkow@10654
   221
of a partial function, you are likely to need @{term while}.
nipkow@10654
   222
*}
nipkow@10654
   223
nipkow@10654
   224
(*<*)end(*>*)