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