| author | paulson | 
| Tue, 23 Dec 2003 18:26:03 +0100 | |
| changeset 14326 | c817dd885a32 | 
| parent 12030 | 46d57d0290a2 | 
| child 14981 | e73f8140af78 | 
| permissions | -rw-r--r-- | 
| 2278 | 1 | (* Title: HOLCF/Up2.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Franz Regensburger | |
| 12030 | 4 | License: GPL (GNU GENERAL PUBLIC LICENSE) | 
| 2278 | 5 | |
| 9245 | 6 | Class Instance u::(pcpo)po | 
| 2278 | 7 | *) | 
| 8 | ||
| 2640 | 9 | (* for compatibility with old HOLCF-Version *) | 
| 9169 | 10 | Goal "(op <<)=(%x1 x2. case Rep_Up(x1) of \ | 
| 2640 | 11 | \ Inl(y1) => True \ | 
| 12 | \ | Inr(y2) => (case Rep_Up(x2) of Inl(z1) => False \ | |
| 9169 | 13 | \ | Inr(z2) => y2<<z2))"; | 
| 14 | by (fold_goals_tac [less_up_def]); | |
| 15 | by (rtac refl 1); | |
| 16 | qed "inst_up_po"; | |
| 2640 | 17 | |
| 2278 | 18 | (* -------------------------------------------------------------------------*) | 
| 19 | (* type ('a)u is pointed                                                    *)
 | |
| 20 | (* ------------------------------------------------------------------------ *) | |
| 21 | ||
| 9169 | 22 | Goal "Abs_Up(Inl ()) << z"; | 
| 23 | by (simp_tac (simpset() addsimps [less_up1a]) 1); | |
| 24 | qed "minimal_up"; | |
| 2640 | 25 | |
| 26 | bind_thm ("UU_up_def",minimal_up RS minimal2UU RS sym);
 | |
| 27 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 28 | Goal "EX x::'a u. ALL y. x<<y"; | 
| 9169 | 29 | by (res_inst_tac [("x","Abs_Up(Inl ())")] exI 1);
 | 
| 30 | by (rtac (minimal_up RS allI) 1); | |
| 31 | qed "least_up"; | |
| 2278 | 32 | |
| 33 | (* -------------------------------------------------------------------------*) | |
| 34 | (* access to less_up in class po *) | |
| 35 | (* ------------------------------------------------------------------------ *) | |
| 36 | ||
| 9169 | 37 | Goal "~ Iup(x) << Abs_Up(Inl ())"; | 
| 38 | by (simp_tac (simpset() addsimps [less_up1b]) 1); | |
| 39 | qed "less_up2b"; | |
| 2278 | 40 | |
| 9169 | 41 | Goal "(Iup(x)<<Iup(y)) = (x<<y)"; | 
| 42 | by (simp_tac (simpset() addsimps [less_up1c]) 1); | |
| 43 | qed "less_up2c"; | |
| 2278 | 44 | |
| 45 | (* ------------------------------------------------------------------------ *) | |
| 46 | (* Iup and Ifup are monotone *) | |
| 47 | (* ------------------------------------------------------------------------ *) | |
| 48 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 49 | Goalw [monofun] "monofun(Iup)"; | 
| 9245 | 50 | by (strip_tac 1); | 
| 51 | by (etac (less_up2c RS iffD2) 1); | |
| 52 | qed "monofun_Iup"; | |
| 53 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 54 | Goalw [monofun] "monofun(Ifup)"; | 
| 9245 | 55 | by (strip_tac 1); | 
| 56 | by (rtac (less_fun RS iffD2) 1); | |
| 57 | by (strip_tac 1); | |
| 58 | by (res_inst_tac [("p","xa")] upE 1);
 | |
| 59 | by (asm_simp_tac Up0_ss 1); | |
| 60 | by (asm_simp_tac Up0_ss 1); | |
| 61 | by (etac monofun_cfun_fun 1); | |
| 62 | qed "monofun_Ifup1"; | |
| 2278 | 63 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 64 | Goalw [monofun] "monofun(Ifup(f))"; | 
| 9245 | 65 | by (strip_tac 1); | 
| 66 | by (res_inst_tac [("p","x")] upE 1);
 | |
| 67 | by (asm_simp_tac Up0_ss 1); | |
| 68 | by (asm_simp_tac Up0_ss 1); | |
| 69 | by (res_inst_tac [("p","y")] upE 1);
 | |
| 70 | by (hyp_subst_tac 1); | |
| 71 | by (rtac notE 1); | |
| 72 | by (rtac less_up2b 1); | |
| 73 | by (atac 1); | |
| 74 | by (asm_simp_tac Up0_ss 1); | |
| 75 | by (rtac monofun_cfun_arg 1); | |
| 76 | by (hyp_subst_tac 1); | |
| 77 | by (etac (less_up2c RS iffD1) 1); | |
| 78 | qed "monofun_Ifup2"; | |
| 2278 | 79 | |
| 80 | (* ------------------------------------------------------------------------ *) | |
| 81 | (* Some kind of surjectivity lemma *) | |
| 82 | (* ------------------------------------------------------------------------ *) | |
| 83 | ||
| 9169 | 84 | Goal "z=Iup(x) ==> Iup(Ifup(LAM x. x)(z)) = z"; | 
| 85 | by (asm_simp_tac Up0_ss 1); | |
| 86 | qed "up_lemma1"; | |
| 2278 | 87 | |
| 88 | (* ------------------------------------------------------------------------ *) | |
| 89 | (* ('a)u is a cpo                                                           *)
 | |
| 90 | (* ------------------------------------------------------------------------ *) | |
| 91 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 92 | Goal "[|chain(Y);EX i x. Y(i)=Iup(x)|] \ | 
| 9169 | 93 | \ ==> range(Y) <<| Iup(lub(range(%i.(Ifup (LAM x. x) (Y(i))))))"; | 
| 94 | by (rtac is_lubI 1); | |
| 95 | by (rtac ub_rangeI 1); | |
| 96 | by (res_inst_tac [("p","Y(i)")] upE 1);
 | |
| 97 | by (res_inst_tac [("s","Abs_Up (Inl ())"),("t","Y(i)")] subst 1);
 | |
| 98 | by (etac sym 1); | |
| 99 | by (rtac minimal_up 1); | |
| 100 | by (res_inst_tac [("t","Y(i)")] (up_lemma1 RS subst) 1);
 | |
| 101 | by (atac 1); | |
| 102 | by (rtac (less_up2c RS iffD2) 1); | |
| 103 | by (rtac is_ub_thelub 1); | |
| 104 | by (etac (monofun_Ifup2 RS ch2ch_monofun) 1); | |
| 105 | by (strip_tac 1); | |
| 106 | by (res_inst_tac [("p","u")] upE 1);
 | |
| 107 | by (etac exE 1); | |
| 108 | by (etac exE 1); | |
| 109 | by (res_inst_tac [("P","Y(i)<<Abs_Up (Inl ())")] notE 1);
 | |
| 110 | by (res_inst_tac [("s","Iup(x)"),("t","Y(i)")] ssubst 1);
 | |
| 111 | by (atac 1); | |
| 112 | by (rtac less_up2b 1); | |
| 113 | by (hyp_subst_tac 1); | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 114 | by (etac ub_rangeD 1); | 
| 9169 | 115 | by (res_inst_tac [("t","u")] (up_lemma1 RS subst) 1);
 | 
| 116 | by (atac 1); | |
| 117 | by (rtac (less_up2c RS iffD2) 1); | |
| 118 | by (rtac is_lub_thelub 1); | |
| 119 | by (etac (monofun_Ifup2 RS ch2ch_monofun) 1); | |
| 120 | by (etac (monofun_Ifup2 RS ub2ub_monofun) 1); | |
| 121 | qed "lub_up1a"; | |
| 2278 | 122 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 123 | Goal "[|chain(Y); ALL i x. Y(i)~=Iup(x)|] ==> range(Y) <<| Abs_Up (Inl ())"; | 
| 9169 | 124 | by (rtac is_lubI 1); | 
| 125 | by (rtac ub_rangeI 1); | |
| 126 | by (res_inst_tac [("p","Y(i)")] upE 1);
 | |
| 127 | by (res_inst_tac [("s","Abs_Up (Inl ())"),("t","Y(i)")] ssubst 1);
 | |
| 128 | by (atac 1); | |
| 129 | by (rtac refl_less 1); | |
| 130 | by (rtac notE 1); | |
| 131 | by (dtac spec 1); | |
| 132 | by (dtac spec 1); | |
| 133 | by (atac 1); | |
| 134 | by (atac 1); | |
| 135 | by (strip_tac 1); | |
| 136 | by (rtac minimal_up 1); | |
| 137 | qed "lub_up1b"; | |
| 2278 | 138 | |
| 139 | bind_thm ("thelub_up1a", lub_up1a RS thelubI);
 | |
| 140 | (* | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 141 | [| chain ?Y1; EX i x. ?Y1 i = Iup x |] ==> | 
| 2278 | 142 | lub (range ?Y1) = Iup (lub (range (%i. Iup (LAM x. x) (?Y1 i)))) | 
| 143 | *) | |
| 144 | ||
| 145 | bind_thm ("thelub_up1b", lub_up1b RS thelubI);
 | |
| 146 | (* | |
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
4098diff
changeset | 147 | [| chain ?Y1; ! i x. ?Y1 i ~= Iup x |] ==> | 
| 2278 | 148 | lub (range ?Y1) = UU_up | 
| 149 | *) | |
| 150 | ||
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 151 | Goal "chain(Y::nat=>('a)u) ==> EX x. range(Y) <<|x";
 | 
| 9169 | 152 | by (rtac disjE 1); | 
| 153 | by (rtac exI 2); | |
| 154 | by (etac lub_up1a 2); | |
| 155 | by (atac 2); | |
| 156 | by (rtac exI 2); | |
| 157 | by (etac lub_up1b 2); | |
| 158 | by (atac 2); | |
| 159 | by (fast_tac HOL_cs 1); | |
| 160 | qed "cpo_up"; | |
| 2278 | 161 |