author | huffman |
Fri, 27 May 2005 00:24:02 +0200 | |
changeset 16094 | a92ee2833938 |
parent 16067 | c57725e8055a |
child 16121 | a80aa66d2271 |
permissions | -rw-r--r-- |
2640 | 1 |
(* Title: HOLCF/Lift.thy |
2 |
ID: $Id$ |
|
12026 | 3 |
Author: Olaf Mueller |
2640 | 4 |
*) |
5 |
||
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
6 |
header {* Lifting types of class type to flat pcpo's *} |
12026 | 7 |
|
15577 | 8 |
theory Lift |
9 |
imports Cprod |
|
10 |
begin |
|
12026 | 11 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
12 |
defaultsort type |
12026 | 13 |
|
14 |
||
15 |
typedef 'a lift = "UNIV :: 'a option set" .. |
|
16 |
||
17 |
constdefs |
|
18 |
Undef :: "'a lift" |
|
19 |
"Undef == Abs_lift None" |
|
20 |
Def :: "'a => 'a lift" |
|
21 |
"Def x == Abs_lift (Some x)" |
|
22 |
||
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
23 |
instance lift :: (type) sq_ord .. |
12026 | 24 |
|
25 |
defs (overloaded) |
|
26 |
less_lift_def: "x << y == (x=y | x=Undef)" |
|
27 |
||
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
28 |
instance lift :: (type) po |
12026 | 29 |
proof |
30 |
fix x y z :: "'a lift" |
|
31 |
show "x << x" by (unfold less_lift_def) blast |
|
32 |
{ assume "x << y" and "y << x" thus "x = y" by (unfold less_lift_def) blast } |
|
33 |
{ assume "x << y" and "y << z" thus "x << z" by (unfold less_lift_def) blast } |
|
34 |
qed |
|
35 |
||
36 |
lemma inst_lift_po: "(op <<) = (\<lambda>x y. x = y | x = Undef)" |
|
37 |
-- {* For compatibility with old HOLCF-Version. *} |
|
38 |
by (simp only: less_lift_def [symmetric]) |
|
39 |
||
40 |
||
41 |
subsection {* Type lift is pointed *} |
|
42 |
||
43 |
lemma minimal_lift [iff]: "Undef << x" |
|
44 |
by (simp add: inst_lift_po) |
|
45 |
||
15930
145651bc64a8
Replaced all unnecessary uses of SOME with THE or LEAST
huffman
parents:
15651
diff
changeset
|
46 |
lemma UU_lift_def: "(THE u. \<forall>y. u \<sqsubseteq> y) = Undef" |
12026 | 47 |
apply (rule minimal2UU [symmetric]) |
48 |
apply (rule minimal_lift) |
|
49 |
done |
|
50 |
||
51 |
lemma least_lift: "EX x::'a lift. ALL y. x << y" |
|
52 |
apply (rule_tac x = Undef in exI) |
|
53 |
apply (rule minimal_lift [THEN allI]) |
|
54 |
done |
|
55 |
||
56 |
||
57 |
subsection {* Type lift is a cpo *} |
|
58 |
||
59 |
text {* |
|
60 |
The following lemmas have already been proved in @{text Pcpo.ML} and |
|
61 |
@{text Fix.ML}, but there class @{text pcpo} is assumed, although |
|
62 |
only @{text po} is necessary and a least element. Therefore they are |
|
63 |
redone here for the @{text po} lift with least element @{text |
|
64 |
Undef}. |
|
65 |
*} |
|
66 |
||
67 |
lemma flat_imp_chfin_poo: "(ALL Y. chain(Y::nat=>('a)lift)-->(EX n. max_in_chain n Y))" |
|
68 |
-- {* Tailoring @{text flat_imp_chfin} of @{text Fix.ML} to @{text lift} *} |
|
69 |
apply (unfold max_in_chain_def) |
|
16067 | 70 |
apply clarify |
71 |
apply (case_tac "ALL i. Y i = Undef") |
|
72 |
apply simp |
|
73 |
apply simp |
|
74 |
apply (erule exE) |
|
75 |
apply (rule_tac x=i in exI) |
|
76 |
apply clarify |
|
77 |
apply (drule chain_mono3, assumption) |
|
78 |
apply (simp add: less_lift_def) |
|
12026 | 79 |
done |
80 |
||
16067 | 81 |
instance lift :: (type) chfin |
82 |
apply intro_classes |
|
83 |
apply (rule flat_imp_chfin_poo) |
|
12026 | 84 |
done |
85 |
||
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
86 |
instance lift :: (type) pcpo |
12026 | 87 |
apply intro_classes |
88 |
apply (rule least_lift) |
|
89 |
done |
|
90 |
||
91 |
lemma inst_lift_pcpo: "UU = Undef" |
|
92 |
-- {* For compatibility with old HOLCF-Version. *} |
|
93 |
by (simp add: UU_def UU_lift_def) |
|
94 |
||
95 |
||
96 |
subsection {* Lift as a datatype *} |
|
97 |
||
98 |
lemma lift_distinct1: "UU ~= Def x" |
|
99 |
by (simp add: Undef_def Def_def Abs_lift_inject lift_def inst_lift_pcpo) |
|
100 |
||
101 |
lemma lift_distinct2: "Def x ~= UU" |
|
102 |
by (simp add: Undef_def Def_def Abs_lift_inject lift_def inst_lift_pcpo) |
|
103 |
||
104 |
lemma Def_inject: "(Def x = Def x') = (x = x')" |
|
105 |
by (simp add: Def_def Abs_lift_inject lift_def) |
|
106 |
||
107 |
lemma lift_induct: "P UU ==> (!!x. P (Def x)) ==> P y" |
|
108 |
apply (induct y) |
|
109 |
apply (induct_tac y) |
|
110 |
apply (simp_all add: Undef_def Def_def inst_lift_pcpo) |
|
111 |
done |
|
112 |
||
113 |
rep_datatype lift |
|
114 |
distinct lift_distinct1 lift_distinct2 |
|
115 |
inject Def_inject |
|
116 |
induction lift_induct |
|
117 |
||
118 |
lemma Def_not_UU: "Def a ~= UU" |
|
119 |
by simp |
|
120 |
||
121 |
||
122 |
subsection {* Further operations *} |
|
123 |
||
124 |
consts |
|
125 |
flift1 :: "('a => 'b::pcpo) => ('a lift -> 'b)" |
|
126 |
flift2 :: "('a => 'b) => ('a lift -> 'b lift)" |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
127 |
liftpair ::"'a::type lift * 'b::type lift => ('a * 'b) lift" |
2640 | 128 |
|
12026 | 129 |
defs |
130 |
flift1_def: |
|
131 |
"flift1 f == (LAM x. (case x of |
|
132 |
UU => UU |
|
133 |
| Def y => (f y)))" |
|
134 |
flift2_def: |
|
135 |
"flift2 f == (LAM x. (case x of |
|
136 |
UU => UU |
|
137 |
| Def y => Def (f y)))" |
|
138 |
liftpair_def: |
|
139 |
"liftpair x == (case (cfst$x) of |
|
140 |
UU => UU |
|
141 |
| Def x1 => (case (csnd$x) of |
|
142 |
UU => UU |
|
143 |
| Def x2 => Def (x1,x2)))" |
|
144 |
||
145 |
||
146 |
declare inst_lift_pcpo [symmetric, simp] |
|
147 |
||
148 |
||
149 |
lemma less_lift: "(x::'a lift) << y = (x=y | x=UU)" |
|
150 |
by (simp add: inst_lift_po) |
|
151 |
||
152 |
||
153 |
text {* @{text UU} and @{text Def} *} |
|
154 |
||
155 |
lemma Lift_exhaust: "x = UU | (EX y. x = Def y)" |
|
156 |
by (induct x) simp_all |
|
157 |
||
158 |
lemma Lift_cases: "[| x = UU ==> P; ? a. x = Def a ==> P |] ==> P" |
|
159 |
by (insert Lift_exhaust) blast |
|
160 |
||
161 |
lemma not_Undef_is_Def: "(x ~= UU) = (EX y. x = Def y)" |
|
162 |
by (cases x) simp_all |
|
163 |
||
164 |
text {* |
|
165 |
For @{term "x ~= UU"} in assumptions @{text def_tac} replaces @{text |
|
166 |
x} by @{text "Def a"} in conclusion. *} |
|
167 |
||
15930
145651bc64a8
Replaced all unnecessary uses of SOME with THE or LEAST
huffman
parents:
15651
diff
changeset
|
168 |
ML_setup {* |
12026 | 169 |
local val not_Undef_is_Def = thm "not_Undef_is_Def" |
170 |
in val def_tac = SIMPSET' (fn ss => |
|
171 |
etac (not_Undef_is_Def RS iffD1 RS exE) THEN' asm_simp_tac ss) |
|
172 |
end; |
|
173 |
*} |
|
174 |
||
175 |
lemma Undef_eq_UU: "Undef = UU" |
|
176 |
by (rule inst_lift_pcpo [symmetric]) |
|
177 |
||
178 |
lemma DefE: "Def x = UU ==> R" |
|
179 |
by simp |
|
180 |
||
181 |
lemma DefE2: "[| x = Def s; x = UU |] ==> R" |
|
182 |
by simp |
|
183 |
||
184 |
lemma Def_inject_less_eq: "Def x << Def y = (x = y)" |
|
185 |
by (simp add: less_lift_def) |
|
186 |
||
187 |
lemma Def_less_is_eq [simp]: "Def x << y = (Def x = y)" |
|
188 |
by (simp add: less_lift) |
|
189 |
||
190 |
||
191 |
subsection {* Lift is flat *} |
|
192 |
||
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
12026
diff
changeset
|
193 |
instance lift :: (type) flat |
12026 | 194 |
proof |
195 |
show "ALL x y::'a lift. x << y --> x = UU | x = y" |
|
196 |
by (simp add: less_lift) |
|
197 |
qed |
|
198 |
||
199 |
defaultsort pcpo |
|
200 |
||
201 |
||
202 |
text {* |
|
203 |
\medskip Two specific lemmas for the combination of LCF and HOL |
|
204 |
terms. |
|
205 |
*} |
|
206 |
||
207 |
lemma cont_Rep_CFun_app: "[|cont g; cont f|] ==> cont(%x. ((f x)$(g x)) s)" |
|
208 |
apply (rule cont2cont_CF1L) |
|
209 |
apply (tactic "resolve_tac cont_lemmas1 1")+ |
|
210 |
apply auto |
|
211 |
done |
|
212 |
||
213 |
lemma cont_Rep_CFun_app_app: "[|cont g; cont f|] ==> cont(%x. ((f x)$(g x)) s t)" |
|
214 |
apply (rule cont2cont_CF1L) |
|
215 |
apply (erule cont_Rep_CFun_app) |
|
216 |
apply assumption |
|
217 |
done |
|
2640 | 218 |
|
12026 | 219 |
text {* Continuity of if-then-else. *} |
220 |
||
221 |
lemma cont_if: "[| cont f1; cont f2 |] ==> cont (%x. if b then f1 x else f2 x)" |
|
222 |
by (cases b) simp_all |
|
223 |
||
224 |
||
225 |
subsection {* Continuity Proofs for flift1, flift2, if *} |
|
226 |
||
227 |
text {* Need the instance of @{text flat}. *} |
|
228 |
||
229 |
lemma cont_flift1_arg: "cont (lift_case UU f)" |
|
230 |
-- {* @{text flift1} is continuous in its argument itself. *} |
|
231 |
apply (rule flatdom_strict2cont) |
|
232 |
apply simp |
|
233 |
done |
|
234 |
||
235 |
lemma cont_flift1_not_arg: "!!f. [| !! a. cont (%y. (f y) a) |] ==> |
|
236 |
cont (%y. lift_case UU (f y))" |
|
237 |
-- {* @{text flift1} is continuous in a variable that occurs only |
|
238 |
in the @{text Def} branch. *} |
|
239 |
apply (rule cont2cont_CF1L_rev) |
|
240 |
apply (intro strip) |
|
241 |
apply (case_tac y) |
|
242 |
apply simp |
|
243 |
apply simp |
|
244 |
done |
|
245 |
||
246 |
lemma cont_flift1_arg_and_not_arg: "!!f. [| !! a. cont (%y. (f y) a); cont g|] ==> |
|
247 |
cont (%y. lift_case UU (f y) (g y))" |
|
248 |
-- {* @{text flift1} is continuous in a variable that occurs either |
|
249 |
in the @{text Def} branch or in the argument. *} |
|
250 |
apply (rule_tac tt = g in cont2cont_app) |
|
251 |
apply (rule cont_flift1_not_arg) |
|
252 |
apply auto |
|
253 |
apply (rule cont_flift1_arg) |
|
254 |
done |
|
255 |
||
256 |
lemma cont_flift2_arg: "cont (lift_case UU (%y. Def (f y)))" |
|
257 |
-- {* @{text flift2} is continuous in its argument itself. *} |
|
258 |
apply (rule flatdom_strict2cont) |
|
259 |
apply simp |
|
260 |
done |
|
261 |
||
262 |
text {* |
|
14096 | 263 |
\medskip Extension of @{text cont_tac} and installation of simplifier. |
12026 | 264 |
*} |
265 |
||
266 |
lemmas cont_lemmas_ext [simp] = |
|
267 |
cont_flift1_arg cont_flift2_arg |
|
268 |
cont_flift1_arg_and_not_arg cont2cont_CF1L_rev2 |
|
269 |
cont_Rep_CFun_app cont_Rep_CFun_app_app cont_if |
|
270 |
||
271 |
ML_setup {* |
|
272 |
val cont_lemmas2 = cont_lemmas1 @ thms "cont_lemmas_ext"; |
|
273 |
||
274 |
fun cont_tac i = resolve_tac cont_lemmas2 i; |
|
275 |
fun cont_tacR i = REPEAT (cont_tac i); |
|
276 |
||
277 |
local val flift1_def = thm "flift1_def" and flift2_def = thm "flift2_def" |
|
278 |
in fun cont_tacRs i = |
|
279 |
simp_tac (simpset() addsimps [flift1_def, flift2_def]) i THEN |
|
280 |
REPEAT (cont_tac i) |
|
281 |
end; |
|
15651
4b393520846e
Replaced continuity solver with new continuity simproc. Also removed cont lemmas from simp set, so that the simproc actually gets used.
huffman
parents:
15577
diff
changeset
|
282 |
*} |
12026 | 283 |
|
284 |
||
285 |
subsection {* flift1, flift2 *} |
|
286 |
||
287 |
lemma flift1_Def [simp]: "flift1 f$(Def x) = (f x)" |
|
288 |
by (simp add: flift1_def) |
|
289 |
||
290 |
lemma flift2_Def [simp]: "flift2 f$(Def x) = Def (f x)" |
|
291 |
by (simp add: flift2_def) |
|
292 |
||
293 |
lemma flift1_UU [simp]: "flift1 f$UU = UU" |
|
294 |
by (simp add: flift1_def) |
|
295 |
||
296 |
lemma flift2_UU [simp]: "flift2 f$UU = UU" |
|
297 |
by (simp add: flift2_def) |
|
298 |
||
299 |
lemma flift2_nUU [simp]: "x~=UU ==> (flift2 f)$x~=UU" |
|
300 |
by (tactic "def_tac 1") |
|
2640 | 301 |
|
302 |
end |