| author | wenzelm | 
| Thu, 02 Mar 2000 18:18:10 +0100 | |
| changeset 8328 | efbcec3eb02f | 
| parent 5143 | b94cd208f073 | 
| child 9245 | 428385c4bc50 | 
| permissions | -rw-r--r-- | 
| 2357 | 1  | 
(* Title: HOLCF/Lift2.ML  | 
2  | 
ID: $Id$  | 
|
| 3033 | 3  | 
Author: Olaf Mueller  | 
| 2357 | 4  | 
Copyright 1996 Technische Universitaet Muenchen  | 
5  | 
||
6  | 
Theorems for Lift2.thy  | 
|
7  | 
*)  | 
|
8  | 
||
| 2640 | 9  | 
open Lift2;  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
10  | 
|
| 2640 | 11  | 
(* for compatibility with old HOLCF-Version *)  | 
| 3842 | 12  | 
qed_goal "inst_lift_po" thy "(op <<)=(%x y. x=y|x=Undef)"  | 
| 2640 | 13  | 
(fn prems =>  | 
14  | 
[  | 
|
| 
3323
 
194ae2e0c193
eliminated the constant less by the introduction of the axclass sq_ord
 
slotosch 
parents: 
3033 
diff
changeset
 | 
15  | 
(fold_goals_tac [less_lift_def]),  | 
| 2640 | 16  | 
(rtac refl 1)  | 
17  | 
]);  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
19  | 
(* -------------------------------------------------------------------------*)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
20  | 
(* type ('a)lift is pointed                                                *)
 | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
21  | 
(* ------------------------------------------------------------------------ *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
22  | 
|
| 5068 | 23  | 
Goal "Undef << x";  | 
| 4098 | 24  | 
by (simp_tac (simpset() addsimps [inst_lift_po]) 1);  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
25  | 
qed"minimal_lift";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
26  | 
|
| 2640 | 27  | 
bind_thm ("UU_lift_def",minimal_lift RS minimal2UU RS sym);
 | 
28  | 
||
| 3842 | 29  | 
qed_goal "least_lift" thy "? x::'a lift.!y. x<<y"  | 
| 2640 | 30  | 
(fn prems =>  | 
31  | 
[  | 
|
32  | 
        (res_inst_tac [("x","Undef")] exI 1),
 | 
|
33  | 
(rtac (minimal_lift RS allI) 1)  | 
|
34  | 
]);  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
35  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
36  | 
(* ------------------------------------------------------------------------ *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
37  | 
(* ('a)lift is a cpo                                                       *)
 | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
38  | 
(* ------------------------------------------------------------------------ *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
40  | 
(* The following Lemmata have already been proved in Pcpo.ML and Fix.ML,  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
41  | 
but there class pcpo is assumed, although only po is necessary and a  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
42  | 
least element. Therefore they are redone here for the po lift with  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
43  | 
least element Undef. *)  | 
| 
 
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  | 
(* Tailoring notUU_I of Pcpo.ML to Undef *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
46  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5068 
diff
changeset
 | 
47  | 
Goal "[| x<<y; ~x=Undef |] ==> ~y=Undef";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
48  | 
by (etac contrapos 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
49  | 
by (hyp_subst_tac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
50  | 
by (rtac antisym_less 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
51  | 
by (atac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
52  | 
by (rtac minimal_lift 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
53  | 
qed"notUndef_I";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
56  | 
(* Tailoring chain_mono2 of Pcpo.ML to Undef *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
57  | 
|
| 5068 | 58  | 
Goal  | 
| 
4721
 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 
oheimb 
parents: 
4098 
diff
changeset
 | 
59  | 
"!!Y. [|? j.~Y(j)=Undef;chain(Y::nat=>('a)lift)|] \
 | 
| 3842 | 60  | 
\ ==> ? j.!i. j<i-->~Y(i)=Undef";  | 
| 3726 | 61  | 
by Safe_tac;  | 
62  | 
by (Step_tac 1);  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
63  | 
by (strip_tac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
64  | 
by (rtac notUndef_I 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
65  | 
by (atac 2);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
66  | 
by (etac (chain_mono RS mp) 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
67  | 
by (atac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
68  | 
qed"chain_mono2_po";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
70  | 
|
| 
4721
 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 
oheimb 
parents: 
4098 
diff
changeset
 | 
71  | 
(* Tailoring flat_imp_chfin of Fix.ML to lift *)  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
72  | 
|
| 5068 | 73  | 
Goal  | 
| 
4721
 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 
oheimb 
parents: 
4098 
diff
changeset
 | 
74  | 
        "(! Y. chain(Y::nat=>('a)lift)-->(? n. max_in_chain n Y))";
 | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
75  | 
by (rewtac max_in_chain_def);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
76  | 
by (strip_tac 1);  | 
| 3842 | 77  | 
by (res_inst_tac [("P","!i. Y(i)=Undef")] case_split_thm  1);
 | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
78  | 
by (res_inst_tac [("x","0")] exI 1);
 | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
79  | 
by (strip_tac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
80  | 
by (rtac trans 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
81  | 
by (etac spec 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
82  | 
by (rtac sym 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
83  | 
by (etac spec 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
84  | 
|
| 3842 | 85  | 
by (subgoal_tac "!x y. x<<(y::('a)lift) --> x=Undef | x=y" 1);
 | 
| 4098 | 86  | 
by (simp_tac (simpset() addsimps [inst_lift_po]) 2);  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
87  | 
by (rtac (chain_mono2_po RS exE) 1);  | 
| 3726 | 88  | 
by (Fast_tac 1);  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
89  | 
by (atac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
90  | 
by (res_inst_tac [("x","Suc(x)")] exI 1); 
 | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
91  | 
by (strip_tac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
92  | 
by (rtac disjE 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
93  | 
by (atac 3);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
94  | 
by (rtac mp 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
95  | 
by (dtac spec 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
96  | 
by (etac spec 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
97  | 
by (etac (le_imp_less_or_eq RS disjE) 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
98  | 
by (etac (chain_mono RS mp) 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
99  | 
by (atac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
100  | 
by (hyp_subst_tac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
101  | 
by (rtac refl_less 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
102  | 
by (res_inst_tac [("P","Y(Suc(x)) = Undef")] notE 1); 
 | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
103  | 
by (atac 2);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
104  | 
by (rtac mp 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
105  | 
by (etac spec 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
106  | 
by (Asm_simp_tac 1);  | 
| 
4721
 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 
oheimb 
parents: 
4098 
diff
changeset
 | 
107  | 
qed"flat_imp_chfin_poo";  | 
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
108  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
110  | 
(* Main Lemma: cpo_lift *)  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
111  | 
|
| 5068 | 112  | 
Goal  | 
| 
4721
 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 
oheimb 
parents: 
4098 
diff
changeset
 | 
113  | 
  "!!Y. chain(Y::nat=>('a)lift) ==> ? x. range(Y) <<|x";
 | 
| 
 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 
oheimb 
parents: 
4098 
diff
changeset
 | 
114  | 
by (cut_inst_tac [] flat_imp_chfin_poo 1);  | 
| 3726 | 115  | 
by (Step_tac 1);  | 
116  | 
by Safe_tac;  | 
|
117  | 
by (Step_tac 1);  | 
|
| 
2356
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
118  | 
by (rtac exI 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
119  | 
by (rtac lub_finch1 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
120  | 
by (atac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
121  | 
by (atac 1);  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
122  | 
qed"cpo_lift";  | 
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
123  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
124  | 
|
| 
 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 
sandnerr 
parents:  
diff
changeset
 | 
125  |