| author | nipkow | 
| Fri, 28 Jul 2000 16:02:51 +0200 | |
| changeset 9458 | c613cd06d5cf | 
| parent 9248 | e1dee89de037 | 
| 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 | ||
| 9245 | 6 | Class Instance lift::(term)po | 
| 2357 | 7 | *) | 
| 8 | ||
| 2640 | 9 | (* for compatibility with old HOLCF-Version *) | 
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 10 | Goal "(op <<)=(%x y. x=y|x=Undef)"; | 
| 9245 | 11 | by (fold_goals_tac [less_lift_def]); | 
| 12 | by (rtac refl 1); | |
| 13 | qed "inst_lift_po"; | |
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 14 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 15 | (* -------------------------------------------------------------------------*) | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 16 | (* type ('a)lift is pointed                                                *)
 | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 17 | (* ------------------------------------------------------------------------ *) | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 18 | |
| 5068 | 19 | Goal "Undef << x"; | 
| 4098 | 20 | by (simp_tac (simpset() addsimps [inst_lift_po]) 1); | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 21 | qed"minimal_lift"; | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 22 | |
| 2640 | 23 | bind_thm ("UU_lift_def",minimal_lift RS minimal2UU RS sym);
 | 
| 24 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 25 | AddIffs [minimal_lift]; | 
| 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 26 | |
| 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 27 | Goal "EX x::'a lift. ALL y. x<<y"; | 
| 9245 | 28 | by (res_inst_tac [("x","Undef")] exI 1);
 | 
| 29 | by (rtac (minimal_lift RS allI) 1); | |
| 30 | qed "least_lift"; | |
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 31 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 32 | (* ------------------------------------------------------------------------ *) | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 33 | (* ('a)lift is a cpo                                                       *)
 | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 34 | (* ------------------------------------------------------------------------ *) | 
| 
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 | (* 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 | 37 | 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 | 38 | 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 | 39 | least element Undef. *) | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 40 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 41 | (* Tailoring notUU_I of Pcpo.ML to Undef *) | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 42 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5068diff
changeset | 43 | Goal "[| x<<y; ~x=Undef |] ==> ~y=Undef"; | 
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 44 | by (blast_tac (claset() addIs [antisym_less]) 1); | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 45 | qed"notUndef_I"; | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 46 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 47 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 48 | (* Tailoring chain_mono2 of Pcpo.ML to Undef *) | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 49 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 50 | Goal "[| EX j.~Y(j)=Undef; chain(Y::nat=>('a)lift) |] \
 | 
| 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 51 | \ ==> EX j. ALL i. j<i-->~Y(i)=Undef"; | 
| 3726 | 52 | by Safe_tac; | 
| 53 | by (Step_tac 1); | |
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 54 | by (strip_tac 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 55 | by (rtac notUndef_I 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 56 | by (atac 2); | 
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 57 | by (etac (chain_mono) 1); | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 58 | by (atac 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 59 | qed"chain_mono2_po"; | 
| 
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 | |
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
4098diff
changeset | 62 | (* Tailoring flat_imp_chfin of Fix.ML to lift *) | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 63 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 64 | Goal "(ALL Y. chain(Y::nat=>('a)lift)-->(EX n. max_in_chain n Y))";
 | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 65 | by (rewtac max_in_chain_def); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 66 | by (strip_tac 1); | 
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 67 | by (res_inst_tac [("P","ALL i. Y(i)=Undef")] case_split_thm  1);
 | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 68 | by (res_inst_tac [("x","0")] exI 1);
 | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 69 | by (strip_tac 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 70 | by (rtac trans 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 71 | by (etac spec 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 72 | by (rtac sym 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 73 | by (etac spec 1); | 
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 74 | by (subgoal_tac "ALL x y. x<<(y::('a)lift) --> x=Undef | x=y" 1);
 | 
| 4098 | 75 | by (simp_tac (simpset() addsimps [inst_lift_po]) 2); | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 76 | by (rtac (chain_mono2_po RS exE) 1); | 
| 3726 | 77 | by (Fast_tac 1); | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 78 | by (atac 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 79 | by (res_inst_tac [("x","Suc(x)")] exI 1); 
 | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 80 | by (strip_tac 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 81 | by (rtac disjE 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 82 | by (atac 3); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 83 | by (rtac mp 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 84 | by (dtac spec 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 85 | by (etac spec 1); | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 86 | by (etac (le_imp_less_or_eq RS disjE) 1); | 
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 87 | by (etac (chain_mono) 1); | 
| 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 88 | by Auto_tac; | 
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
4098diff
changeset | 89 | qed"flat_imp_chfin_poo"; | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 90 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 91 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 92 | (* Main Lemma: cpo_lift *) | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 93 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 94 | Goal "chain(Y::nat=>('a)lift) ==> EX x. range(Y) <<|x";
 | 
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
4098diff
changeset | 95 | by (cut_inst_tac [] flat_imp_chfin_poo 1); | 
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 96 | by (blast_tac (claset() addIs [lub_finch1]) 1); | 
| 2356 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 97 | qed"cpo_lift"; | 
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 98 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 99 | |
| 
125260ef480c
Theories Lift1, Lift2 and Lift3 inserted below HOLCF.thy
 sandnerr parents: diff
changeset | 100 |