40107
|
1 |
(* Title: HOL/Partial_Function.thy
|
|
2 |
Author: Alexander Krauss, TU Muenchen
|
|
3 |
*)
|
|
4 |
|
|
5 |
header {* Partial Function Definitions *}
|
|
6 |
|
|
7 |
theory Partial_Function
|
|
8 |
imports Complete_Partial_Order Option
|
|
9 |
uses
|
|
10 |
"Tools/Function/function_lib.ML"
|
|
11 |
"Tools/Function/partial_function.ML"
|
|
12 |
begin
|
|
13 |
|
|
14 |
setup Partial_Function.setup
|
|
15 |
|
|
16 |
subsection {* Axiomatic setup *}
|
|
17 |
|
|
18 |
text {* This techical locale constains the requirements for function
|
|
19 |
definitions with ccpo fixed points. *}
|
|
20 |
|
|
21 |
definition "fun_ord ord f g \<longleftrightarrow> (\<forall>x. ord (f x) (g x))"
|
|
22 |
definition "fun_lub L A = (\<lambda>x. L {y. \<exists>f\<in>A. y = f x})"
|
|
23 |
definition "img_ord f ord = (\<lambda>x y. ord (f x) (f y))"
|
|
24 |
definition "img_lub f g Lub = (\<lambda>A. g (Lub (f ` A)))"
|
|
25 |
|
|
26 |
lemma call_mono[partial_function_mono]: "monotone (fun_ord ord) ord (\<lambda>f. f t)"
|
|
27 |
by (rule monotoneI) (auto simp: fun_ord_def)
|
|
28 |
|
|
29 |
lemma if_mono[partial_function_mono]: "monotone orda ordb F
|
|
30 |
\<Longrightarrow> monotone orda ordb G
|
|
31 |
\<Longrightarrow> monotone orda ordb (\<lambda>f. if c then F f else G f)"
|
|
32 |
unfolding monotone_def by simp
|
|
33 |
|
|
34 |
definition "mk_less R = (\<lambda>x y. R x y \<and> \<not> R y x)"
|
|
35 |
|
|
36 |
locale partial_function_definitions =
|
|
37 |
fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
|
|
38 |
fixes lub :: "'a set \<Rightarrow> 'a"
|
|
39 |
assumes leq_refl: "leq x x"
|
|
40 |
assumes leq_trans: "leq x y \<Longrightarrow> leq y z \<Longrightarrow> leq x z"
|
|
41 |
assumes leq_antisym: "leq x y \<Longrightarrow> leq y x \<Longrightarrow> x = y"
|
|
42 |
assumes lub_upper: "chain leq A \<Longrightarrow> x \<in> A \<Longrightarrow> leq x (lub A)"
|
|
43 |
assumes lub_least: "chain leq A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> leq x z) \<Longrightarrow> leq (lub A) z"
|
|
44 |
|
|
45 |
lemma partial_function_lift:
|
|
46 |
assumes "partial_function_definitions ord lb"
|
|
47 |
shows "partial_function_definitions (fun_ord ord) (fun_lub lb)" (is "partial_function_definitions ?ordf ?lubf")
|
|
48 |
proof -
|
|
49 |
interpret partial_function_definitions ord lb by fact
|
|
50 |
|
|
51 |
{ fix A a assume A: "chain ?ordf A"
|
|
52 |
have "chain ord {y. \<exists>f\<in>A. y = f a}" (is "chain ord ?C")
|
|
53 |
proof (rule chainI)
|
|
54 |
fix x y assume "x \<in> ?C" "y \<in> ?C"
|
|
55 |
then obtain f g where fg: "f \<in> A" "g \<in> A"
|
|
56 |
and [simp]: "x = f a" "y = g a" by blast
|
|
57 |
from chainD[OF A fg]
|
|
58 |
show "ord x y \<or> ord y x" unfolding fun_ord_def by auto
|
|
59 |
qed }
|
|
60 |
note chain_fun = this
|
|
61 |
|
|
62 |
show ?thesis
|
|
63 |
proof
|
|
64 |
fix x show "?ordf x x"
|
|
65 |
unfolding fun_ord_def by (auto simp: leq_refl)
|
|
66 |
next
|
|
67 |
fix x y z assume "?ordf x y" "?ordf y z"
|
|
68 |
thus "?ordf x z" unfolding fun_ord_def
|
|
69 |
by (force dest: leq_trans)
|
|
70 |
next
|
|
71 |
fix x y assume "?ordf x y" "?ordf y x"
|
|
72 |
thus "x = y" unfolding fun_ord_def
|
|
73 |
by (force intro!: ext dest: leq_antisym)
|
|
74 |
next
|
|
75 |
fix A f assume f: "f \<in> A" and A: "chain ?ordf A"
|
|
76 |
thus "?ordf f (?lubf A)"
|
|
77 |
unfolding fun_lub_def fun_ord_def
|
|
78 |
by (blast intro: lub_upper chain_fun[OF A] f)
|
|
79 |
next
|
|
80 |
fix A :: "('b \<Rightarrow> 'a) set" and g :: "'b \<Rightarrow> 'a"
|
|
81 |
assume A: "chain ?ordf A" and g: "\<And>f. f \<in> A \<Longrightarrow> ?ordf f g"
|
|
82 |
show "?ordf (?lubf A) g" unfolding fun_lub_def fun_ord_def
|
|
83 |
by (blast intro: lub_least chain_fun[OF A] dest: g[unfolded fun_ord_def])
|
|
84 |
qed
|
|
85 |
qed
|
|
86 |
|
|
87 |
lemma ccpo: assumes "partial_function_definitions ord lb"
|
|
88 |
shows "class.ccpo ord (mk_less ord) lb"
|
|
89 |
using assms unfolding partial_function_definitions_def mk_less_def
|
|
90 |
by unfold_locales blast+
|
|
91 |
|
|
92 |
lemma partial_function_image:
|
|
93 |
assumes "partial_function_definitions ord Lub"
|
|
94 |
assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y"
|
|
95 |
assumes inv: "\<And>x. f (g x) = x"
|
|
96 |
shows "partial_function_definitions (img_ord f ord) (img_lub f g Lub)"
|
|
97 |
proof -
|
|
98 |
let ?iord = "img_ord f ord"
|
|
99 |
let ?ilub = "img_lub f g Lub"
|
|
100 |
|
|
101 |
interpret partial_function_definitions ord Lub by fact
|
|
102 |
show ?thesis
|
|
103 |
proof
|
|
104 |
fix A x assume "chain ?iord A" "x \<in> A"
|
|
105 |
then have "chain ord (f ` A)" "f x \<in> f ` A"
|
|
106 |
by (auto simp: img_ord_def intro: chainI dest: chainD)
|
|
107 |
thus "?iord x (?ilub A)"
|
|
108 |
unfolding inv img_lub_def img_ord_def by (rule lub_upper)
|
|
109 |
next
|
|
110 |
fix A x assume "chain ?iord A"
|
|
111 |
and 1: "\<And>z. z \<in> A \<Longrightarrow> ?iord z x"
|
|
112 |
then have "chain ord (f ` A)"
|
|
113 |
by (auto simp: img_ord_def intro: chainI dest: chainD)
|
|
114 |
thus "?iord (?ilub A) x"
|
|
115 |
unfolding inv img_lub_def img_ord_def
|
|
116 |
by (rule lub_least) (auto dest: 1[unfolded img_ord_def])
|
|
117 |
qed (auto simp: img_ord_def intro: leq_refl dest: leq_trans leq_antisym inj)
|
|
118 |
qed
|
|
119 |
|
|
120 |
context partial_function_definitions
|
|
121 |
begin
|
|
122 |
|
|
123 |
abbreviation "le_fun \<equiv> fun_ord leq"
|
|
124 |
abbreviation "lub_fun \<equiv> fun_lub lub"
|
|
125 |
abbreviation "fixp_fun == ccpo.fixp le_fun lub_fun"
|
|
126 |
abbreviation "mono_body \<equiv> monotone le_fun leq"
|
|
127 |
|
|
128 |
text {* Interpret manually, to avoid flooding everything with facts about
|
|
129 |
orders *}
|
|
130 |
|
|
131 |
lemma ccpo: "class.ccpo le_fun (mk_less le_fun) lub_fun"
|
|
132 |
apply (rule ccpo)
|
|
133 |
apply (rule partial_function_lift)
|
|
134 |
apply (rule partial_function_definitions_axioms)
|
|
135 |
done
|
|
136 |
|
|
137 |
text {* The crucial fixed-point theorem *}
|
|
138 |
|
|
139 |
lemma mono_body_fixp:
|
|
140 |
"(\<And>x. mono_body (\<lambda>f. F f x)) \<Longrightarrow> fixp_fun F = F (fixp_fun F)"
|
|
141 |
by (rule ccpo.fixp_unfold[OF ccpo]) (auto simp: monotone_def fun_ord_def)
|
|
142 |
|
|
143 |
text {* Version with curry/uncurry combinators, to be used by package *}
|
|
144 |
|
|
145 |
lemma fixp_rule_uc:
|
|
146 |
fixes F :: "'c \<Rightarrow> 'c" and
|
|
147 |
U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" and
|
|
148 |
C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c"
|
|
149 |
assumes mono: "\<And>x. mono_body (\<lambda>f. U (F (C f)) x)"
|
|
150 |
assumes eq: "f \<equiv> C (fixp_fun (\<lambda>f. U (F (C f))))"
|
|
151 |
assumes inverse: "\<And>f. C (U f) = f"
|
|
152 |
shows "f = F f"
|
|
153 |
proof -
|
|
154 |
have "f = C (fixp_fun (\<lambda>f. U (F (C f))))" by (simp add: eq)
|
|
155 |
also have "... = C (U (F (C (fixp_fun (\<lambda>f. U (F (C f)))))))"
|
|
156 |
by (subst mono_body_fixp[of "%f. U (F (C f))", OF mono]) (rule refl)
|
|
157 |
also have "... = F (C (fixp_fun (\<lambda>f. U (F (C f)))))" by (rule inverse)
|
|
158 |
also have "... = F f" by (simp add: eq)
|
|
159 |
finally show "f = F f" .
|
|
160 |
qed
|
|
161 |
|
|
162 |
text {* Rules for @{term mono_body}: *}
|
|
163 |
|
|
164 |
lemma const_mono[partial_function_mono]: "monotone ord leq (\<lambda>f. c)"
|
|
165 |
by (rule monotoneI) (rule leq_refl)
|
|
166 |
|
|
167 |
declaration {* Partial_Function.init @{term fixp_fun}
|
|
168 |
@{term mono_body} @{thm fixp_rule_uc} *}
|
|
169 |
|
|
170 |
end
|
|
171 |
|
|
172 |
|
|
173 |
subsection {* Flat interpretation: tailrec and option *}
|
|
174 |
|
|
175 |
definition
|
|
176 |
"flat_ord b x y \<longleftrightarrow> x = b \<or> x = y"
|
|
177 |
|
|
178 |
definition
|
|
179 |
"flat_lub b A = (if A \<subseteq> {b} then b else (THE x. x \<in> A - {b}))"
|
|
180 |
|
|
181 |
lemma flat_interpretation:
|
|
182 |
"partial_function_definitions (flat_ord b) (flat_lub b)"
|
|
183 |
proof
|
|
184 |
fix A x assume 1: "chain (flat_ord b) A" "x \<in> A"
|
|
185 |
show "flat_ord b x (flat_lub b A)"
|
|
186 |
proof cases
|
|
187 |
assume "x = b"
|
|
188 |
thus ?thesis by (simp add: flat_ord_def)
|
|
189 |
next
|
|
190 |
assume "x \<noteq> b"
|
|
191 |
with 1 have "A - {b} = {x}"
|
|
192 |
by (auto elim: chainE simp: flat_ord_def)
|
|
193 |
then have "flat_lub b A = x"
|
|
194 |
by (auto simp: flat_lub_def)
|
|
195 |
thus ?thesis by (auto simp: flat_ord_def)
|
|
196 |
qed
|
|
197 |
next
|
|
198 |
fix A z assume A: "chain (flat_ord b) A"
|
|
199 |
and z: "\<And>x. x \<in> A \<Longrightarrow> flat_ord b x z"
|
|
200 |
show "flat_ord b (flat_lub b A) z"
|
|
201 |
proof cases
|
|
202 |
assume "A \<subseteq> {b}"
|
|
203 |
thus ?thesis
|
|
204 |
by (auto simp: flat_lub_def flat_ord_def)
|
|
205 |
next
|
|
206 |
assume nb: "\<not> A \<subseteq> {b}"
|
|
207 |
then obtain y where y: "y \<in> A" "y \<noteq> b" by auto
|
|
208 |
with A have "A - {b} = {y}"
|
|
209 |
by (auto elim: chainE simp: flat_ord_def)
|
|
210 |
with nb have "flat_lub b A = y"
|
|
211 |
by (auto simp: flat_lub_def)
|
|
212 |
with z y show ?thesis by auto
|
|
213 |
qed
|
|
214 |
qed (auto simp: flat_ord_def)
|
|
215 |
|
|
216 |
interpretation tailrec!:
|
|
217 |
partial_function_definitions "flat_ord undefined" "flat_lub undefined"
|
|
218 |
by (rule flat_interpretation)
|
|
219 |
|
|
220 |
interpretation option!:
|
|
221 |
partial_function_definitions "flat_ord None" "flat_lub None"
|
|
222 |
by (rule flat_interpretation)
|
|
223 |
|
|
224 |
abbreviation "option_ord \<equiv> flat_ord None"
|
|
225 |
abbreviation "mono_option \<equiv> monotone (fun_ord option_ord) option_ord"
|
|
226 |
|
|
227 |
lemma bind_mono[partial_function_mono]:
|
|
228 |
assumes mf: "mono_option B" and mg: "\<And>y. mono_option (\<lambda>f. C y f)"
|
|
229 |
shows "mono_option (\<lambda>f. Option.bind (B f) (\<lambda>y. C y f))"
|
|
230 |
proof (rule monotoneI)
|
|
231 |
fix f g :: "'a \<Rightarrow> 'b option" assume fg: "fun_ord option_ord f g"
|
|
232 |
with mf
|
|
233 |
have "option_ord (B f) (B g)" by (rule monotoneD[of _ _ _ f g])
|
|
234 |
then have "option_ord (Option.bind (B f) (\<lambda>y. C y f)) (Option.bind (B g) (\<lambda>y. C y f))"
|
|
235 |
unfolding flat_ord_def by auto
|
|
236 |
also from mg
|
|
237 |
have "\<And>y'. option_ord (C y' f) (C y' g)"
|
|
238 |
by (rule monotoneD) (rule fg)
|
|
239 |
then have "option_ord (Option.bind (B g) (\<lambda>y'. C y' f)) (Option.bind (B g) (\<lambda>y'. C y' g))"
|
|
240 |
unfolding flat_ord_def by (cases "B g") auto
|
|
241 |
finally (option.leq_trans)
|
|
242 |
show "option_ord (Option.bind (B f) (\<lambda>y. C y f)) (Option.bind (B g) (\<lambda>y'. C y' g))" .
|
|
243 |
qed
|
|
244 |
|
|
245 |
|
|
246 |
end
|
|
247 |
|