| author | wenzelm | 
| Sat, 02 Sep 2000 21:49:51 +0200 | |
| changeset 9803 | bc883b390d91 | 
| parent 9248 | e1dee89de037 | 
| child 12030 | 46d57d0290a2 | 
| permissions | -rw-r--r-- | 
| 2640 | 1 | (* Title: HOLCF/Ssum2.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 | |
| 9169 | 6 | Class Instance ++::(pcpo,pcpo)po | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 7 | *) | 
| 
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 *) | 
| 9169 | 10 | Goal "(op <<)=(%s1 s2.@z.\ | 
| 3842 | 11 | \ (! u x. s1=Isinl u & s2=Isinl x --> z = u << x)\ | 
| 12 | \ &(! v y. s1=Isinr v & s2=Isinr y --> z = v << y)\ | |
| 13 | \ &(! u y. s1=Isinl u & s2=Isinr y --> z = (u = UU))\ | |
| 9169 | 14 | \ &(! v x. s1=Isinr v & s2=Isinl x --> z = (v = UU)))"; | 
| 15 | by (fold_goals_tac [less_ssum_def]); | |
| 16 | by (rtac refl 1); | |
| 17 | qed "inst_ssum_po"; | |
| 2640 | 18 | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 19 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 20 | (* access to less_ssum in class po *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 21 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 22 | |
| 9169 | 23 | Goal "Isinl x << Isinl y = x << y"; | 
| 24 | by (simp_tac (simpset() addsimps [less_ssum2a]) 1); | |
| 25 | qed "less_ssum3a"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 26 | |
| 9169 | 27 | Goal "Isinr x << Isinr y = x << y"; | 
| 28 | by (simp_tac (simpset() addsimps [less_ssum2b]) 1); | |
| 29 | qed "less_ssum3b"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 30 | |
| 9169 | 31 | Goal "Isinl x << Isinr y = (x = UU)"; | 
| 32 | by (simp_tac (simpset() addsimps [less_ssum2c]) 1); | |
| 33 | qed "less_ssum3c"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 34 | |
| 9169 | 35 | Goal "Isinr x << Isinl y = (x = UU)"; | 
| 36 | by (simp_tac (simpset() addsimps [less_ssum2d]) 1); | |
| 37 | qed "less_ssum3d"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 38 | |
| 
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 | (* type ssum ++ is pointed *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 41 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 42 | |
| 9169 | 43 | Goal "Isinl UU << s"; | 
| 44 | by (res_inst_tac [("p","s")] IssumE2 1);
 | |
| 45 | by (hyp_subst_tac 1); | |
| 46 | by (rtac (less_ssum3a RS iffD2) 1); | |
| 47 | by (rtac minimal 1); | |
| 48 | by (hyp_subst_tac 1); | |
| 49 | by (stac strict_IsinlIsinr 1); | |
| 50 | by (rtac (less_ssum3b RS iffD2) 1); | |
| 51 | by (rtac minimal 1); | |
| 52 | qed "minimal_ssum"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 53 | |
| 2640 | 54 | bind_thm ("UU_ssum_def",minimal_ssum RS minimal2UU RS sym);
 | 
| 55 | ||
| 9169 | 56 | Goal "? x::'a++'b.!y. x<<y"; | 
| 57 | by (res_inst_tac [("x","Isinl UU")] exI 1);
 | |
| 58 | by (rtac (minimal_ssum RS allI) 1); | |
| 59 | qed "least_ssum"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 60 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 61 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 62 | (* Isinl, Isinr are monotone *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 63 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 64 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 65 | Goalw [monofun] "monofun(Isinl)"; | 
| 9245 | 66 | by (strip_tac 1); | 
| 67 | by (etac (less_ssum3a RS iffD2) 1); | |
| 68 | qed "monofun_Isinl"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 69 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 70 | Goalw [monofun] "monofun(Isinr)"; | 
| 9245 | 71 | by (strip_tac 1); | 
| 72 | by (etac (less_ssum3b RS iffD2) 1); | |
| 73 | qed "monofun_Isinr"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 74 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 75 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 76 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 77 | (* Iwhen is monotone in all arguments *) | 
| 
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 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 81 | Goalw [monofun] "monofun(Iwhen)"; | 
| 9245 | 82 | by (strip_tac 1); | 
| 83 | by (rtac (less_fun RS iffD2) 1); | |
| 84 | by (strip_tac 1); | |
| 85 | by (rtac (less_fun RS iffD2) 1); | |
| 86 | by (strip_tac 1); | |
| 87 | by (res_inst_tac [("p","xb")] IssumE 1);
 | |
| 88 | by (hyp_subst_tac 1); | |
| 89 | by (asm_simp_tac Ssum0_ss 1); | |
| 90 | by (asm_simp_tac Ssum0_ss 1); | |
| 91 | by (etac monofun_cfun_fun 1); | |
| 92 | by (asm_simp_tac Ssum0_ss 1); | |
| 93 | qed "monofun_Iwhen1"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 94 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 95 | Goalw [monofun] "monofun(Iwhen(f))"; | 
| 9245 | 96 | by (strip_tac 1); | 
| 97 | by (rtac (less_fun RS iffD2) 1); | |
| 98 | by (strip_tac 1); | |
| 99 | by (res_inst_tac [("p","xa")] IssumE 1);
 | |
| 100 | by (hyp_subst_tac 1); | |
| 101 | by (asm_simp_tac Ssum0_ss 1); | |
| 102 | by (asm_simp_tac Ssum0_ss 1); | |
| 103 | by (asm_simp_tac Ssum0_ss 1); | |
| 104 | by (etac monofun_cfun_fun 1); | |
| 105 | qed "monofun_Iwhen2"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 106 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 107 | Goalw [monofun] "monofun(Iwhen(f)(g))"; | 
| 9245 | 108 | by (strip_tac 1); | 
| 109 | by (res_inst_tac [("p","x")] IssumE 1);
 | |
| 110 | by (hyp_subst_tac 1); | |
| 111 | by (asm_simp_tac Ssum0_ss 1); | |
| 112 | by (hyp_subst_tac 1); | |
| 113 | by (res_inst_tac [("p","y")] IssumE 1);
 | |
| 114 | by (hyp_subst_tac 1); | |
| 115 | by (asm_simp_tac Ssum0_ss 1); | |
| 116 | by (res_inst_tac  [("P","xa=UU")] notE 1);
 | |
| 117 | by (atac 1); | |
| 118 | by (rtac UU_I 1); | |
| 119 | by (rtac (less_ssum3a RS iffD1) 1); | |
| 120 | by (atac 1); | |
| 121 | by (hyp_subst_tac 1); | |
| 122 | by (asm_simp_tac Ssum0_ss 1); | |
| 123 | by (rtac monofun_cfun_arg 1); | |
| 124 | by (etac (less_ssum3a RS iffD1) 1); | |
| 125 | by (hyp_subst_tac 1); | |
| 126 | by (res_inst_tac [("s","UU"),("t","xa")] subst 1);
 | |
| 127 | by (etac (less_ssum3c RS iffD1 RS sym) 1); | |
| 128 | by (asm_simp_tac Ssum0_ss 1); | |
| 129 | by (hyp_subst_tac 1); | |
| 130 | by (res_inst_tac [("p","y")] IssumE 1);
 | |
| 131 | by (hyp_subst_tac 1); | |
| 132 | by (res_inst_tac [("s","UU"),("t","ya")] subst 1);
 | |
| 133 | by (etac (less_ssum3d RS iffD1 RS sym) 1); | |
| 134 | by (asm_simp_tac Ssum0_ss 1); | |
| 135 | by (hyp_subst_tac 1); | |
| 136 | by (res_inst_tac [("s","UU"),("t","ya")] subst 1);
 | |
| 137 | by (etac (less_ssum3d RS iffD1 RS sym) 1); | |
| 138 | by (asm_simp_tac Ssum0_ss 1); | |
| 139 | by (hyp_subst_tac 1); | |
| 140 | by (asm_simp_tac Ssum0_ss 1); | |
| 141 | by (rtac monofun_cfun_arg 1); | |
| 142 | by (etac (less_ssum3b RS iffD1) 1); | |
| 143 | qed "monofun_Iwhen3"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 144 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 145 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 146 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 147 | (* some kind of exhaustion rules for chains in 'a ++ 'b *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 148 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 149 | |
| 9169 | 150 | Goal "[|~(!i.? x. Y(i::nat)=Isinl(x))|] ==> (? i.! x. Y(i)~=Isinl(x))"; | 
| 151 | by (fast_tac HOL_cs 1); | |
| 152 | qed "ssum_lemma1"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 153 | |
| 9169 | 154 | Goal "[|(? i.!x.(Y::nat => 'a++'b)(i::nat)~=Isinl(x::'a))|] \ | 
| 155 | \ ==> (? i y. (Y::nat => 'a++'b)(i::nat)=Isinr(y::'b) & y~=UU)"; | |
| 156 | by (etac exE 1); | |
| 157 | by (res_inst_tac [("p","Y(i)")] IssumE 1);
 | |
| 158 | by (dtac spec 1); | |
| 159 | by (contr_tac 1); | |
| 160 | by (dtac spec 1); | |
| 161 | by (contr_tac 1); | |
| 162 | by (fast_tac HOL_cs 1); | |
| 163 | qed "ssum_lemma2"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 164 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 165 | |
| 9169 | 166 | Goal "[|chain(Y);(? i x. Y(i)=Isinr(x::'b) & (x::'b)~=UU)|] \ | 
| 167 | \ ==> (!i.? y. Y(i)=Isinr(y))"; | |
| 168 | by (etac exE 1); | |
| 169 | by (etac exE 1); | |
| 170 | by (rtac allI 1); | |
| 171 | by (res_inst_tac [("p","Y(ia)")] IssumE 1);
 | |
| 172 | by (rtac exI 1); | |
| 173 | by (rtac trans 1); | |
| 174 | by (rtac strict_IsinlIsinr 2); | |
| 175 | by (atac 1); | |
| 176 | by (etac exI 2); | |
| 177 | by (etac conjE 1); | |
| 178 | by (res_inst_tac [("m","i"),("n","ia")] nat_less_cases 1);
 | |
| 179 | by (hyp_subst_tac 2); | |
| 180 | by (etac exI 2); | |
| 181 | by (eres_inst_tac [("P","x=UU")] notE 1);
 | |
| 182 | by (rtac (less_ssum3d RS iffD1) 1); | |
| 183 | by (eres_inst_tac [("s","Y(i)"),("t","Isinr(x)::'a++'b")] subst 1);
 | |
| 184 | by (eres_inst_tac [("s","Y(ia)"),("t","Isinl(xa)::'a++'b")] subst 1);
 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 185 | by (etac (chain_mono) 1); | 
| 9169 | 186 | by (atac 1); | 
| 187 | by (eres_inst_tac [("P","xa=UU")] notE 1);
 | |
| 188 | by (rtac (less_ssum3c RS iffD1) 1); | |
| 189 | by (eres_inst_tac [("s","Y(i)"),("t","Isinr(x)::'a++'b")] subst 1);
 | |
| 190 | by (eres_inst_tac [("s","Y(ia)"),("t","Isinl(xa)::'a++'b")] subst 1);
 | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 191 | by (etac (chain_mono) 1); | 
| 9169 | 192 | by (atac 1); | 
| 193 | qed "ssum_lemma3"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 194 | |
| 9169 | 195 | Goal "chain(Y) ==> (!i.? x. Y(i)=Isinl(x))|(!i.? y. Y(i)=Isinr(y))"; | 
| 196 | by (rtac case_split_thm 1); | |
| 197 | by (etac disjI1 1); | |
| 198 | by (rtac disjI2 1); | |
| 199 | by (etac ssum_lemma3 1); | |
| 200 | by (rtac ssum_lemma2 1); | |
| 201 | by (etac ssum_lemma1 1); | |
| 202 | qed "ssum_lemma4"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 203 | |
| 
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 | (* restricted surjectivity of Isinl *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 207 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 208 | |
| 9169 | 209 | Goal "z=Isinl(x)==> Isinl((Iwhen (LAM x. x) (LAM y. UU))(z)) = z"; | 
| 210 | by (hyp_subst_tac 1); | |
| 211 | by (case_tac "x=UU" 1); | |
| 212 | by (asm_simp_tac Ssum0_ss 1); | |
| 213 | by (asm_simp_tac Ssum0_ss 1); | |
| 214 | qed "ssum_lemma5"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 215 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 216 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 217 | (* restricted surjectivity of Isinr *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 218 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 219 | |
| 9169 | 220 | Goal "z=Isinr(x)==> Isinr((Iwhen (LAM y. UU) (LAM x. x))(z)) = z"; | 
| 221 | by (hyp_subst_tac 1); | |
| 222 | by (case_tac "x=UU" 1); | |
| 223 | by (asm_simp_tac Ssum0_ss 1); | |
| 224 | by (asm_simp_tac Ssum0_ss 1); | |
| 225 | qed "ssum_lemma6"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 226 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 227 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 228 | (* technical lemmas *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 229 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 230 | |
| 9169 | 231 | Goal "[|Isinl(x) << z; x~=UU|] ==> ? y. z=Isinl(y) & y~=UU"; | 
| 232 | by (res_inst_tac [("p","z")] IssumE 1);
 | |
| 233 | by (hyp_subst_tac 1); | |
| 234 | by (etac notE 1); | |
| 235 | by (rtac antisym_less 1); | |
| 236 | by (etac (less_ssum3a RS iffD1) 1); | |
| 237 | by (rtac minimal 1); | |
| 238 | by (fast_tac HOL_cs 1); | |
| 239 | by (hyp_subst_tac 1); | |
| 240 | by (rtac notE 1); | |
| 241 | by (etac (less_ssum3c RS iffD1) 2); | |
| 242 | by (atac 1); | |
| 243 | qed "ssum_lemma7"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 244 | |
| 9169 | 245 | Goal "[|Isinr(x) << z; x~=UU|] ==> ? y. z=Isinr(y) & y~=UU"; | 
| 246 | by (res_inst_tac [("p","z")] IssumE 1);
 | |
| 247 | by (hyp_subst_tac 1); | |
| 248 | by (etac notE 1); | |
| 249 | by (etac (less_ssum3d RS iffD1) 1); | |
| 250 | by (hyp_subst_tac 1); | |
| 251 | by (rtac notE 1); | |
| 252 | by (etac (less_ssum3d RS iffD1) 2); | |
| 253 | by (atac 1); | |
| 254 | by (fast_tac HOL_cs 1); | |
| 255 | qed "ssum_lemma8"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 256 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 257 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 258 | (* the type 'a ++ 'b is a cpo in three steps *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 259 | (* ------------------------------------------------------------------------ *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 260 | |
| 9169 | 261 | Goal "[|chain(Y);(!i.? x. Y(i)=Isinl(x))|] ==>\ | 
| 262 | \ range(Y) <<| Isinl(lub(range(%i.(Iwhen (LAM x. x) (LAM y. UU))(Y i))))"; | |
| 263 | by (rtac is_lubI 1); | |
| 264 | by (rtac ub_rangeI 1); | |
| 265 | by (etac allE 1); | |
| 266 | by (etac exE 1); | |
| 267 | by (res_inst_tac [("t","Y(i)")] (ssum_lemma5 RS subst) 1);
 | |
| 268 | by (atac 1); | |
| 269 | by (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1); | |
| 270 | by (rtac is_ub_thelub 1); | |
| 271 | by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1); | |
| 272 | by (strip_tac 1); | |
| 273 | by (res_inst_tac [("p","u")] IssumE2 1);
 | |
| 274 | by (res_inst_tac [("t","u")] (ssum_lemma5 RS subst) 1);
 | |
| 275 | by (atac 1); | |
| 276 | by (rtac (monofun_Isinl RS monofunE RS spec RS spec RS mp) 1); | |
| 277 | by (rtac is_lub_thelub 1); | |
| 278 | by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1); | |
| 279 | by (etac (monofun_Iwhen3 RS ub2ub_monofun) 1); | |
| 280 | by (hyp_subst_tac 1); | |
| 281 | by (rtac (less_ssum3c RS iffD2) 1); | |
| 282 | by (rtac chain_UU_I_inverse 1); | |
| 283 | by (rtac allI 1); | |
| 284 | by (res_inst_tac [("p","Y(i)")] IssumE 1);
 | |
| 285 | by (asm_simp_tac Ssum0_ss 1); | |
| 286 | by (asm_simp_tac Ssum0_ss 2); | |
| 287 | by (etac notE 1); | |
| 288 | by (rtac (less_ssum3c RS iffD1) 1); | |
| 289 | by (res_inst_tac [("t","Isinl(x)")] subst 1);
 | |
| 290 | by (atac 1); | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 291 | by (etac (ub_rangeD) 1); | 
| 9169 | 292 | qed "lub_ssum1a"; | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 293 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 294 | |
| 9169 | 295 | Goal "[|chain(Y);(!i.? x. Y(i)=Isinr(x))|] ==>\ | 
| 296 | \ range(Y) <<| Isinr(lub(range(%i.(Iwhen (LAM y. UU) (LAM x. x))(Y i))))"; | |
| 297 | by (rtac is_lubI 1); | |
| 298 | by (rtac ub_rangeI 1); | |
| 299 | by (etac allE 1); | |
| 300 | by (etac exE 1); | |
| 301 | by (res_inst_tac [("t","Y(i)")] (ssum_lemma6 RS subst) 1);
 | |
| 302 | by (atac 1); | |
| 303 | by (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1); | |
| 304 | by (rtac is_ub_thelub 1); | |
| 305 | by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1); | |
| 306 | by (strip_tac 1); | |
| 307 | by (res_inst_tac [("p","u")] IssumE2 1);
 | |
| 308 | by (hyp_subst_tac 1); | |
| 309 | by (rtac (less_ssum3d RS iffD2) 1); | |
| 310 | by (rtac chain_UU_I_inverse 1); | |
| 311 | by (rtac allI 1); | |
| 312 | by (res_inst_tac [("p","Y(i)")] IssumE 1);
 | |
| 313 | by (asm_simp_tac Ssum0_ss 1); | |
| 314 | by (asm_simp_tac Ssum0_ss 1); | |
| 315 | by (etac notE 1); | |
| 316 | by (rtac (less_ssum3d RS iffD1) 1); | |
| 317 | by (res_inst_tac [("t","Isinr(y)")] subst 1);
 | |
| 318 | by (atac 1); | |
| 9248 
e1dee89de037
massive tidy-up: goal -> Goal, remove use of prems, etc.
 paulson parents: 
9245diff
changeset | 319 | by (etac (ub_rangeD) 1); | 
| 9169 | 320 | by (res_inst_tac [("t","u")] (ssum_lemma6 RS subst) 1);
 | 
| 321 | by (atac 1); | |
| 322 | by (rtac (monofun_Isinr RS monofunE RS spec RS spec RS mp) 1); | |
| 323 | by (rtac is_lub_thelub 1); | |
| 324 | by (etac (monofun_Iwhen3 RS ch2ch_monofun) 1); | |
| 325 | by (etac (monofun_Iwhen3 RS ub2ub_monofun) 1); | |
| 326 | qed "lub_ssum1b"; | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 327 | |
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 328 | |
| 1779 | 329 | bind_thm ("thelub_ssum1a", lub_ssum1a RS thelubI);
 | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 330 | (* | 
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
4098diff
changeset | 331 | [| chain ?Y1; ! i. ? x. ?Y1 i = Isinl x |] ==> | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 332 | lub (range ?Y1) = Isinl | 
| 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 333 | (lub (range (%i. Iwhen (LAM x. x) (LAM y. UU) (?Y1 i)))) | 
| 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 334 | *) | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 335 | |
| 1779 | 336 | bind_thm ("thelub_ssum1b", lub_ssum1b RS thelubI);
 | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 337 | (* | 
| 4721 
c8a8482a8124
renamed is_chain to chain, is_tord to tord, replaced chain_finite by chfin
 oheimb parents: 
4098diff
changeset | 338 | [| chain ?Y1; ! i. ? x. ?Y1 i = Isinr x |] ==> | 
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 339 | lub (range ?Y1) = Isinr | 
| 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 340 | (lub (range (%i. Iwhen (LAM y. UU) (LAM x. x) (?Y1 i)))) | 
| 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 341 | *) | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 342 | |
| 9169 | 343 | Goal "chain(Y::nat=>'a ++'b) ==> ? x. range(Y) <<|x"; | 
| 344 | by (rtac (ssum_lemma4 RS disjE) 1); | |
| 345 | by (atac 1); | |
| 346 | by (rtac exI 1); | |
| 347 | by (etac lub_ssum1a 1); | |
| 348 | by (atac 1); | |
| 349 | by (rtac exI 1); | |
| 350 | by (etac lub_ssum1b 1); | |
| 351 | by (atac 1); | |
| 352 | qed "cpo_ssum"; | |
| 1168 
74be52691d62
The curried version of HOLCF is now just called HOLCF. The old
 regensbu parents: 
961diff
changeset | 353 |