author | wenzelm |
Wed, 12 Dec 2001 20:37:31 +0100 | |
changeset 12484 | 7ad150f5fc10 |
parent 12030 | 46d57d0290a2 |
child 14981 | e73f8140af78 |
permissions | -rw-r--r-- |
9245 | 1 |
(* Title: HOLCF/Cfun3 |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
2 |
ID: $Id$ |
1461 | 3 |
Author: Franz Regensburger |
12030 | 4 |
License: GPL (GNU GENERAL PUBLIC LICENSE) |
9245 | 5 |
|
6 |
Class instance of -> for class pcpo |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
7 |
*) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
8 |
|
2640 | 9 |
(* for compatibility with old HOLCF-Version *) |
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
10 |
Goal "UU = Abs_CFun(%x. UU)"; |
9245 | 11 |
by (simp_tac (HOL_ss addsimps [UU_def,UU_cfun_def]) 1); |
12 |
qed "inst_cfun_pcpo"; |
|
2640 | 13 |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
14 |
(* ------------------------------------------------------------------------ *) |
5291 | 15 |
(* the contlub property for Rep_CFun its 'first' argument *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
16 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
17 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
18 |
Goal "contlub(Rep_CFun)"; |
9245 | 19 |
by (rtac contlubI 1); |
20 |
by (strip_tac 1); |
|
21 |
by (rtac (expand_fun_eq RS iffD2) 1); |
|
22 |
by (strip_tac 1); |
|
23 |
by (stac thelub_cfun 1); |
|
24 |
by (atac 1); |
|
25 |
by (stac Cfunapp2 1); |
|
26 |
by (etac cont_lubcfun 1); |
|
27 |
by (stac thelub_fun 1); |
|
28 |
by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); |
|
29 |
by (rtac refl 1); |
|
30 |
qed "contlub_Rep_CFun1"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
31 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
32 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
33 |
(* ------------------------------------------------------------------------ *) |
5291 | 34 |
(* the cont property for Rep_CFun in its first argument *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
35 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
36 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
37 |
Goal "cont(Rep_CFun)"; |
9245 | 38 |
by (rtac monocontlub2cont 1); |
39 |
by (rtac monofun_Rep_CFun1 1); |
|
40 |
by (rtac contlub_Rep_CFun1 1); |
|
41 |
qed "cont_Rep_CFun1"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
42 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
43 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
44 |
(* ------------------------------------------------------------------------ *) |
5291 | 45 |
(* contlub, cont properties of Rep_CFun in its first argument in mixfix _[_] *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
46 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
47 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
48 |
Goal |
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
4423
diff
changeset
|
49 |
"chain(FY) ==>\ |
10834 | 50 |
\ lub(range FY)$x = lub(range (%i. FY(i)$x))"; |
9245 | 51 |
by (rtac trans 1); |
52 |
by (etac (contlub_Rep_CFun1 RS contlubE RS spec RS mp RS fun_cong) 1); |
|
53 |
by (stac thelub_fun 1); |
|
54 |
by (etac (monofun_Rep_CFun1 RS ch2ch_monofun) 1); |
|
55 |
by (rtac refl 1); |
|
56 |
qed "contlub_cfun_fun"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
57 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
58 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
59 |
Goal |
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
4423
diff
changeset
|
60 |
"chain(FY) ==>\ |
10834 | 61 |
\ range(%i. FY(i)$x) <<| lub(range FY)$x"; |
9245 | 62 |
by (rtac thelubE 1); |
63 |
by (etac ch2ch_Rep_CFunL 1); |
|
64 |
by (etac (contlub_cfun_fun RS sym) 1); |
|
65 |
qed "cont_cfun_fun"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
66 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
67 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
68 |
(* ------------------------------------------------------------------------ *) |
5291 | 69 |
(* contlub, cont properties of Rep_CFun in both argument in mixfix _[_] *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
70 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
71 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
72 |
Goal |
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
4423
diff
changeset
|
73 |
"[|chain(FY);chain(TY)|] ==>\ |
10834 | 74 |
\ (lub(range FY))$(lub(range TY)) = lub(range(%i. FY(i)$(TY i)))"; |
9245 | 75 |
by (rtac contlub_CF2 1); |
76 |
by (rtac cont_Rep_CFun1 1); |
|
77 |
by (rtac allI 1); |
|
78 |
by (rtac cont_Rep_CFun2 1); |
|
79 |
by (atac 1); |
|
80 |
by (atac 1); |
|
81 |
qed "contlub_cfun"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
82 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
83 |
Goal |
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
4423
diff
changeset
|
84 |
"[|chain(FY);chain(TY)|] ==>\ |
10834 | 85 |
\ range(%i.(FY i)$(TY i)) <<| (lub (range FY))$(lub(range TY))"; |
9245 | 86 |
by (rtac thelubE 1); |
87 |
by (rtac (monofun_Rep_CFun1 RS ch2ch_MF2LR) 1); |
|
88 |
by (rtac allI 1); |
|
89 |
by (rtac monofun_Rep_CFun2 1); |
|
90 |
by (atac 1); |
|
91 |
by (atac 1); |
|
92 |
by (etac (contlub_cfun RS sym) 1); |
|
93 |
by (atac 1); |
|
94 |
qed "cont_cfun"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
95 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
96 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
97 |
(* ------------------------------------------------------------------------ *) |
5291 | 98 |
(* cont2cont lemma for Rep_CFun *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
99 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
100 |
|
10834 | 101 |
Goal "[|cont(%x. ft x);cont(%x. tt x)|] ==> cont(%x. (ft x)$(tt x))"; |
9245 | 102 |
by (best_tac (claset() addIs [cont2cont_app2, cont_const, cont_Rep_CFun1, |
103 |
cont_Rep_CFun2]) 1); |
|
104 |
qed "cont2cont_Rep_CFun"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
105 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
106 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
107 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
108 |
(* ------------------------------------------------------------------------ *) |
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
961
diff
changeset
|
109 |
(* cont2mono Lemma for %x. LAM y. c1(x)(y) *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
110 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
111 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
112 |
val [p1,p2] = Goal |
9245 | 113 |
"[| !!x. cont(c1 x); !!y. monofun(%x. c1 x y)|] ==> monofun(%x. LAM y. c1 x y)"; |
114 |
by (rtac monofunI 1); |
|
115 |
by (strip_tac 1); |
|
116 |
by (stac less_cfun 1); |
|
117 |
by (stac less_fun 1); |
|
118 |
by (rtac allI 1); |
|
119 |
by (stac beta_cfun 1); |
|
120 |
by (rtac p1 1); |
|
121 |
by (stac beta_cfun 1); |
|
122 |
by (rtac p1 1); |
|
123 |
by (etac (p2 RS monofunE RS spec RS spec RS mp) 1); |
|
124 |
qed "cont2mono_LAM"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
125 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
126 |
(* ------------------------------------------------------------------------ *) |
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
961
diff
changeset
|
127 |
(* cont2cont Lemma for %x. LAM y. c1 x y) *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
128 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
129 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
130 |
val [p1,p2] = Goal |
9245 | 131 |
"[| !!x. cont(c1 x); !!y. cont(%x. c1 x y) |] ==> cont(%x. LAM y. c1 x y)"; |
132 |
by (rtac monocontlub2cont 1); |
|
133 |
by (rtac (p1 RS cont2mono_LAM) 1); |
|
134 |
by (rtac (p2 RS cont2mono) 1); |
|
135 |
by (rtac contlubI 1); |
|
136 |
by (strip_tac 1); |
|
137 |
by (stac thelub_cfun 1); |
|
138 |
by (rtac (p1 RS cont2mono_LAM RS ch2ch_monofun) 1); |
|
139 |
by (rtac (p2 RS cont2mono) 1); |
|
140 |
by (atac 1); |
|
141 |
by (res_inst_tac [("f","Abs_CFun")] arg_cong 1); |
|
142 |
by (rtac ext 1); |
|
143 |
by (stac (p1 RS beta_cfun RS ext) 1); |
|
144 |
by (etac (p2 RS cont2contlub RS contlubE RS spec RS mp) 1); |
|
145 |
qed "cont2cont_LAM"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
146 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
147 |
(* ------------------------------------------------------------------------ *) |
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
961
diff
changeset
|
148 |
(* cont2cont tactic *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
149 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
150 |
|
5291 | 151 |
val cont_lemmas1 = [cont_const, cont_id, cont_Rep_CFun2, |
152 |
cont2cont_Rep_CFun,cont2cont_LAM]; |
|
2566 | 153 |
|
154 |
Addsimps cont_lemmas1; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
155 |
|
4004 | 156 |
(* HINT: cont_tac is now installed in simplifier in Lift.ML ! *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
157 |
|
2566 | 158 |
(*val cont_tac = (fn i => (resolve_tac cont_lemmas i));*) |
159 |
(*val cont_tacR = (fn i => (REPEAT (cont_tac i)));*) |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
160 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
161 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
162 |
(* function application _[_] is strict in its first arguments *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
163 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
164 |
|
10834 | 165 |
Goal "(UU::'a::cpo->'b)$x = (UU::'b)"; |
9245 | 166 |
by (stac inst_cfun_pcpo 1); |
167 |
by (stac beta_cfun 1); |
|
168 |
by (Simp_tac 1); |
|
169 |
by (rtac refl 1); |
|
170 |
qed "strict_Rep_CFun1"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
171 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
172 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
173 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
174 |
(* results about strictify *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
175 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
176 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
177 |
Goalw [Istrictify_def] |
9245 | 178 |
"Istrictify(f)(UU)= (UU)"; |
179 |
by (Simp_tac 1); |
|
180 |
qed "Istrictify1"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
181 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
182 |
Goalw [Istrictify_def] |
10834 | 183 |
"~x=UU ==> Istrictify(f)(x)=f$x"; |
9245 | 184 |
by (Asm_simp_tac 1); |
185 |
qed "Istrictify2"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
186 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
187 |
Goal "monofun(Istrictify)"; |
9245 | 188 |
by (rtac monofunI 1); |
189 |
by (strip_tac 1); |
|
190 |
by (rtac (less_fun RS iffD2) 1); |
|
191 |
by (strip_tac 1); |
|
192 |
by (res_inst_tac [("Q","xa=UU")] (excluded_middle RS disjE) 1); |
|
193 |
by (stac Istrictify2 1); |
|
194 |
by (atac 1); |
|
195 |
by (stac Istrictify2 1); |
|
196 |
by (atac 1); |
|
197 |
by (rtac monofun_cfun_fun 1); |
|
198 |
by (atac 1); |
|
199 |
by (hyp_subst_tac 1); |
|
200 |
by (stac Istrictify1 1); |
|
201 |
by (stac Istrictify1 1); |
|
202 |
by (rtac refl_less 1); |
|
203 |
qed "monofun_Istrictify1"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
204 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
205 |
Goal "monofun(Istrictify(f))"; |
9245 | 206 |
by (rtac monofunI 1); |
207 |
by (strip_tac 1); |
|
208 |
by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); |
|
209 |
by (stac Istrictify2 1); |
|
210 |
by (etac notUU_I 1); |
|
211 |
by (atac 1); |
|
212 |
by (stac Istrictify2 1); |
|
213 |
by (atac 1); |
|
214 |
by (rtac monofun_cfun_arg 1); |
|
215 |
by (atac 1); |
|
216 |
by (hyp_subst_tac 1); |
|
217 |
by (stac Istrictify1 1); |
|
218 |
by (rtac minimal 1); |
|
219 |
qed "monofun_Istrictify2"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
220 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
221 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
222 |
Goal "contlub(Istrictify)"; |
9245 | 223 |
by (rtac contlubI 1); |
224 |
by (strip_tac 1); |
|
225 |
by (rtac (expand_fun_eq RS iffD2) 1); |
|
226 |
by (strip_tac 1); |
|
227 |
by (stac thelub_fun 1); |
|
228 |
by (etac (monofun_Istrictify1 RS ch2ch_monofun) 1); |
|
229 |
by (res_inst_tac [("Q","x=UU")] (excluded_middle RS disjE) 1); |
|
230 |
by (stac Istrictify2 1); |
|
231 |
by (atac 1); |
|
232 |
by (stac (Istrictify2 RS ext) 1); |
|
233 |
by (atac 1); |
|
234 |
by (stac thelub_cfun 1); |
|
235 |
by (atac 1); |
|
236 |
by (stac beta_cfun 1); |
|
237 |
by (rtac cont_lubcfun 1); |
|
238 |
by (atac 1); |
|
239 |
by (rtac refl 1); |
|
240 |
by (hyp_subst_tac 1); |
|
241 |
by (stac Istrictify1 1); |
|
242 |
by (stac (Istrictify1 RS ext) 1); |
|
243 |
by (rtac (chain_UU_I_inverse RS sym) 1); |
|
244 |
by (rtac (refl RS allI) 1); |
|
245 |
qed "contlub_Istrictify1"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
246 |
|
9245 | 247 |
Goal "contlub(Istrictify(f::'a -> 'b))"; |
248 |
by (rtac contlubI 1); |
|
249 |
by (strip_tac 1); |
|
250 |
by (case_tac "lub(range(Y))=(UU::'a)" 1); |
|
251 |
by (asm_simp_tac (simpset() addsimps [Istrictify1, chain_UU_I_inverse, chain_UU_I, Istrictify1]) 1); |
|
252 |
by (stac Istrictify2 1); |
|
253 |
by (atac 1); |
|
10834 | 254 |
by (res_inst_tac [("s","lub(range(%i. f$(Y i)))")] trans 1); |
9245 | 255 |
by (rtac contlub_cfun_arg 1); |
256 |
by (atac 1); |
|
257 |
by (rtac lub_equal2 1); |
|
258 |
by (best_tac (claset() addIs [ch2ch_monofun, monofun_Istrictify2]) 3); |
|
259 |
by (best_tac (claset() addIs [ch2ch_monofun, monofun_Rep_CFun2]) 2); |
|
260 |
by (rtac (chain_mono2 RS exE) 1); |
|
261 |
by (atac 2); |
|
262 |
by (etac chain_UU_I_inverse2 1); |
|
263 |
by (blast_tac (claset() addIs [Istrictify2 RS sym]) 1); |
|
264 |
qed "contlub_Istrictify2"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
265 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
266 |
|
1779 | 267 |
bind_thm ("cont_Istrictify1", contlub_Istrictify1 RS |
1461 | 268 |
(monofun_Istrictify1 RS monocontlub2cont)); |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
269 |
|
1779 | 270 |
bind_thm ("cont_Istrictify2", contlub_Istrictify2 RS |
1461 | 271 |
(monofun_Istrictify2 RS monocontlub2cont)); |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
272 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
273 |
|
10834 | 274 |
Goalw [strictify_def] "strictify$f$UU=UU"; |
9245 | 275 |
by (stac beta_cfun 1); |
276 |
by (simp_tac (simpset() addsimps [cont_Istrictify2,cont_Istrictify1, cont2cont_CF1L]) 1); |
|
277 |
by (stac beta_cfun 1); |
|
278 |
by (rtac cont_Istrictify2 1); |
|
279 |
by (rtac Istrictify1 1); |
|
280 |
qed "strictify1"; |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
281 |
|
10834 | 282 |
Goalw [strictify_def] "~x=UU ==> strictify$f$x=f$x"; |
9245 | 283 |
by (stac beta_cfun 1); |
284 |
by (simp_tac (simpset() addsimps [cont_Istrictify2,cont_Istrictify1, cont2cont_CF1L]) 1); |
|
285 |
by (stac beta_cfun 1); |
|
286 |
by (rtac cont_Istrictify2 1); |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
287 |
by (etac Istrictify2 1); |
9245 | 288 |
qed "strictify2"; |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
289 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
290 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
291 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
292 |
(* Instantiate the simplifier *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
293 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
294 |
|
5291 | 295 |
Addsimps [minimal,refl_less,beta_cfun,strict_Rep_CFun1,strictify1, strictify2]; |
1267 | 296 |
|
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
297 |
|
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
298 |
(* ------------------------------------------------------------------------ *) |
1168
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
regensbu
parents:
961
diff
changeset
|
299 |
(* use cont_tac as autotac. *) |
243
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
300 |
(* ------------------------------------------------------------------------ *) |
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
nipkow
parents:
diff
changeset
|
301 |
|
4004 | 302 |
(* HINT: cont_tac is now installed in simplifier in Lift.ML ! *) |
4098 | 303 |
(*simpset_ref() := simpset() addsolver (K (DEPTH_SOLVE_1 o cont_tac));*) |
3326 | 304 |
|
305 |
(* ------------------------------------------------------------------------ *) |
|
4721
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
oheimb
parents:
4423
diff
changeset
|
306 |
(* some lemmata for functions with flat/chfin domain/range types *) |
3326 | 307 |
(* ------------------------------------------------------------------------ *) |
308 |
||
9245 | 309 |
Goal "chain (Y::nat => 'a::cpo->'b::chfin) \ |
10834 | 310 |
\ ==> !s. ? n. lub(range(Y))$s = Y n$s"; |
9245 | 311 |
by (rtac allI 1); |
312 |
by (stac contlub_cfun_fun 1); |
|
313 |
by (atac 1); |
|
314 |
by (fast_tac (HOL_cs addSIs [thelubI,chfin,lub_finch2,chfin2finch,ch2ch_Rep_CFunL])1); |
|
315 |
qed "chfin_Rep_CFunR"; |
|
3326 | 316 |
|
317 |
(* ------------------------------------------------------------------------ *) |
|
318 |
(* continuous isomorphisms are strict *) |
|
319 |
(* a prove for embedding projection pairs is similar *) |
|
320 |
(* ------------------------------------------------------------------------ *) |
|
321 |
||
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
322 |
Goal |
10834 | 323 |
"!!f g.[|!y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a) |] \ |
324 |
\ ==> f$UU=UU & g$UU=UU"; |
|
9245 | 325 |
by (rtac conjI 1); |
326 |
by (rtac UU_I 1); |
|
10834 | 327 |
by (res_inst_tac [("s","f$(g$(UU::'b))"),("t","UU::'b")] subst 1); |
9245 | 328 |
by (etac spec 1); |
329 |
by (rtac (minimal RS monofun_cfun_arg) 1); |
|
330 |
by (rtac UU_I 1); |
|
10834 | 331 |
by (res_inst_tac [("s","g$(f$(UU::'a))"),("t","UU::'a")] subst 1); |
9245 | 332 |
by (etac spec 1); |
333 |
by (rtac (minimal RS monofun_cfun_arg) 1); |
|
334 |
qed "iso_strict"; |
|
3326 | 335 |
|
336 |
||
10834 | 337 |
Goal "[|!x. rep$(ab$x)=x;!y. ab$(rep$y)=y; z~=UU|] ==> rep$z ~= UU"; |
10230 | 338 |
by (etac contrapos_nn 1); |
9245 | 339 |
by (dres_inst_tac [("f","ab")] cfun_arg_cong 1); |
340 |
by (etac box_equals 1); |
|
341 |
by (fast_tac HOL_cs 1); |
|
342 |
by (etac (iso_strict RS conjunct1) 1); |
|
343 |
by (atac 1); |
|
344 |
qed "isorep_defined"; |
|
3326 | 345 |
|
10834 | 346 |
Goal "[|!x. rep$(ab$x) = x;!y. ab$(rep$y)=y ; z~=UU|] ==> ab$z ~= UU"; |
10230 | 347 |
by (etac contrapos_nn 1); |
9245 | 348 |
by (dres_inst_tac [("f","rep")] cfun_arg_cong 1); |
349 |
by (etac box_equals 1); |
|
350 |
by (fast_tac HOL_cs 1); |
|
351 |
by (etac (iso_strict RS conjunct2) 1); |
|
352 |
by (atac 1); |
|
353 |
qed "isoabs_defined"; |
|
3326 | 354 |
|
355 |
(* ------------------------------------------------------------------------ *) |
|
356 |
(* propagation of flatness and chainfiniteness by continuous isomorphisms *) |
|
357 |
(* ------------------------------------------------------------------------ *) |
|
358 |
||
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
359 |
Goal "!!f g.[|! Y::nat=>'a. chain Y --> (? n. max_in_chain n Y); \ |
10834 | 360 |
\ !y. f$(g$y)=(y::'b) ; !x. g$(f$x)=(x::'a::chfin) |] \ |
9245 | 361 |
\ ==> ! Y::nat=>'b. chain Y --> (? n. max_in_chain n Y)"; |
362 |
by (rewtac max_in_chain_def); |
|
363 |
by (strip_tac 1); |
|
364 |
by (rtac exE 1); |
|
10834 | 365 |
by (res_inst_tac [("P","chain(%i. g$(Y i))")] mp 1); |
9245 | 366 |
by (etac spec 1); |
367 |
by (etac ch2ch_Rep_CFunR 1); |
|
368 |
by (rtac exI 1); |
|
369 |
by (strip_tac 1); |
|
10834 | 370 |
by (res_inst_tac [("s","f$(g$(Y x))"),("t","Y(x)")] subst 1); |
9245 | 371 |
by (etac spec 1); |
10834 | 372 |
by (res_inst_tac [("s","f$(g$(Y j))"),("t","Y(j)")] subst 1); |
9245 | 373 |
by (etac spec 1); |
374 |
by (rtac cfun_arg_cong 1); |
|
375 |
by (rtac mp 1); |
|
376 |
by (etac spec 1); |
|
377 |
by (atac 1); |
|
378 |
qed "chfin2chfin"; |
|
3326 | 379 |
|
380 |
||
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
381 |
Goal "!!f g.[|!x y::'a. x<<y --> x=UU | x=y; \ |
10834 | 382 |
\ !y. f$(g$y)=(y::'b); !x. g$(f$x)=(x::'a)|] ==> !x y::'b. x<<y --> x=UU | x=y"; |
9245 | 383 |
by (strip_tac 1); |
384 |
by (rtac disjE 1); |
|
10834 | 385 |
by (res_inst_tac [("P","g$x<<g$y")] mp 1); |
9245 | 386 |
by (etac monofun_cfun_arg 2); |
387 |
by (dtac spec 1); |
|
388 |
by (etac spec 1); |
|
389 |
by (rtac disjI1 1); |
|
390 |
by (rtac trans 1); |
|
10834 | 391 |
by (res_inst_tac [("s","f$(g$x)"),("t","x")] subst 1); |
9245 | 392 |
by (etac spec 1); |
393 |
by (etac cfun_arg_cong 1); |
|
394 |
by (rtac (iso_strict RS conjunct1) 1); |
|
395 |
by (atac 1); |
|
396 |
by (atac 1); |
|
397 |
by (rtac disjI2 1); |
|
10834 | 398 |
by (res_inst_tac [("s","f$(g$x)"),("t","x")] subst 1); |
9245 | 399 |
by (etac spec 1); |
10834 | 400 |
by (res_inst_tac [("s","f$(g$y)"),("t","y")] subst 1); |
9245 | 401 |
by (etac spec 1); |
402 |
by (etac cfun_arg_cong 1); |
|
403 |
qed "flat2flat"; |
|
3326 | 404 |
|
405 |
(* ------------------------------------------------------------------------- *) |
|
406 |
(* a result about functions with flat codomain *) |
|
407 |
(* ------------------------------------------------------------------------- *) |
|
408 |
||
10834 | 409 |
Goal "f$(x::'a)=(c::'b::flat) ==> f$(UU::'a)=(UU::'b) | (!z. f$(z::'a)=c)"; |
410 |
by (case_tac "f$(x::'a)=(UU::'b)" 1); |
|
9245 | 411 |
by (rtac disjI1 1); |
412 |
by (rtac UU_I 1); |
|
10834 | 413 |
by (res_inst_tac [("s","f$(x)"),("t","UU::'b")] subst 1); |
9245 | 414 |
by (atac 1); |
415 |
by (rtac (minimal RS monofun_cfun_arg) 1); |
|
10834 | 416 |
by (case_tac "f$(UU::'a)=(UU::'b)" 1); |
9245 | 417 |
by (etac disjI1 1); |
418 |
by (rtac disjI2 1); |
|
419 |
by (rtac allI 1); |
|
420 |
by (hyp_subst_tac 1); |
|
10834 | 421 |
by (res_inst_tac [("a","f$(UU::'a)")] (refl RS box_equals) 1); |
9245 | 422 |
by (res_inst_tac [("fo5","f")] ((minimal RS monofun_cfun_arg) RS (ax_flat RS spec RS spec RS mp) RS disjE) 1); |
423 |
by (contr_tac 1); |
|
424 |
by (atac 1); |
|
425 |
by (res_inst_tac [("fo5","f")] ((minimal RS monofun_cfun_arg) RS (ax_flat RS spec RS spec RS mp) RS disjE) 1); |
|
426 |
by (contr_tac 1); |
|
427 |
by (atac 1); |
|
428 |
qed "flat_codom"; |
|
3326 | 429 |
|
3327 | 430 |
|
431 |
(* ------------------------------------------------------------------------ *) |
|
432 |
(* Access to definitions *) |
|
433 |
(* ------------------------------------------------------------------------ *) |
|
434 |
||
435 |
||
10834 | 436 |
Goalw [ID_def] "ID$x=x"; |
9245 | 437 |
by (stac beta_cfun 1); |
438 |
by (rtac cont_id 1); |
|
439 |
by (rtac refl 1); |
|
440 |
qed "ID1"; |
|
3327 | 441 |
|
10834 | 442 |
Goalw [oo_def] "(f oo g)=(LAM x. f$(g$x))"; |
9245 | 443 |
by (stac beta_cfun 1); |
444 |
by (Simp_tac 1); |
|
445 |
by (stac beta_cfun 1); |
|
446 |
by (Simp_tac 1); |
|
447 |
by (rtac refl 1); |
|
448 |
qed "cfcomp1"; |
|
3327 | 449 |
|
10834 | 450 |
Goal "(f oo g)$x=f$(g$x)"; |
9245 | 451 |
by (stac cfcomp1 1); |
452 |
by (stac beta_cfun 1); |
|
453 |
by (Simp_tac 1); |
|
454 |
by (rtac refl 1); |
|
455 |
qed "cfcomp2"; |
|
3327 | 456 |
|
457 |
||
458 |
(* ------------------------------------------------------------------------ *) |
|
459 |
(* Show that interpretation of (pcpo,_->_) is a category *) |
|
460 |
(* The class of objects is interpretation of syntactical class pcpo *) |
|
461 |
(* The class of arrows between objects 'a and 'b is interpret. of 'a -> 'b *) |
|
462 |
(* The identity arrow is interpretation of ID *) |
|
463 |
(* The composition of f and g is interpretation of oo *) |
|
464 |
(* ------------------------------------------------------------------------ *) |
|
465 |
||
466 |
||
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
467 |
Goal "f oo ID = f "; |
9245 | 468 |
by (rtac ext_cfun 1); |
469 |
by (stac cfcomp2 1); |
|
470 |
by (stac ID1 1); |
|
471 |
by (rtac refl 1); |
|
472 |
qed "ID2"; |
|
3327 | 473 |
|
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
474 |
Goal "ID oo f = f "; |
9245 | 475 |
by (rtac ext_cfun 1); |
476 |
by (stac cfcomp2 1); |
|
477 |
by (stac ID1 1); |
|
478 |
by (rtac refl 1); |
|
479 |
qed "ID3"; |
|
3327 | 480 |
|
481 |
||
9248
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
paulson
parents:
9245
diff
changeset
|
482 |
Goal "f oo (g oo h) = (f oo g) oo h"; |
9245 | 483 |
by (rtac ext_cfun 1); |
10834 | 484 |
by (res_inst_tac [("s","f$(g$(h$x))")] trans 1); |
9245 | 485 |
by (stac cfcomp2 1); |
486 |
by (stac cfcomp2 1); |
|
487 |
by (rtac refl 1); |
|
488 |
by (stac cfcomp2 1); |
|
489 |
by (stac cfcomp2 1); |
|
490 |
by (rtac refl 1); |
|
491 |
qed "assoc_oo"; |
|
3327 | 492 |
|
493 |
(* ------------------------------------------------------------------------ *) |
|
494 |
(* Merge the different rewrite rules for the simplifier *) |
|
495 |
(* ------------------------------------------------------------------------ *) |
|
496 |
||
497 |
Addsimps ([ID1,ID2,ID3,cfcomp2]); |
|
498 |
||
499 |