| author | wenzelm | 
| Thu, 02 Mar 2000 18:18:10 +0100 | |
| changeset 8328 | efbcec3eb02f | 
| parent 5291 | 5706f0ef1d43 | 
| child 9245 | 428385c4bc50 | 
| permissions | -rw-r--r-- | 
| 2357 | 1  | 
(* Title: HOLCF/Lift3.ML  | 
2  | 
ID: $Id$  | 
|
| 3035 | 3  | 
Author: Olaf Mueller  | 
| 2357 | 4  | 
Copyright 1996 Technische Universitaet Muenchen  | 
5  | 
||
6  | 
Theorems for Lift3.thy  | 
|
7  | 
*)  | 
|
8  | 
||
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
9  | 
|
| 2640 | 10  | 
(* for compatibility with old HOLCF-Version *)  | 
11  | 
qed_goal "inst_lift_pcpo" thy "UU = Undef"  | 
|
12  | 
(fn prems =>  | 
|
13  | 
[  | 
|
14  | 
(simp_tac (HOL_ss addsimps [UU_def,UU_lift_def]) 1)  | 
|
15  | 
]);  | 
|
16  | 
||
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
17  | 
(* ----------------------------------------------------------- *)  | 
| 3035 | 18  | 
(* In lift.simps Undef is replaced by UU *)  | 
19  | 
(* Undef should be invisible from now on *)  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
20  | 
(* ----------------------------------------------------------- *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
21  | 
|
| 3035 | 22  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
23  | 
Addsimps [inst_lift_pcpo];  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
25  | 
local  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
26  | 
|
| 3035 | 27  | 
val case1' = prove_goal thy "lift_case f1 f2 UU = f1"  | 
| 4098 | 28  | 
(fn _ => [simp_tac (simpset() addsimps lift.simps) 1]);  | 
| 3035 | 29  | 
val case2' = prove_goal thy "lift_case f1 f2 (Def a) = f2 a"  | 
| 3041 | 30  | 
(fn _ => [Simp_tac 1]);  | 
| 3035 | 31  | 
val distinct1' = prove_goal thy "UU ~= Def a"  | 
| 3041 | 32  | 
(fn _ => [Simp_tac 1]);  | 
| 3035 | 33  | 
val distinct2' = prove_goal thy "Def a ~= UU"  | 
| 3041 | 34  | 
(fn _ => [Simp_tac 1]);  | 
| 3035 | 35  | 
val inject' = prove_goal thy "Def a = Def aa = (a = aa)"  | 
| 3041 | 36  | 
(fn _ => [Simp_tac 1]);  | 
| 3035 | 37  | 
val rec1' = prove_goal thy "lift_rec f1 f2 UU = f1"  | 
| 3041 | 38  | 
(fn _ => [Simp_tac 1]);  | 
| 3035 | 39  | 
val rec2' = prove_goal thy "lift_rec f1 f2 (Def a) = f2 a"  | 
| 3041 | 40  | 
(fn _ => [Simp_tac 1]);  | 
| 3035 | 41  | 
val induct' = prove_goal thy "[| P UU; !a. P (Def a) |] ==> P lift"  | 
| 3041 | 42  | 
(fn prems => [cut_facts_tac prems 1, Asm_full_simp_tac 1,  | 
43  | 
etac Lift1.lift.induct 1,fast_tac HOL_cs 1]);  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
45  | 
in  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
46  | 
|
| 3035 | 47  | 
val Def_not_UU = distinct2';  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
49  | 
structure lift =  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
50  | 
struct  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
51  | 
val cases = [case1',case2'];  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
52  | 
val distinct = [distinct1',distinct2'];  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
53  | 
val inject = [inject'];  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
54  | 
val induct = allI RSN(2,induct');  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
55  | 
val recs = [rec1',rec2'];  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
56  | 
val simps = cases@distinct@inject@recs;  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
57  | 
fun induct_tac (s:string) (i:int) =  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
58  | 
    (res_inst_tac [("lift",s)] induct i);
 | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
59  | 
end;  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
61  | 
end; (* local *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
62  | 
|
| 3250 | 63  | 
Delsimps Lift1.lift.simps;  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
64  | 
Delsimps [inst_lift_pcpo];  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
65  | 
Addsimps [inst_lift_pcpo RS sym];  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
66  | 
Addsimps lift.simps;  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
68  | 
|
| 3035 | 69  | 
(* --------------------------------------------------------*)  | 
70  | 
section"less_lift";  | 
|
71  | 
(* --------------------------------------------------------*)  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
72  | 
|
| 5068 | 73  | 
Goal "(x::'a lift) << y = (x=y | x=UU)";  | 
| 3457 | 74  | 
by (stac inst_lift_po 1);  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
75  | 
by (Simp_tac 1);  | 
| 3035 | 76  | 
qed"less_lift";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
78  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
79  | 
(* ---------------------------------------------------------- *)  | 
| 3035 | 80  | 
section"UU and Def";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
81  | 
(* ---------------------------------------------------------- *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
82  | 
|
| 5068 | 83  | 
Goal "x=UU | (? y. x=Def y)";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
84  | 
by (lift.induct_tac "x" 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
85  | 
by (Asm_simp_tac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
86  | 
by (rtac disjI2 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
87  | 
by (rtac exI 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
88  | 
by (Asm_simp_tac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
89  | 
qed"Lift_exhaust";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
90  | 
|
| 3035 | 91  | 
val prems = goal thy  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
92  | 
"[| x = UU ==> P; ? a. x = Def a ==> P |] ==> P";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
93  | 
by (cut_facts_tac [Lift_exhaust] 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
94  | 
by (fast_tac (HOL_cs addSEs prems) 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
95  | 
qed"Lift_cases";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
96  | 
|
| 5068 | 97  | 
Goal "(x~=UU)=(? y. x=Def y)";  | 
| 3457 | 98  | 
by (rtac iffI 1);  | 
99  | 
by (rtac Lift_cases 1);  | 
|
| 3035 | 100  | 
by (REPEAT (fast_tac (HOL_cs addSIs lift.distinct) 1));  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
101  | 
qed"not_Undef_is_Def";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
102  | 
|
| 3035 | 103  | 
(* For x~=UU in assumptions def_tac replaces x by (Def a) in conclusion *)  | 
104  | 
val def_tac = etac (not_Undef_is_Def RS iffD1 RS exE) THEN' Asm_simp_tac;  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
105  | 
|
| 3035 | 106  | 
bind_thm("Undef_eq_UU", inst_lift_pcpo RS sym);
 | 
107  | 
||
108  | 
val DefE = prove_goal thy "Def x = UU ==> R"  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
109  | 
(fn prems => [  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
110  | 
cut_facts_tac prems 1,  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
111  | 
asm_full_simp_tac (HOL_ss addsimps [Def_not_UU]) 1]);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
112  | 
|
| 3035 | 113  | 
val prems = goal thy "[| x = Def s; x = UU |] ==> R";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
114  | 
by (cut_facts_tac prems 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
115  | 
by (fast_tac (HOL_cs addSDs [DefE]) 1);  | 
| 3035 | 116  | 
qed"DefE2";  | 
117  | 
||
| 5068 | 118  | 
Goal "Def x << Def y = (x = y)";  | 
| 3035 | 119  | 
by (stac (hd lift.inject RS sym) 1);  | 
120  | 
back();  | 
|
121  | 
by (rtac iffI 1);  | 
|
| 4098 | 122  | 
by (asm_full_simp_tac (simpset() addsimps [inst_lift_po] ) 1);  | 
| 3457 | 123  | 
by (etac (antisym_less_inverse RS conjunct1) 1);  | 
| 3035 | 124  | 
qed"Def_inject_less_eq";  | 
125  | 
||
| 5068 | 126  | 
Goal "Def x << y = (Def x = y)";  | 
| 4098 | 127  | 
by (simp_tac (simpset() addsimps [less_lift]) 1);  | 
| 3035 | 128  | 
qed"Def_less_is_eq";  | 
129  | 
||
130  | 
Addsimps [Def_less_is_eq];  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
131  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
132  | 
(* ---------------------------------------------------------- *)  | 
| 3035 | 133  | 
section"Lift is flat";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
134  | 
(* ---------------------------------------------------------- *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
135  | 
|
| 5068 | 136  | 
Goal "! x y::'a lift. x << y --> x = UU | x = y";  | 
| 4098 | 137  | 
by (simp_tac (simpset() addsimps [less_lift]) 1);  | 
| 3324 | 138  | 
qed"ax_flat_lift";  | 
| 3035 | 139  | 
|
140  | 
(* Two specific lemmas for the combination of LCF and HOL terms *)  | 
|
141  | 
||
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5068 
diff
changeset
 | 
142  | 
Goal "[|cont g; cont f|] ==> cont(%x. ((f x)`(g x)) s)";  | 
| 3457 | 143  | 
by (rtac cont2cont_CF1L 1);  | 
| 3035 | 144  | 
by (REPEAT (resolve_tac cont_lemmas1 1));  | 
| 
4477
 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 
paulson 
parents: 
4098 
diff
changeset
 | 
145  | 
by Auto_tac;  | 
| 5291 | 146  | 
qed"cont_Rep_CFun_app";  | 
| 3035 | 147  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5068 
diff
changeset
 | 
148  | 
Goal "[|cont g; cont f|] ==> cont(%x. ((f x)`(g x)) s t)";  | 
| 3457 | 149  | 
by (rtac cont2cont_CF1L 1);  | 
| 5291 | 150  | 
by (etac cont_Rep_CFun_app 1);  | 
| 3457 | 151  | 
by (assume_tac 1);  | 
| 5291 | 152  | 
qed"cont_Rep_CFun_app_app";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
154  | 
|
| 3035 | 155  | 
(* continuity of if then else *)  | 
156  | 
val prems = goal thy "[| cont f1; cont f2 |] ==> \  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
157  | 
\ cont (%x. if b then f1 x else f2 x)";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
158  | 
by (cut_facts_tac prems 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
159  | 
by (case_tac "b" 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
160  | 
by (TRYALL (fast_tac (HOL_cs addss HOL_ss)));  | 
| 3035 | 161  | 
qed"cont_if";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
162  |