76650
|
1 |
theory Partial_Functions
|
76673
|
2 |
imports Setup "HOL-Library.Monad_Syntax"
|
76650
|
3 |
begin
|
|
4 |
|
|
5 |
section \<open>Partial Functions \label{sec:partial}\<close>
|
|
6 |
|
76673
|
7 |
text \<open>We demonstrate three approaches to defining executable partial recursive functions,
|
|
8 |
i.e.\ functions that do not terminate for all inputs.
|
76650
|
9 |
The main points are the definitions of the functions and the inductive proofs about them.
|
|
10 |
|
|
11 |
Our concrete example represents a typical termination problem: following a data structure that
|
|
12 |
may contain cycles. We want to follow a mapping from \<open>nat\<close> to \<open>nat\<close> to the end
|
|
13 |
(until we leave its domain). The mapping is represented by a list \<open>ns :: nat list\<close>
|
|
14 |
that maps \<open>n\<close> to \<open>ns ! n\<close>.
|
|
15 |
The details of the example are in some sense irrelevant but make the exposition more realistic.
|
|
16 |
However, we hide most proofs or show only the characteristic opening.\<close>
|
|
17 |
|
|
18 |
text \<open>The list representation of the mapping can be abstracted to a relation.
|
|
19 |
The order @{term "(ns!n, n)"} is the order that @{const wf} expects.\<close>
|
|
20 |
|
|
21 |
definition rel :: "nat list \<Rightarrow> (nat * nat) set" where
|
|
22 |
"rel ns = set(zip ns [0..<length ns])"
|
|
23 |
|
|
24 |
lemma finite_rel[simp]: "finite(rel ns)"
|
|
25 |
(*<*)by(simp add: rel_def)(*>*)
|
|
26 |
|
76673
|
27 |
text \<open>\noindent This relation should be acyclic
|
76650
|
28 |
to guarantee termination of the partial functions defined below.\<close>
|
|
29 |
|
|
30 |
|
|
31 |
subsection \<open>Tail recursion\<close>
|
|
32 |
|
|
33 |
text \<open>If a function is tail-recursive, an executable definition is easy:\<close>
|
|
34 |
|
|
35 |
partial_function (tailrec) follow :: "nat list \<Rightarrow> nat \<Rightarrow> nat" where
|
|
36 |
"follow ns n = (if n < length ns then follow ns (ns!n) else n)"
|
|
37 |
|
|
38 |
text \<open>Informing the code generator:\<close>
|
|
39 |
|
|
40 |
declare follow.simps[code]
|
|
41 |
|
|
42 |
text \<open>Now @{const follow} is executable:\<close>
|
|
43 |
|
|
44 |
value "follow [1,2,3] 0"
|
|
45 |
|
|
46 |
text \<open>For proofs about @{const follow} we need a @{const wf} relation on @{term "(ns,n)"} pairs
|
|
47 |
that decreases with each recursive call. The first component stays the same but must be acyclic.
|
|
48 |
The second component must decrease w.r.t @{const rel}:\<close>
|
|
49 |
|
|
50 |
definition "rel_follow = same_fst (acyclic o rel) rel"
|
|
51 |
|
|
52 |
lemma wf_follow: "wf rel_follow"
|
|
53 |
(*<*)
|
|
54 |
by (auto intro: wf_same_fst simp: rel_follow_def finite_acyclic_wf)
|
|
55 |
|
|
56 |
text \<open>A more explicit formulation of \<open>rel_follow\<close>:
|
|
57 |
The first component stays the same but must be acyclic.
|
|
58 |
The second component decreases w.r.t \<open>rel\<close>:\<close>
|
|
59 |
|
|
60 |
lemma rel_follow_step:
|
|
61 |
"\<lbrakk> m < length ms; acyclic (rel ms) \<rbrakk> \<Longrightarrow> ((ms, ms ! m), (ms, m)) \<in> rel_follow"
|
|
62 |
by(force simp: rel_follow_def rel_def in_set_zip)
|
|
63 |
(*>*)
|
|
64 |
|
|
65 |
text \<open>This is how you start an inductive proof about @{const follow} along @{const rel_follow}:\<close>
|
|
66 |
|
|
67 |
lemma "acyclic(rel ms) \<Longrightarrow> follow ms m = n \<Longrightarrow> length ms \<le> n"
|
|
68 |
proof (induction "(ms,m)" arbitrary: m n rule: wf_induct_rule[OF wf_follow])
|
|
69 |
(*<*)
|
|
70 |
case 1
|
|
71 |
thus ?case using follow.simps rel_follow_step by fastforce
|
|
72 |
qed
|
|
73 |
(*>*)
|
|
74 |
|
|
75 |
|
|
76 |
subsection \<open>Option\<close>
|
|
77 |
|
|
78 |
text \<open>If the function is not tail-recursive, not all is lost: if we rewrite it to return an \<open>option\<close>
|
|
79 |
type, it can still be defined. In this case @{term "Some x"} represents the result \<open>x\<close>
|
|
80 |
and @{term None} represents represents nontermination.
|
|
81 |
For example, counting the length of the chain represented by \<open>ns\<close> can be defined like this:\<close>
|
|
82 |
|
|
83 |
partial_function (option) count :: "nat list \<Rightarrow> nat \<Rightarrow> nat option" where
|
|
84 |
"count ns n
|
|
85 |
= (if n < length ns then do {k \<leftarrow> count ns (ns!n); Some (k+1)} else Some 0)"
|
|
86 |
|
76673
|
87 |
text \<open>\noindent We use a Haskell-like \<open>do\<close> notation (import @{theory "HOL-Library.Monad_Syntax"})
|
76650
|
88 |
to abbreviate the clumsy explicit
|
|
89 |
|
|
90 |
\noindent
|
|
91 |
@{term "case count ns (ns!n) of Some k \<Rightarrow> Some(k+1) | None \<Rightarrow> None"}.
|
|
92 |
|
|
93 |
\noindent
|
|
94 |
The branch \<open>None \<Rightarrow> None\<close> represents the requirement
|
|
95 |
that nontermination of a recursive call must lead to nontermination of the function.
|
|
96 |
|
|
97 |
Now we can prove that @{const count} terminates for all acyclic maps:\<close>
|
|
98 |
|
|
99 |
lemma "acyclic(rel ms) \<Longrightarrow> \<exists>k. count ms m = Some k"
|
|
100 |
proof (induction "(ms,m)" arbitrary: ms m rule: wf_induct_rule[OF wf_follow])
|
|
101 |
(*<*)
|
|
102 |
case 1
|
|
103 |
thus ?case
|
|
104 |
by (metis bind.bind_lunit count.simps rel_follow_step)
|
|
105 |
qed
|
|
106 |
(*>*)
|
|
107 |
|
|
108 |
|
|
109 |
subsection \<open>Subtype\<close>
|
|
110 |
|
|
111 |
text \<open>In this approach we define a new type that contains only elements on which the function
|
|
112 |
in question terminates.
|
|
113 |
In our example that is the subtype of all \<open>ns :: nat list\<close> such that @{term "rel ns"} is acyclic.
|
|
114 |
Then @{const follow} can be defined as a total function on that subtype.\<close>
|
|
115 |
|
|
116 |
text \<open>The subtype is not empty:\<close>
|
|
117 |
|
|
118 |
lemma acyclic_rel_Nil: "acyclic(rel [])"
|
|
119 |
(*<*)by (simp add: rel_def acyclic_def)(*>*)
|
|
120 |
|
|
121 |
text \<open>Definition of subtype \<open>acyc\<close>:\<close>
|
|
122 |
|
|
123 |
typedef acyc = "{ns. acyclic(rel ns)}"
|
|
124 |
morphisms rep_acyc abs_acyc
|
|
125 |
using acyclic_rel_Nil by auto
|
|
126 |
|
76673
|
127 |
text \<open>\noindent This defines two functions @{term [source] "rep_acyc :: acyc \<Rightarrow> nat list"}
|
|
128 |
and @{const abs_acyc} \<open>::\<close> \mbox{@{typ "nat list \<Rightarrow> acyc"}}.
|
|
129 |
Function @{const abs_acyc} is only defined on acyclic lists and not executable for that reason.
|
|
130 |
Type \<open>dlist\<close> in Section~\ref{sec:partiality} is defined in the same manner.
|
|
131 |
|
|
132 |
The following command sets up infrastructure for lifting functions on @{typ "nat list"}
|
76987
|
133 |
to @{typ acyc} (used by @{command_def lift_definition} below) \<^cite>\<open>"isabelle-isar-ref"\<close>.\<close>
|
76650
|
134 |
|
|
135 |
setup_lifting type_definition_acyc
|
|
136 |
|
|
137 |
text \<open>This is how @{const follow} can be defined on @{typ acyc}:\<close>
|
|
138 |
|
|
139 |
function follow2 :: "acyc \<Rightarrow> nat \<Rightarrow> nat" where
|
|
140 |
"follow2 ac n
|
|
141 |
= (let ns = rep_acyc ac in if n < length ns then follow2 ac (ns!n) else n)"
|
|
142 |
by pat_completeness auto
|
|
143 |
|
76673
|
144 |
text \<open>Now we prepare for the termination proof.
|
76650
|
145 |
Relation \<open>rel_follow2\<close> is almost identical to @{const rel_follow}.\<close>
|
|
146 |
|
|
147 |
definition "rel_follow2 = same_fst (acyclic o rel o rep_acyc) (rel o rep_acyc)"
|
|
148 |
|
|
149 |
lemma wf_follow2: "wf rel_follow2"
|
|
150 |
(*<*)
|
|
151 |
by (auto intro: wf_same_fst simp add: rel_follow2_def finite_acyclic_wf)
|
|
152 |
|
|
153 |
text \<open>A more explicit formulation of \<open>rel_follow2\<close>:\<close>
|
|
154 |
|
|
155 |
lemma rel_follow2_step:
|
|
156 |
"\<lbrakk> ns = rep_acyc ac; m < length ns; acyclic (rel ns) \<rbrakk> \<Longrightarrow> ((ac, ns ! m), (ac, m)) \<in> rel_follow2"
|
|
157 |
by(force simp add: rel_follow2_def rel_def in_set_zip)
|
|
158 |
(*>*)
|
|
159 |
|
|
160 |
text \<open>Here comes the actual termination proof:\<close>
|
|
161 |
|
|
162 |
termination follow2
|
|
163 |
proof
|
|
164 |
show "wf rel_follow2"
|
|
165 |
(*<*) by (auto intro: wf_same_fst simp add: rel_follow2_def finite_acyclic_wf)(*>*)
|
|
166 |
|
|
167 |
next
|
|
168 |
show "\<And>ac n ns. \<lbrakk> ns = rep_acyc ac; n < length ns \<rbrakk>
|
|
169 |
\<Longrightarrow> ((ac, ns ! n), (ac, n)) \<in> rel_follow2"
|
|
170 |
(*<*) using rel_follow2_step rep_acyc by simp(*>*)
|
|
171 |
|
|
172 |
qed
|
|
173 |
|
|
174 |
text \<open>Inductive proofs about @{const follow2} can now simply use computation induction:\<close>
|
|
175 |
|
|
176 |
lemma "follow2 ac m = n \<Longrightarrow> length (rep_acyc ac) \<le> n"
|
|
177 |
proof (induction ac m arbitrary: n rule: follow2.induct)
|
|
178 |
(*<*)
|
|
179 |
case 1
|
|
180 |
thus ?case by (metis (full_types) follow2.simps linorder_not_le)
|
|
181 |
qed
|
|
182 |
(*>*)
|
|
183 |
|
|
184 |
text \<open>A complication with the subtype approach is that injection into the subtype
|
|
185 |
(function \<open>abs_acyc\<close> in our example) is not executable. But to call @{const follow2},
|
|
186 |
we need an argument of type \<open>acyc\<close> and we need to obtain it in an executable manner.
|
|
187 |
There are two approaches.\<close>
|
|
188 |
|
|
189 |
text \<open>In the first approach we check wellformedness (i.e. acyclicity) explicitly.
|
|
190 |
This check needs to be executable (which @{const acyclic} and @{const rel} are).
|
|
191 |
If the check fails, @{term "[]"} is returned (which is acyclic).\<close>
|
|
192 |
|
|
193 |
lift_definition is_acyc :: "nat list \<Rightarrow> acyc" is
|
|
194 |
"\<lambda>ns. if acyclic(rel ns) then ns else []"
|
|
195 |
(*<*)by (auto simp: acyclic_rel_Nil)(*>*)
|
|
196 |
|
76673
|
197 |
text \<open>\noindent This works because we can easily prove that for any \<open>ns\<close>,
|
76650
|
198 |
the \<open>\<lambda>\<close>-term produces an acyclic list.
|
|
199 |
But it requires the possibly expensive check @{prop "acyclic(rel ns)"}.\<close>
|
|
200 |
|
|
201 |
definition "follow_test ns n = follow2 (is_acyc ns) n"
|
|
202 |
|
|
203 |
text \<open>The relation is acyclic (a chain):\<close>
|
|
204 |
|
|
205 |
value "follow_test [1,2,3] 1"
|
|
206 |
|
|
207 |
text \<open>In the second approach, wellformedness of the argument is guaranteed by construction.
|
|
208 |
In our example \<open>[1..<n+1]\<close> represents an acyclic chain \<open>i \<mapsto> i+1\<close>\<close>
|
|
209 |
|
|
210 |
lemma acyclic_chain: "acyclic (rel [1..<n+1])"
|
|
211 |
(*<*)
|
|
212 |
apply(induction n)
|
|
213 |
apply (simp add: rel_def acyclic_def)
|
|
214 |
apply (auto simp add: rel_def)
|
|
215 |
by (metis atLeast_upt lessI lessThan_iff order_less_asym' rtranclE set_zip_rightD)
|
|
216 |
(*>*)
|
|
217 |
|
|
218 |
text \<open>\<close>
|
|
219 |
lift_definition acyc_chain :: "nat \<Rightarrow> acyc" is "\<lambda>n. [1..<n+1]"
|
|
220 |
(*<*)by (use acyclic_chain in auto)(*>*)
|
|
221 |
|
|
222 |
text \<open>\<close>
|
|
223 |
definition "follow_chain m n = follow2 (acyc_chain m) n"
|
|
224 |
|
|
225 |
value "follow_chain 5 1"
|
|
226 |
|
|
227 |
text \<open>The subtype approach requires neither tail-recursion nor \<open>option\<close> but you cannot easily modify
|
|
228 |
arguments of the subtype except via existing functions on the subtype. Otherwise you need to inject
|
|
229 |
some value into the subtype and that injection is not computable.
|
|
230 |
\<close>
|
|
231 |
(*<*)
|
|
232 |
text \<open>The problem with subtypes: the Abs function must not be used explicitly.
|
|
233 |
This can be avoided by using \<open>lift_definition\<close>. Example:\<close>
|
|
234 |
|
|
235 |
typedef nat1 = "{n::nat. n > 0}"
|
|
236 |
by auto
|
|
237 |
print_theorems
|
|
238 |
|
|
239 |
setup_lifting type_definition_nat1
|
|
240 |
|
|
241 |
(* just boiler plate *)
|
|
242 |
lift_definition mk1 :: "nat \<Rightarrow> nat1" is
|
|
243 |
"\<lambda>n. if n>0 then n else 1"
|
|
244 |
by auto
|
|
245 |
|
|
246 |
lift_definition g :: "nat1 \<Rightarrow> nat1" is
|
|
247 |
"\<lambda>n. if n \<ge> 2 then n-1 else n"
|
|
248 |
by auto
|
|
249 |
print_theorems
|
|
250 |
text \<open>This generates
|
|
251 |
\<open>g.rep_eq: Rep_nat1 (g x) = (if 2 \<le> Rep_nat1 x then Rep_nat1 x - 1 else Rep_nat1 x)\<close>
|
|
252 |
which is acceptable as an abstract code eqn. The manual definition of
|
|
253 |
\<open>g n1 = (let n = Rep_nat1 n1 in if 2 \<le> n then Abs_nat1 (n - 1) else Abs_nat1 n)\<close>
|
|
254 |
looks non-executable but \<open>g.rep_eq\<close> can be derived from it.\<close>
|
|
255 |
|
|
256 |
value "g (mk1 3)"
|
|
257 |
|
|
258 |
text \<open>However, this trick does not work when passing an Abs-term as an argument,
|
|
259 |
eg in a recursive call, because the Abs-term is `hidden' by the function call.\<close>
|
|
260 |
(*>*)
|
|
261 |
end |