| author | wenzelm | 
| Wed, 16 Apr 1997 18:53:36 +0200 | |
| changeset 2967 | 89db5eedecab | 
| parent 2950 | 5d2e0865ecf3 | 
| child 2985 | 824e18a114c9 | 
| permissions | -rw-r--r-- | 
| 1465 | 1 | (* Title: HOL/llist | 
| 969 | 2 | ID: $Id$ | 
| 1465 | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 969 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | SHOULD LListD_Fun_CONS_I, etc., be equations (for rewriting)? | |
| 7 | *) | |
| 8 | ||
| 9 | open LList; | |
| 10 | ||
| 11 | (** Simplification **) | |
| 12 | ||
| 1266 | 13 | simpset := !simpset setloop split_tac [expand_split, expand_sum_case]; | 
| 969 | 14 | |
| 15 | (*For adding _eqI rules to a simpset; we must remove Pair_eq because | |
| 16 | it may turn an instance of reflexivity into a conjunction!*) | |
| 17 | fun add_eqI ss = ss addsimps [range_eqI, image_eqI] | |
| 18 | delsimps [Pair_eq]; | |
| 19 | ||
| 20 | ||
| 21 | (*This justifies using llist in other recursive type definitions*) | |
| 22 | goalw LList.thy llist.defs "!!A B. A<=B ==> llist(A) <= llist(B)"; | |
| 23 | by (rtac gfp_mono 1); | |
| 24 | by (REPEAT (ares_tac basic_monos 1)); | |
| 25 | qed "llist_mono"; | |
| 26 | ||
| 27 | ||
| 28 | goal LList.thy "llist(A) = {Numb(0)} <+> (A <*> llist(A))";
 | |
| 29 | let val rew = rewrite_rule [NIL_def, CONS_def] in | |
| 1820 | 30 | by (fast_tac (!claset addSIs (equalityI :: map rew llist.intrs) | 
| 969 | 31 | addEs [rew llist.elim]) 1) | 
| 32 | end; | |
| 33 | qed "llist_unfold"; | |
| 34 | ||
| 35 | ||
| 36 | (*** Type checking by coinduction, using list_Fun | |
| 37 | THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! | |
| 38 | ***) | |
| 39 | ||
| 40 | goalw LList.thy [list_Fun_def] | |
| 41 | "!!M. [| M : X; X <= list_Fun A (X Un llist(A)) |] ==> M : llist(A)"; | |
| 1046 | 42 | by (etac llist.coinduct 1); | 
| 43 | by (etac (subsetD RS CollectD) 1); | |
| 44 | by (assume_tac 1); | |
| 969 | 45 | qed "llist_coinduct"; | 
| 46 | ||
| 47 | goalw LList.thy [list_Fun_def, NIL_def] "NIL: list_Fun A X"; | |
| 1820 | 48 | by (Fast_tac 1); | 
| 969 | 49 | qed "list_Fun_NIL_I"; | 
| 50 | ||
| 51 | goalw LList.thy [list_Fun_def,CONS_def] | |
| 52 | "!!M N. [| M: A; N: X |] ==> CONS M N : list_Fun A X"; | |
| 1820 | 53 | by (Fast_tac 1); | 
| 969 | 54 | qed "list_Fun_CONS_I"; | 
| 55 | ||
| 56 | (*Utilise the "strong" part, i.e. gfp(f)*) | |
| 57 | goalw LList.thy (llist.defs @ [list_Fun_def]) | |
| 58 | "!!M N. M: llist(A) ==> M : list_Fun A (X Un llist(A))"; | |
| 59 | by (etac (llist.mono RS gfp_fun_UnI2) 1); | |
| 60 | qed "list_Fun_llist_I"; | |
| 61 | ||
| 62 | (*** LList_corec satisfies the desired recurion equation ***) | |
| 63 | ||
| 64 | (*A continuity result?*) | |
| 65 | goalw LList.thy [CONS_def] "CONS M (UN x.f(x)) = (UN x. CONS M (f x))"; | |
| 1266 | 66 | by (simp_tac (!simpset addsimps [In1_UN1, Scons_UN1_y]) 1); | 
| 969 | 67 | qed "CONS_UN1"; | 
| 68 | ||
| 69 | (*UNUSED; obsolete? | |
| 70 | goal Prod.thy "split p (%x y.UN z.f x y z) = (UN z. split p (%x y.f x y z))"; | |
| 1266 | 71 | by (simp_tac (!simpset setloop (split_tac [expand_split])) 1); | 
| 969 | 72 | qed "split_UN1"; | 
| 73 | ||
| 74 | goal Sum.thy "sum_case s f (%y.UN z.g y z) = (UN z.sum_case s f (%y.g y z))"; | |
| 1266 | 75 | by (simp_tac (!simpset setloop (split_tac [expand_sum_case])) 1); | 
| 969 | 76 | qed "sum_case2_UN1"; | 
| 77 | *) | |
| 78 | ||
| 79 | val prems = goalw LList.thy [CONS_def] | |
| 80 | "[| M<=M'; N<=N' |] ==> CONS M N <= CONS M' N'"; | |
| 81 | by (REPEAT (resolve_tac ([In1_mono,Scons_mono]@prems) 1)); | |
| 82 | qed "CONS_mono"; | |
| 83 | ||
| 1266 | 84 | Addsimps [LList_corec_fun_def RS def_nat_rec_0, | 
| 1465 | 85 | LList_corec_fun_def RS def_nat_rec_Suc]; | 
| 969 | 86 | |
| 87 | (** The directions of the equality are proved separately **) | |
| 88 | ||
| 89 | goalw LList.thy [LList_corec_def] | |
| 90 | "LList_corec a f <= sum_case (%u.NIL) \ | |
| 1465 | 91 | \ (split(%z w. CONS z (LList_corec w f))) (f a)"; | 
| 969 | 92 | by (rtac UN1_least 1); | 
| 93 | by (res_inst_tac [("n","k")] natE 1);
 | |
| 1266 | 94 | by (ALLGOALS (Asm_simp_tac)); | 
| 969 | 95 | by (REPEAT (resolve_tac [allI, impI, subset_refl RS CONS_mono, UN1_upper] 1)); | 
| 96 | qed "LList_corec_subset1"; | |
| 97 | ||
| 98 | goalw LList.thy [LList_corec_def] | |
| 99 | "sum_case (%u.NIL) (split(%z w. CONS z (LList_corec w f))) (f a) <= \ | |
| 100 | \ LList_corec a f"; | |
| 1266 | 101 | by (simp_tac (!simpset addsimps [CONS_UN1]) 1); | 
| 1820 | 102 | by (safe_tac (!claset)); | 
| 1266 | 103 | by (ALLGOALS (res_inst_tac [("x","Suc(?k)")] UN1_I THEN' Asm_simp_tac));
 | 
| 969 | 104 | qed "LList_corec_subset2"; | 
| 105 | ||
| 106 | (*the recursion equation for LList_corec -- NOT SUITABLE FOR REWRITING!*) | |
| 107 | goal LList.thy | |
| 108 | "LList_corec a f = sum_case (%u. NIL) \ | |
| 1465 | 109 | \ (split(%z w. CONS z (LList_corec w f))) (f a)"; | 
| 969 | 110 | by (REPEAT (resolve_tac [equalityI, LList_corec_subset1, | 
| 1465 | 111 | LList_corec_subset2] 1)); | 
| 969 | 112 | qed "LList_corec"; | 
| 113 | ||
| 114 | (*definitional version of same*) | |
| 115 | val [rew] = goal LList.thy | |
| 1465 | 116 | "[| !!x. h(x) == LList_corec x f |] ==> \ | 
| 969 | 117 | \ h(a) = sum_case (%u.NIL) (split(%z w. CONS z (h w))) (f a)"; | 
| 118 | by (rewtac rew); | |
| 119 | by (rtac LList_corec 1); | |
| 120 | qed "def_LList_corec"; | |
| 121 | ||
| 122 | (*A typical use of co-induction to show membership in the gfp. | |
| 123 | Bisimulation is range(%x. LList_corec x f) *) | |
| 124 | goal LList.thy "LList_corec a f : llist({u.True})";
 | |
| 125 | by (res_inst_tac [("X", "range(%x.LList_corec x ?g)")] llist_coinduct 1);
 | |
| 126 | by (rtac rangeI 1); | |
| 1820 | 127 | by (safe_tac (!claset)); | 
| 969 | 128 | by (stac LList_corec 1); | 
| 1266 | 129 | by (simp_tac (!simpset addsimps [list_Fun_NIL_I, list_Fun_CONS_I, CollectI] | 
| 969 | 130 | |> add_eqI) 1); | 
| 131 | qed "LList_corec_type"; | |
| 132 | ||
| 133 | (*Lemma for the proof of llist_corec*) | |
| 134 | goal LList.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 135 | "LList_corec a (%z.sum_case Inl (split(%v w.Inr((Leaf(v),w)))) (f z)) : \ | 
| 1642 | 136 | \ llist(range Leaf)"; | 
| 969 | 137 | by (res_inst_tac [("X", "range(%x.LList_corec x ?g)")] llist_coinduct 1);
 | 
| 138 | by (rtac rangeI 1); | |
| 1820 | 139 | by (safe_tac (!claset)); | 
| 969 | 140 | by (stac LList_corec 1); | 
| 1266 | 141 | by (asm_simp_tac (!simpset addsimps [list_Fun_NIL_I]) 1); | 
| 1820 | 142 | by (fast_tac (!claset addSIs [list_Fun_CONS_I]) 1); | 
| 969 | 143 | qed "LList_corec_type2"; | 
| 144 | ||
| 145 | ||
| 146 | (**** llist equality as a gfp; the bisimulation principle ****) | |
| 147 | ||
| 148 | (*This theorem is actually used, unlike the many similar ones in ZF*) | |
| 149 | goal LList.thy "LListD(r) = diag({Numb(0)}) <++> (r <**> LListD(r))";
 | |
| 150 | let val rew = rewrite_rule [NIL_def, CONS_def] in | |
| 1820 | 151 | by (fast_tac (!claset addSIs (equalityI :: map rew LListD.intrs) | 
| 969 | 152 | addEs [rew LListD.elim]) 1) | 
| 153 | end; | |
| 154 | qed "LListD_unfold"; | |
| 155 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 156 | goal LList.thy "!M N. (M,N) : LListD(diag(A)) --> ntrunc k M = ntrunc k N"; | 
| 969 | 157 | by (res_inst_tac [("n", "k")] less_induct 1);
 | 
| 1820 | 158 | by (safe_tac ((claset_of "Fun") delrules [equalityI])); | 
| 969 | 159 | by (etac LListD.elim 1); | 
| 1820 | 160 | by (safe_tac (((claset_of "Prod") delrules [equalityI]) addSEs [diagE])); | 
| 969 | 161 | by (res_inst_tac [("n", "n")] natE 1);
 | 
| 1266 | 162 | by (asm_simp_tac (!simpset addsimps [ntrunc_0]) 1); | 
| 969 | 163 | by (rename_tac "n'" 1); | 
| 164 | by (res_inst_tac [("n", "n'")] natE 1);
 | |
| 1266 | 165 | by (asm_simp_tac (!simpset addsimps [CONS_def, ntrunc_one_In1]) 1); | 
| 1673 
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
 oheimb parents: 
1642diff
changeset | 166 | by (asm_simp_tac (!simpset addsimps [CONS_def, ntrunc_In1, ntrunc_Scons, less_Suc_eq]) 1); | 
| 969 | 167 | qed "LListD_implies_ntrunc_equality"; | 
| 168 | ||
| 169 | (*The domain of the LListD relation*) | |
| 170 | goalw LList.thy (llist.defs @ [NIL_def, CONS_def]) | |
| 171 | "fst``LListD(diag(A)) <= llist(A)"; | |
| 172 | by (rtac gfp_upperbound 1); | |
| 173 | (*avoids unfolding LListD on the rhs*) | |
| 174 | by (res_inst_tac [("P", "%x. fst``x <= ?B")] (LListD_unfold RS ssubst) 1);
 | |
| 1266 | 175 | by (Simp_tac 1); | 
| 1820 | 176 | by (Fast_tac 1); | 
| 969 | 177 | qed "fst_image_LListD"; | 
| 178 | ||
| 179 | (*This inclusion justifies the use of coinduction to show M=N*) | |
| 180 | goal LList.thy "LListD(diag(A)) <= diag(llist(A))"; | |
| 181 | by (rtac subsetI 1); | |
| 182 | by (res_inst_tac [("p","x")] PairE 1);
 | |
| 1820 | 183 | by (safe_tac (!claset)); | 
| 969 | 184 | by (rtac diag_eqI 1); | 
| 185 | by (rtac (LListD_implies_ntrunc_equality RS spec RS spec RS mp RS | |
| 1465 | 186 | ntrunc_equality) 1); | 
| 969 | 187 | by (assume_tac 1); | 
| 188 | by (etac (fst_imageI RS (fst_image_LListD RS subsetD)) 1); | |
| 189 | qed "LListD_subset_diag"; | |
| 190 | ||
| 191 | (** Coinduction, using LListD_Fun | |
| 192 | THE COINDUCTIVE DEFINITION PACKAGE COULD DO THIS! | |
| 193 | **) | |
| 194 | ||
| 195 | goalw LList.thy [LListD_Fun_def] | |
| 196 | "!!M. [| M : X; X <= LListD_Fun r (X Un LListD(r)) |] ==> M : LListD(r)"; | |
| 1046 | 197 | by (etac LListD.coinduct 1); | 
| 198 | by (etac (subsetD RS CollectD) 1); | |
| 199 | by (assume_tac 1); | |
| 969 | 200 | qed "LListD_coinduct"; | 
| 201 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 202 | goalw LList.thy [LListD_Fun_def,NIL_def] "(NIL,NIL) : LListD_Fun r s"; | 
| 1820 | 203 | by (Fast_tac 1); | 
| 969 | 204 | qed "LListD_Fun_NIL_I"; | 
| 205 | ||
| 206 | goalw LList.thy [LListD_Fun_def,CONS_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 207 | "!!x. [| x:A; (M,N):s |] ==> (CONS x M, CONS x N) : LListD_Fun (diag A) s"; | 
| 1820 | 208 | by (Fast_tac 1); | 
| 969 | 209 | qed "LListD_Fun_CONS_I"; | 
| 210 | ||
| 211 | (*Utilise the "strong" part, i.e. gfp(f)*) | |
| 212 | goalw LList.thy (LListD.defs @ [LListD_Fun_def]) | |
| 213 | "!!M N. M: LListD(r) ==> M : LListD_Fun r (X Un LListD(r))"; | |
| 214 | by (etac (LListD.mono RS gfp_fun_UnI2) 1); | |
| 215 | qed "LListD_Fun_LListD_I"; | |
| 216 | ||
| 217 | ||
| 218 | (*This converse inclusion helps to strengthen LList_equalityI*) | |
| 219 | goal LList.thy "diag(llist(A)) <= LListD(diag(A))"; | |
| 220 | by (rtac subsetI 1); | |
| 221 | by (etac LListD_coinduct 1); | |
| 222 | by (rtac subsetI 1); | |
| 1046 | 223 | by (etac diagE 1); | 
| 224 | by (etac ssubst 1); | |
| 969 | 225 | by (eresolve_tac [llist.elim] 1); | 
| 226 | by (ALLGOALS | |
| 1266 | 227 | (asm_simp_tac (!simpset addsimps [diagI, LListD_Fun_NIL_I, | 
| 1465 | 228 | LListD_Fun_CONS_I]))); | 
| 969 | 229 | qed "diag_subset_LListD"; | 
| 230 | ||
| 231 | goal LList.thy "LListD(diag(A)) = diag(llist(A))"; | |
| 232 | by (REPEAT (resolve_tac [equalityI, LListD_subset_diag, | |
| 1465 | 233 | diag_subset_LListD] 1)); | 
| 969 | 234 | qed "LListD_eq_diag"; | 
| 235 | ||
| 236 | goal LList.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 237 | "!!M N. M: llist(A) ==> (M,M) : LListD_Fun (diag A) (X Un diag(llist(A)))"; | 
| 969 | 238 | by (rtac (LListD_eq_diag RS subst) 1); | 
| 1046 | 239 | by (rtac LListD_Fun_LListD_I 1); | 
| 1266 | 240 | by (asm_simp_tac (!simpset addsimps [LListD_eq_diag, diagI]) 1); | 
| 969 | 241 | qed "LListD_Fun_diag_I"; | 
| 242 | ||
| 243 | ||
| 244 | (** To show two LLists are equal, exhibit a bisimulation! | |
| 245 | [also admits true equality] | |
| 246 |    Replace "A" by some particular set, like {x.True}??? *)
 | |
| 247 | goal LList.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 248 | "!!r. [| (M,N) : r; r <= LListD_Fun (diag A) (r Un diag(llist(A))) \ | 
| 969 | 249 | \ |] ==> M=N"; | 
| 250 | by (rtac (LListD_subset_diag RS subsetD RS diagE) 1); | |
| 251 | by (etac LListD_coinduct 1); | |
| 1266 | 252 | by (asm_simp_tac (!simpset addsimps [LListD_eq_diag]) 1); | 
| 1820 | 253 | by (safe_tac (!claset)); | 
| 969 | 254 | qed "LList_equalityI"; | 
| 255 | ||
| 256 | ||
| 257 | (*** Finality of llist(A): Uniqueness of functions defined by corecursion ***) | |
| 258 | ||
| 259 | (*abstract proof using a bisimulation*) | |
| 260 | val [prem1,prem2] = goal LList.thy | |
| 261 | "[| !!x. h1(x) = sum_case (%u.NIL) (split(%z w. CONS z (h1 w))) (f x); \ | |
| 262 | \ !!x. h2(x) = sum_case (%u.NIL) (split(%z w. CONS z (h2 w))) (f x) |]\ | |
| 263 | \ ==> h1=h2"; | |
| 264 | by (rtac ext 1); | |
| 265 | (*next step avoids an unknown (and flexflex pair) in simplification*) | |
| 266 | by (res_inst_tac [("A", "{u.True}"),
 | |
| 1465 | 267 |                   ("r", "range(%u. (h1(u),h2(u)))")] LList_equalityI 1);
 | 
| 969 | 268 | by (rtac rangeI 1); | 
| 1820 | 269 | by (safe_tac (!claset)); | 
| 969 | 270 | by (stac prem1 1); | 
| 271 | by (stac prem2 1); | |
| 1266 | 272 | by (simp_tac (!simpset addsimps [LListD_Fun_NIL_I, | 
| 1465 | 273 | CollectI RS LListD_Fun_CONS_I] | 
| 274 | |> add_eqI) 1); | |
| 969 | 275 | qed "LList_corec_unique"; | 
| 276 | ||
| 277 | val [prem] = goal LList.thy | |
| 278 | "[| !!x. h(x) = sum_case (%u.NIL) (split(%z w. CONS z (h w))) (f x) |] \ | |
| 279 | \ ==> h = (%x.LList_corec x f)"; | |
| 280 | by (rtac (LList_corec RS (prem RS LList_corec_unique)) 1); | |
| 281 | qed "equals_LList_corec"; | |
| 282 | ||
| 283 | ||
| 284 | (** Obsolete LList_corec_unique proof: complete induction, not coinduction **) | |
| 285 | ||
| 286 | goalw LList.thy [CONS_def] "ntrunc (Suc 0) (CONS M N) = {}";
 | |
| 287 | by (rtac ntrunc_one_In1 1); | |
| 288 | qed "ntrunc_one_CONS"; | |
| 289 | ||
| 290 | goalw LList.thy [CONS_def] | |
| 291 | "ntrunc (Suc(Suc(k))) (CONS M N) = CONS (ntrunc k M) (ntrunc k N)"; | |
| 1266 | 292 | by (simp_tac (!simpset addsimps [ntrunc_Scons,ntrunc_In1]) 1); | 
| 969 | 293 | qed "ntrunc_CONS"; | 
| 294 | ||
| 295 | val [prem1,prem2] = goal LList.thy | |
| 296 | "[| !!x. h1(x) = sum_case (%u.NIL) (split(%z w. CONS z (h1 w))) (f x); \ | |
| 297 | \ !!x. h2(x) = sum_case (%u.NIL) (split(%z w. CONS z (h2 w))) (f x) |]\ | |
| 298 | \ ==> h1=h2"; | |
| 299 | by (rtac (ntrunc_equality RS ext) 1); | |
| 2596 
3b4ad6c7726f
Modified proofs because of added "triv_forall_equality".
 nipkow parents: 
2031diff
changeset | 300 | by (rename_tac "x k" 1); | 
| 969 | 301 | by (res_inst_tac [("x", "x")] spec 1);
 | 
| 302 | by (res_inst_tac [("n", "k")] less_induct 1);
 | |
| 2596 
3b4ad6c7726f
Modified proofs because of added "triv_forall_equality".
 nipkow parents: 
2031diff
changeset | 303 | by (rename_tac "n" 1); | 
| 969 | 304 | by (rtac allI 1); | 
| 2596 
3b4ad6c7726f
Modified proofs because of added "triv_forall_equality".
 nipkow parents: 
2031diff
changeset | 305 | by (rename_tac "y" 1); | 
| 969 | 306 | by (stac prem1 1); | 
| 307 | by (stac prem2 1); | |
| 1303 
010be89a7541
Hat to modify a proof because of chaned simpset for Prod.
 nipkow parents: 
1266diff
changeset | 308 | by (simp_tac (!simpset setloop (split_tac [expand_sum_case])) 1); | 
| 969 | 309 | by (strip_tac 1); | 
| 310 | by (res_inst_tac [("n", "n")] natE 1);
 | |
| 2596 
3b4ad6c7726f
Modified proofs because of added "triv_forall_equality".
 nipkow parents: 
2031diff
changeset | 311 | by (rename_tac "m" 2); | 
| 
3b4ad6c7726f
Modified proofs because of added "triv_forall_equality".
 nipkow parents: 
2031diff
changeset | 312 | by (res_inst_tac [("n", "m")] natE 2);
 | 
| 1266 | 313 | by (ALLGOALS(asm_simp_tac(!simpset addsimps | 
| 1673 
d22110ddd0af
repaired critical proofs depending on the order inside non-confluent SimpSets
 oheimb parents: 
1642diff
changeset | 314 | [ntrunc_0,ntrunc_one_CONS,ntrunc_CONS, less_Suc_eq]))); | 
| 969 | 315 | result(); | 
| 316 | ||
| 317 | ||
| 318 | (*** Lconst -- defined directly using lfp, but equivalent to a LList_corec ***) | |
| 319 | ||
| 320 | goal LList.thy "mono(CONS(M))"; | |
| 321 | by (REPEAT (ares_tac [monoI, subset_refl, CONS_mono] 1)); | |
| 322 | qed "Lconst_fun_mono"; | |
| 323 | ||
| 324 | (* Lconst(M) = CONS M (Lconst M) *) | |
| 325 | bind_thm ("Lconst", (Lconst_fun_mono RS (Lconst_def RS def_lfp_Tarski)));
 | |
| 326 | ||
| 327 | (*A typical use of co-induction to show membership in the gfp. | |
| 328 |   The containing set is simply the singleton {Lconst(M)}. *)
 | |
| 329 | goal LList.thy "!!M A. M:A ==> Lconst(M): llist(A)"; | |
| 330 | by (rtac (singletonI RS llist_coinduct) 1); | |
| 1820 | 331 | by (safe_tac (!claset)); | 
| 969 | 332 | by (res_inst_tac [("P", "%u. u: ?A")] (Lconst RS ssubst) 1);
 | 
| 333 | by (REPEAT (ares_tac [list_Fun_CONS_I, singletonI, UnI1] 1)); | |
| 334 | qed "Lconst_type"; | |
| 335 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 336 | goal LList.thy "Lconst(M) = LList_corec M (%x.Inr((x,x)))"; | 
| 969 | 337 | by (rtac (equals_LList_corec RS fun_cong) 1); | 
| 1266 | 338 | by (Simp_tac 1); | 
| 969 | 339 | by (rtac Lconst 1); | 
| 340 | qed "Lconst_eq_LList_corec"; | |
| 341 | ||
| 342 | (*Thus we could have used gfp in the definition of Lconst*) | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 343 | goal LList.thy "gfp(%N. CONS M N) = LList_corec M (%x.Inr((x,x)))"; | 
| 969 | 344 | by (rtac (equals_LList_corec RS fun_cong) 1); | 
| 1266 | 345 | by (Simp_tac 1); | 
| 969 | 346 | by (rtac (Lconst_fun_mono RS gfp_Tarski) 1); | 
| 347 | qed "gfp_Lconst_eq_LList_corec"; | |
| 348 | ||
| 349 | ||
| 350 | (*** Isomorphisms ***) | |
| 351 | ||
| 352 | goal LList.thy "inj(Rep_llist)"; | |
| 353 | by (rtac inj_inverseI 1); | |
| 354 | by (rtac Rep_llist_inverse 1); | |
| 355 | qed "inj_Rep_llist"; | |
| 356 | ||
| 1642 | 357 | goal LList.thy "inj_onto Abs_llist (llist(range Leaf))"; | 
| 969 | 358 | by (rtac inj_onto_inverseI 1); | 
| 359 | by (etac Abs_llist_inverse 1); | |
| 360 | qed "inj_onto_Abs_llist"; | |
| 361 | ||
| 362 | (** Distinctness of constructors **) | |
| 363 | ||
| 364 | goalw LList.thy [LNil_def,LCons_def] "~ LCons x xs = LNil"; | |
| 365 | by (rtac (CONS_not_NIL RS (inj_onto_Abs_llist RS inj_onto_contraD)) 1); | |
| 366 | by (REPEAT (resolve_tac (llist.intrs @ [rangeI, Rep_llist]) 1)); | |
| 367 | qed "LCons_not_LNil"; | |
| 368 | ||
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1824diff
changeset | 369 | bind_thm ("LNil_not_LCons", LCons_not_LNil RS not_sym);
 | 
| 969 | 370 | |
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1824diff
changeset | 371 | AddIffs [LCons_not_LNil, LNil_not_LCons]; | 
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1824diff
changeset | 372 | |
| 969 | 373 | |
| 374 | (** llist constructors **) | |
| 375 | ||
| 376 | goalw LList.thy [LNil_def] | |
| 377 | "Rep_llist(LNil) = NIL"; | |
| 378 | by (rtac (llist.NIL_I RS Abs_llist_inverse) 1); | |
| 379 | qed "Rep_llist_LNil"; | |
| 380 | ||
| 381 | goalw LList.thy [LCons_def] | |
| 382 | "Rep_llist(LCons x l) = CONS (Leaf x) (Rep_llist l)"; | |
| 383 | by (REPEAT (resolve_tac [llist.CONS_I RS Abs_llist_inverse, | |
| 1465 | 384 | rangeI, Rep_llist] 1)); | 
| 969 | 385 | qed "Rep_llist_LCons"; | 
| 386 | ||
| 387 | (** Injectiveness of CONS and LCons **) | |
| 388 | ||
| 389 | goalw LList.thy [CONS_def] "(CONS M N=CONS M' N') = (M=M' & N=N')"; | |
| 1820 | 390 | by (fast_tac (!claset addSEs [Scons_inject, make_elim In1_inject]) 1); | 
| 1266 | 391 | qed "CONS_CONS_eq2"; | 
| 969 | 392 | |
| 393 | bind_thm ("CONS_inject", (CONS_CONS_eq RS iffD1 RS conjE));
 | |
| 394 | ||
| 395 | ||
| 396 | (*For reasoning about abstract llist constructors*) | |
| 1820 | 397 | |
| 398 | AddIs ([Rep_llist]@llist.intrs); | |
| 399 | AddSDs [inj_onto_Abs_llist RS inj_ontoD, | |
| 2031 | 400 | inj_Rep_llist RS injD, Leaf_inject]; | 
| 969 | 401 | |
| 402 | goalw LList.thy [LCons_def] "(LCons x xs=LCons y ys) = (x=y & xs=ys)"; | |
| 1820 | 403 | by (Fast_tac 1); | 
| 969 | 404 | qed "LCons_LCons_eq"; | 
| 1985 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1824diff
changeset | 405 | |
| 
84cf16192e03
Tidied many proofs, using AddIffs to let equivalences take
 paulson parents: 
1824diff
changeset | 406 | AddIffs [LCons_LCons_eq]; | 
| 969 | 407 | |
| 408 | val [major] = goal LList.thy "CONS M N: llist(A) ==> M: A & N: llist(A)"; | |
| 409 | by (rtac (major RS llist.elim) 1); | |
| 410 | by (etac CONS_neq_NIL 1); | |
| 1820 | 411 | by (Fast_tac 1); | 
| 1266 | 412 | qed "CONS_D2"; | 
| 969 | 413 | |
| 414 | ||
| 415 | (****** Reasoning about llist(A) ******) | |
| 416 | ||
| 1266 | 417 | Addsimps [List_case_NIL, List_case_CONS]; | 
| 969 | 418 | |
| 419 | (*A special case of list_equality for functions over lazy lists*) | |
| 420 | val [Mlist,gMlist,NILcase,CONScase] = goal LList.thy | |
| 1465 | 421 | "[| M: llist(A); g(NIL): llist(A); \ | 
| 422 | \ f(NIL)=g(NIL); \ | |
| 423 | \ !!x l. [| x:A; l: llist(A) |] ==> \ | |
| 424 | \ (f(CONS x l),g(CONS x l)) : \ | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 425 | \ LListD_Fun (diag A) ((%u.(f(u),g(u)))``llist(A) Un \ | 
| 1465 | 426 | \ diag(llist(A))) \ | 
| 969 | 427 | \ |] ==> f(M) = g(M)"; | 
| 428 | by (rtac LList_equalityI 1); | |
| 1046 | 429 | by (rtac (Mlist RS imageI) 1); | 
| 969 | 430 | by (rtac subsetI 1); | 
| 431 | by (etac imageE 1); | |
| 432 | by (etac ssubst 1); | |
| 433 | by (etac llist.elim 1); | |
| 434 | by (etac ssubst 1); | |
| 435 | by (stac NILcase 1); | |
| 1046 | 436 | by (rtac (gMlist RS LListD_Fun_diag_I) 1); | 
| 969 | 437 | by (etac ssubst 1); | 
| 438 | by (REPEAT (ares_tac [CONScase] 1)); | |
| 439 | qed "LList_fun_equalityI"; | |
| 440 | ||
| 441 | ||
| 442 | (*** The functional "Lmap" ***) | |
| 443 | ||
| 444 | goal LList.thy "Lmap f NIL = NIL"; | |
| 445 | by (rtac (Lmap_def RS def_LList_corec RS trans) 1); | |
| 1266 | 446 | by (Simp_tac 1); | 
| 969 | 447 | qed "Lmap_NIL"; | 
| 448 | ||
| 449 | goal LList.thy "Lmap f (CONS M N) = CONS (f M) (Lmap f N)"; | |
| 450 | by (rtac (Lmap_def RS def_LList_corec RS trans) 1); | |
| 1266 | 451 | by (Simp_tac 1); | 
| 969 | 452 | qed "Lmap_CONS"; | 
| 453 | ||
| 454 | (*Another type-checking proof by coinduction*) | |
| 455 | val [major,minor] = goal LList.thy | |
| 456 | "[| M: llist(A); !!x. x:A ==> f(x):B |] ==> Lmap f M: llist(B)"; | |
| 457 | by (rtac (major RS imageI RS llist_coinduct) 1); | |
| 1820 | 458 | by (safe_tac (!claset)); | 
| 969 | 459 | by (etac llist.elim 1); | 
| 1266 | 460 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [Lmap_NIL,Lmap_CONS]))); | 
| 969 | 461 | by (REPEAT (ares_tac [list_Fun_NIL_I, list_Fun_CONS_I, | 
| 1465 | 462 | minor, imageI, UnI1] 1)); | 
| 969 | 463 | qed "Lmap_type"; | 
| 464 | ||
| 465 | (*This type checking rule synthesises a sufficiently large set for f*) | |
| 466 | val [major] = goal LList.thy "M: llist(A) ==> Lmap f M: llist(f``A)"; | |
| 467 | by (rtac (major RS Lmap_type) 1); | |
| 468 | by (etac imageI 1); | |
| 469 | qed "Lmap_type2"; | |
| 470 | ||
| 471 | (** Two easy results about Lmap **) | |
| 472 | ||
| 473 | val [prem] = goalw LList.thy [o_def] | |
| 474 | "M: llist(A) ==> Lmap (f o g) M = Lmap f (Lmap g M)"; | |
| 475 | by (rtac (prem RS imageI RS LList_equalityI) 1); | |
| 1820 | 476 | by (safe_tac (!claset)); | 
| 969 | 477 | by (etac llist.elim 1); | 
| 1266 | 478 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [Lmap_NIL,Lmap_CONS]))); | 
| 969 | 479 | by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI, UnI1, | 
| 1465 | 480 | rangeI RS LListD_Fun_CONS_I] 1)); | 
| 969 | 481 | qed "Lmap_compose"; | 
| 482 | ||
| 483 | val [prem] = goal LList.thy "M: llist(A) ==> Lmap (%x.x) M = M"; | |
| 484 | by (rtac (prem RS imageI RS LList_equalityI) 1); | |
| 1820 | 485 | by (safe_tac (!claset)); | 
| 969 | 486 | by (etac llist.elim 1); | 
| 1266 | 487 | by (ALLGOALS (asm_simp_tac (!simpset addsimps [Lmap_NIL,Lmap_CONS]))); | 
| 969 | 488 | by (REPEAT (ares_tac [LListD_Fun_NIL_I, imageI RS UnI1, | 
| 1465 | 489 | rangeI RS LListD_Fun_CONS_I] 1)); | 
| 969 | 490 | qed "Lmap_ident"; | 
| 491 | ||
| 492 | ||
| 493 | (*** Lappend -- its two arguments cause some complications! ***) | |
| 494 | ||
| 495 | goalw LList.thy [Lappend_def] "Lappend NIL NIL = NIL"; | |
| 496 | by (rtac (LList_corec RS trans) 1); | |
| 1266 | 497 | by (Simp_tac 1); | 
| 969 | 498 | qed "Lappend_NIL_NIL"; | 
| 499 | ||
| 500 | goalw LList.thy [Lappend_def] | |
| 501 | "Lappend NIL (CONS N N') = CONS N (Lappend NIL N')"; | |
| 502 | by (rtac (LList_corec RS trans) 1); | |
| 1266 | 503 | by (Simp_tac 1); | 
| 969 | 504 | qed "Lappend_NIL_CONS"; | 
| 505 | ||
| 506 | goalw LList.thy [Lappend_def] | |
| 507 | "Lappend (CONS M M') N = CONS M (Lappend M' N)"; | |
| 508 | by (rtac (LList_corec RS trans) 1); | |
| 1266 | 509 | by (Simp_tac 1); | 
| 969 | 510 | qed "Lappend_CONS"; | 
| 511 | ||
| 1266 | 512 | Addsimps [llist.NIL_I, Lappend_NIL_NIL, Lappend_NIL_CONS, | 
| 1465 | 513 | Lappend_CONS, LListD_Fun_CONS_I, range_eqI, image_eqI]; | 
| 1266 | 514 | Delsimps [Pair_eq]; | 
| 969 | 515 | |
| 516 | goal LList.thy "!!M. M: llist(A) ==> Lappend NIL M = M"; | |
| 517 | by (etac LList_fun_equalityI 1); | |
| 1266 | 518 | by (ALLGOALS Asm_simp_tac); | 
| 969 | 519 | qed "Lappend_NIL"; | 
| 520 | ||
| 521 | goal LList.thy "!!M. M: llist(A) ==> Lappend M NIL = M"; | |
| 522 | by (etac LList_fun_equalityI 1); | |
| 1266 | 523 | by (ALLGOALS Asm_simp_tac); | 
| 969 | 524 | qed "Lappend_NIL2"; | 
| 525 | ||
| 526 | (** Alternative type-checking proofs for Lappend **) | |
| 527 | ||
| 528 | (*weak co-induction: bisimulation and case analysis on both variables*) | |
| 529 | goal LList.thy | |
| 530 | "!!M N. [| M: llist(A); N: llist(A) |] ==> Lappend M N: llist(A)"; | |
| 531 | by (res_inst_tac | |
| 532 |     [("X", "UN u:llist(A). UN v: llist(A). {Lappend u v}")] llist_coinduct 1);
 | |
| 1820 | 533 | by (Fast_tac 1); | 
| 534 | by (safe_tac (!claset)); | |
| 969 | 535 | by (eres_inst_tac [("a", "u")] llist.elim 1);
 | 
| 536 | by (eres_inst_tac [("a", "v")] llist.elim 1);
 | |
| 537 | by (ALLGOALS | |
| 1266 | 538 | (Asm_simp_tac THEN' | 
| 1820 | 539 | fast_tac (!claset addSIs [llist.NIL_I, list_Fun_NIL_I, list_Fun_CONS_I]))); | 
| 969 | 540 | qed "Lappend_type"; | 
| 541 | ||
| 542 | (*strong co-induction: bisimulation and case analysis on one variable*) | |
| 543 | goal LList.thy | |
| 544 | "!!M N. [| M: llist(A); N: llist(A) |] ==> Lappend M N: llist(A)"; | |
| 545 | by (res_inst_tac [("X", "(%u.Lappend u N)``llist(A)")] llist_coinduct 1);
 | |
| 1046 | 546 | by (etac imageI 1); | 
| 547 | by (rtac subsetI 1); | |
| 548 | by (etac imageE 1); | |
| 969 | 549 | by (eres_inst_tac [("a", "u")] llist.elim 1);
 | 
| 1266 | 550 | by (asm_simp_tac (!simpset addsimps [Lappend_NIL, list_Fun_llist_I]) 1); | 
| 551 | by (Asm_simp_tac 1); | |
| 1820 | 552 | by (fast_tac (!claset addSIs [list_Fun_CONS_I]) 1); | 
| 969 | 553 | qed "Lappend_type"; | 
| 554 | ||
| 555 | (**** Lazy lists as the type 'a llist -- strongly typed versions of above ****) | |
| 556 | ||
| 557 | (** llist_case: case analysis for 'a llist **) | |
| 558 | ||
| 1266 | 559 | Addsimps ([Abs_llist_inverse, Rep_llist_inverse, | 
| 2911 | 560 | Rep_llist, rangeI, inj_Leaf, inv_f_f] @ llist.intrs); | 
| 969 | 561 | |
| 562 | goalw LList.thy [llist_case_def,LNil_def] "llist_case c d LNil = c"; | |
| 1266 | 563 | by (Simp_tac 1); | 
| 969 | 564 | qed "llist_case_LNil"; | 
| 565 | ||
| 566 | goalw LList.thy [llist_case_def,LCons_def] | |
| 567 | "llist_case c d (LCons M N) = d M N"; | |
| 1266 | 568 | by (Simp_tac 1); | 
| 969 | 569 | qed "llist_case_LCons"; | 
| 570 | ||
| 571 | (*Elimination is case analysis, not induction.*) | |
| 572 | val [prem1,prem2] = goalw LList.thy [NIL_def,CONS_def] | |
| 573 | "[| l=LNil ==> P; !!x l'. l=LCons x l' ==> P \ | |
| 574 | \ |] ==> P"; | |
| 575 | by (rtac (Rep_llist RS llist.elim) 1); | |
| 576 | by (rtac (inj_Rep_llist RS injD RS prem1) 1); | |
| 577 | by (stac Rep_llist_LNil 1); | |
| 578 | by (assume_tac 1); | |
| 579 | by (etac rangeE 1); | |
| 580 | by (rtac (inj_Rep_llist RS injD RS prem2) 1); | |
| 1266 | 581 | by (asm_simp_tac (!simpset delsimps [CONS_CONS_eq] addsimps [Rep_llist_LCons]) 1); | 
| 969 | 582 | by (etac (Abs_llist_inverse RS ssubst) 1); | 
| 583 | by (rtac refl 1); | |
| 584 | qed "llistE"; | |
| 585 | ||
| 586 | (** llist_corec: corecursion for 'a llist **) | |
| 587 | ||
| 588 | goalw LList.thy [llist_corec_def,LNil_def,LCons_def] | |
| 589 | "llist_corec a f = sum_case (%u. LNil) \ | |
| 1465 | 590 | \ (split(%z w. LCons z (llist_corec w f))) (f a)"; | 
| 969 | 591 | by (stac LList_corec 1); | 
| 592 | by (res_inst_tac [("s","f(a)")] sumE 1);
 | |
| 1266 | 593 | by (asm_simp_tac (!simpset addsimps [LList_corec_type2]) 1); | 
| 969 | 594 | by (res_inst_tac [("p","y")] PairE 1);
 | 
| 1266 | 595 | by (asm_simp_tac (!simpset addsimps [LList_corec_type2]) 1); | 
| 969 | 596 | (*FIXME: correct case splits usd to be found automatically: | 
| 1266 | 597 | by (ASM_SIMP_TAC(!simpset addsimps [LList_corec_type2]) 1);*) | 
| 969 | 598 | qed "llist_corec"; | 
| 599 | ||
| 600 | (*definitional version of same*) | |
| 601 | val [rew] = goal LList.thy | |
| 1465 | 602 | "[| !!x. h(x) == llist_corec x f |] ==> \ | 
| 969 | 603 | \ h(a) = sum_case (%u.LNil) (split(%z w. LCons z (h w))) (f a)"; | 
| 604 | by (rewtac rew); | |
| 605 | by (rtac llist_corec 1); | |
| 606 | qed "def_llist_corec"; | |
| 607 | ||
| 608 | (**** Proofs about type 'a llist functions ****) | |
| 609 | ||
| 610 | (*** Deriving llist_equalityI -- llist equality is a bisimulation ***) | |
| 611 | ||
| 612 | goalw LList.thy [LListD_Fun_def] | |
| 1642 | 613 | "!!r A. r <= (llist A) Times (llist A) ==> \ | 
| 614 | \ LListD_Fun (diag A) r <= (llist A) Times (llist A)"; | |
| 969 | 615 | by (stac llist_unfold 1); | 
| 1266 | 616 | by (simp_tac (!simpset addsimps [NIL_def, CONS_def]) 1); | 
| 1820 | 617 | by (Fast_tac 1); | 
| 969 | 618 | qed "LListD_Fun_subset_Sigma_llist"; | 
| 619 | ||
| 620 | goal LList.thy | |
| 621 | "prod_fun Rep_llist Rep_llist `` r <= \ | |
| 1642 | 622 | \ (llist(range Leaf)) Times (llist(range Leaf))"; | 
| 1820 | 623 | by (fast_tac (!claset addIs [Rep_llist]) 1); | 
| 969 | 624 | qed "subset_Sigma_llist"; | 
| 625 | ||
| 626 | val [prem] = goal LList.thy | |
| 1642 | 627 | "r <= (llist(range Leaf)) Times (llist(range Leaf)) ==> \ | 
| 969 | 628 | \ prod_fun (Rep_llist o Abs_llist) (Rep_llist o Abs_llist) `` r <= r"; | 
| 1820 | 629 | by (safe_tac (!claset)); | 
| 969 | 630 | by (rtac (prem RS subsetD RS SigmaE2) 1); | 
| 631 | by (assume_tac 1); | |
| 1266 | 632 | by (asm_simp_tac (!simpset addsimps [o_def,prod_fun,Abs_llist_inverse]) 1); | 
| 969 | 633 | qed "prod_fun_lemma"; | 
| 634 | ||
| 635 | goal LList.thy | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 636 | "prod_fun Rep_llist Rep_llist `` range(%x. (x, x)) = \ | 
| 1642 | 637 | \ diag(llist(range Leaf))"; | 
| 1046 | 638 | by (rtac equalityI 1); | 
| 1820 | 639 | by (fast_tac (!claset addIs [Rep_llist]) 1); | 
| 640 | by (fast_tac (!claset addSEs [Abs_llist_inverse RS subst]) 1); | |
| 969 | 641 | qed "prod_fun_range_eq_diag"; | 
| 642 | ||
| 643 | (** To show two llists are equal, exhibit a bisimulation! | |
| 644 | [also admits true equality] **) | |
| 645 | val [prem1,prem2] = goalw LList.thy [llistD_Fun_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 646 | "[| (l1,l2) : r; r <= llistD_Fun(r Un range(%x.(x,x))) |] ==> l1=l2"; | 
| 969 | 647 | by (rtac (inj_Rep_llist RS injD) 1); | 
| 648 | by (res_inst_tac [("r", "prod_fun Rep_llist Rep_llist ``r"),
 | |
| 1465 | 649 |                   ("A", "range(Leaf)")] 
 | 
| 650 | LList_equalityI 1); | |
| 969 | 651 | by (rtac (prem1 RS prod_fun_imageI) 1); | 
| 652 | by (rtac (prem2 RS image_mono RS subset_trans) 1); | |
| 653 | by (rtac (image_compose RS subst) 1); | |
| 654 | by (rtac (prod_fun_compose RS subst) 1); | |
| 2031 | 655 | by (stac image_Un 1); | 
| 969 | 656 | by (stac prod_fun_range_eq_diag 1); | 
| 657 | by (rtac (LListD_Fun_subset_Sigma_llist RS prod_fun_lemma) 1); | |
| 658 | by (rtac (subset_Sigma_llist RS Un_least) 1); | |
| 659 | by (rtac diag_subset_Sigma 1); | |
| 660 | qed "llist_equalityI"; | |
| 661 | ||
| 662 | (** Rules to prove the 2nd premise of llist_equalityI **) | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 663 | goalw LList.thy [llistD_Fun_def,LNil_def] "(LNil,LNil) : llistD_Fun(r)"; | 
| 969 | 664 | by (rtac (LListD_Fun_NIL_I RS prod_fun_imageI) 1); | 
| 665 | qed "llistD_Fun_LNil_I"; | |
| 666 | ||
| 667 | val [prem] = goalw LList.thy [llistD_Fun_def,LCons_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 668 | "(l1,l2):r ==> (LCons x l1, LCons x l2) : llistD_Fun(r)"; | 
| 969 | 669 | by (rtac (rangeI RS LListD_Fun_CONS_I RS prod_fun_imageI) 1); | 
| 670 | by (rtac (prem RS prod_fun_imageI) 1); | |
| 671 | qed "llistD_Fun_LCons_I"; | |
| 672 | ||
| 673 | (*Utilise the "strong" part, i.e. gfp(f)*) | |
| 674 | goalw LList.thy [llistD_Fun_def] | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 675 | "!!l. (l,l) : llistD_Fun(r Un range(%x.(x,x)))"; | 
| 1046 | 676 | by (rtac (Rep_llist_inverse RS subst) 1); | 
| 677 | by (rtac prod_fun_imageI 1); | |
| 2031 | 678 | by (stac image_Un 1); | 
| 969 | 679 | by (stac prod_fun_range_eq_diag 1); | 
| 1046 | 680 | by (rtac (Rep_llist RS LListD_Fun_diag_I) 1); | 
| 969 | 681 | qed "llistD_Fun_range_I"; | 
| 682 | ||
| 683 | (*A special case of list_equality for functions over lazy lists*) | |
| 684 | val [prem1,prem2] = goal LList.thy | |
| 1465 | 685 | "[| f(LNil)=g(LNil); \ | 
| 686 | \ !!x l. (f(LCons x l),g(LCons x l)) : \ | |
| 687 | \ llistD_Fun(range(%u. (f(u),g(u))) Un range(%v. (v,v))) \ | |
| 688 | \ |] ==> f(l) = (g(l :: 'a llist) :: 'b llist)"; | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 689 | by (res_inst_tac [("r", "range(%u. (f(u),g(u)))")] llist_equalityI 1);
 | 
| 969 | 690 | by (rtac rangeI 1); | 
| 691 | by (rtac subsetI 1); | |
| 692 | by (etac rangeE 1); | |
| 693 | by (etac ssubst 1); | |
| 694 | by (res_inst_tac [("l", "u")] llistE 1);
 | |
| 695 | by (etac ssubst 1); | |
| 696 | by (stac prem1 1); | |
| 697 | by (rtac llistD_Fun_range_I 1); | |
| 698 | by (etac ssubst 1); | |
| 699 | by (rtac prem2 1); | |
| 700 | qed "llist_fun_equalityI"; | |
| 701 | ||
| 702 | (*simpset for llist bisimulations*) | |
| 1266 | 703 | Addsimps [llist_case_LNil, llist_case_LCons, | 
| 1465 | 704 | llistD_Fun_LNil_I, llistD_Fun_LCons_I]; | 
| 969 | 705 | |
| 706 | ||
| 707 | (*** The functional "lmap" ***) | |
| 708 | ||
| 709 | goal LList.thy "lmap f LNil = LNil"; | |
| 710 | by (rtac (lmap_def RS def_llist_corec RS trans) 1); | |
| 1266 | 711 | by (Simp_tac 1); | 
| 969 | 712 | qed "lmap_LNil"; | 
| 713 | ||
| 714 | goal LList.thy "lmap f (LCons M N) = LCons (f M) (lmap f N)"; | |
| 715 | by (rtac (lmap_def RS def_llist_corec RS trans) 1); | |
| 1266 | 716 | by (Simp_tac 1); | 
| 969 | 717 | qed "lmap_LCons"; | 
| 718 | ||
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 719 | Addsimps [lmap_LNil, lmap_LCons]; | 
| 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 720 | |
| 969 | 721 | |
| 722 | (** Two easy results about lmap **) | |
| 723 | ||
| 724 | goal LList.thy "lmap (f o g) l = lmap f (lmap g l)"; | |
| 725 | by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 726 | by (ALLGOALS Simp_tac); | 
| 969 | 727 | qed "lmap_compose"; | 
| 728 | ||
| 729 | goal LList.thy "lmap (%x.x) l = l"; | |
| 730 | by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 731 | by (ALLGOALS Simp_tac); | 
| 969 | 732 | qed "lmap_ident"; | 
| 733 | ||
| 734 | ||
| 735 | (*** iterates -- llist_fun_equalityI cannot be used! ***) | |
| 736 | ||
| 737 | goal LList.thy "iterates f x = LCons x (iterates f (f x))"; | |
| 738 | by (rtac (iterates_def RS def_llist_corec RS trans) 1); | |
| 1266 | 739 | by (Simp_tac 1); | 
| 969 | 740 | qed "iterates"; | 
| 741 | ||
| 742 | goal LList.thy "lmap f (iterates f x) = iterates f (f x)"; | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 743 | by (res_inst_tac [("r", "range(%u.(lmap f (iterates f u),iterates f (f u)))")] 
 | 
| 969 | 744 | llist_equalityI 1); | 
| 745 | by (rtac rangeI 1); | |
| 1820 | 746 | by (safe_tac (!claset)); | 
| 969 | 747 | by (res_inst_tac [("x1", "f(u)")] (iterates RS ssubst) 1);
 | 
| 748 | by (res_inst_tac [("x1", "u")] (iterates RS ssubst) 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 749 | by (Simp_tac 1); | 
| 969 | 750 | qed "lmap_iterates"; | 
| 751 | ||
| 752 | goal LList.thy "iterates f x = LCons x (lmap f (iterates f x))"; | |
| 2031 | 753 | by (stac lmap_iterates 1); | 
| 1046 | 754 | by (rtac iterates 1); | 
| 969 | 755 | qed "iterates_lmap"; | 
| 756 | ||
| 757 | (*** A rather complex proof about iterates -- cf Andy Pitts ***) | |
| 758 | ||
| 759 | (** Two lemmas about natrec n x (%m.g), which is essentially (g^n)(x) **) | |
| 760 | ||
| 761 | goal LList.thy | |
| 1824 | 762 | "nat_rec (LCons b l) (%m. lmap(f)) n = \ | 
| 763 | \ LCons (nat_rec b (%m. f) n) (nat_rec l (%m. lmap(f)) n)"; | |
| 969 | 764 | by (nat_ind_tac "n" 1); | 
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 765 | by (ALLGOALS Asm_simp_tac); | 
| 969 | 766 | qed "fun_power_lmap"; | 
| 767 | ||
| 1824 | 768 | goal Nat.thy "nat_rec (g x) (%m. g) n = nat_rec x (%m. g) (Suc n)"; | 
| 969 | 769 | by (nat_ind_tac "n" 1); | 
| 1266 | 770 | by (ALLGOALS Asm_simp_tac); | 
| 969 | 771 | qed "fun_power_Suc"; | 
| 772 | ||
| 773 | val Pair_cong = read_instantiate_sg (sign_of Prod.thy) | |
| 774 |  [("f","Pair")] (standard(refl RS cong RS cong));
 | |
| 775 | ||
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 776 | (*The bisimulation consists of {(lmap(f)^n (h(u)), lmap(f)^n (iterates(f,u)))}
 | 
| 969 | 777 | for all u and all n::nat.*) | 
| 778 | val [prem] = goal LList.thy | |
| 779 | "(!!x. h(x) = LCons x (lmap f (h x))) ==> h = iterates(f)"; | |
| 1046 | 780 | by (rtac ext 1); | 
| 969 | 781 | by (res_inst_tac [("r", 
 | 
| 1824 | 782 | "UN u. range(%n. (nat_rec (h u) (%m y.lmap f y) n, \ | 
| 783 | \ nat_rec (iterates f u) (%m y.lmap f y) n))")] | |
| 969 | 784 | llist_equalityI 1); | 
| 785 | by (REPEAT (resolve_tac [UN1_I, range_eqI, Pair_cong, nat_rec_0 RS sym] 1)); | |
| 1820 | 786 | by (safe_tac (!claset)); | 
| 969 | 787 | by (stac iterates 1); | 
| 788 | by (stac prem 1); | |
| 789 | by (stac fun_power_lmap 1); | |
| 790 | by (stac fun_power_lmap 1); | |
| 1046 | 791 | by (rtac llistD_Fun_LCons_I 1); | 
| 969 | 792 | by (rtac (lmap_iterates RS subst) 1); | 
| 793 | by (stac fun_power_Suc 1); | |
| 794 | by (stac fun_power_Suc 1); | |
| 1046 | 795 | by (rtac (UN1_I RS UnI1) 1); | 
| 796 | by (rtac rangeI 1); | |
| 969 | 797 | qed "iterates_equality"; | 
| 798 | ||
| 799 | ||
| 800 | (*** lappend -- its two arguments cause some complications! ***) | |
| 801 | ||
| 802 | goalw LList.thy [lappend_def] "lappend LNil LNil = LNil"; | |
| 803 | by (rtac (llist_corec RS trans) 1); | |
| 1266 | 804 | by (Simp_tac 1); | 
| 969 | 805 | qed "lappend_LNil_LNil"; | 
| 806 | ||
| 807 | goalw LList.thy [lappend_def] | |
| 808 | "lappend LNil (LCons l l') = LCons l (lappend LNil l')"; | |
| 809 | by (rtac (llist_corec RS trans) 1); | |
| 1266 | 810 | by (Simp_tac 1); | 
| 969 | 811 | qed "lappend_LNil_LCons"; | 
| 812 | ||
| 813 | goalw LList.thy [lappend_def] | |
| 814 | "lappend (LCons l l') N = LCons l (lappend l' N)"; | |
| 815 | by (rtac (llist_corec RS trans) 1); | |
| 1266 | 816 | by (Simp_tac 1); | 
| 969 | 817 | qed "lappend_LCons"; | 
| 818 | ||
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 819 | Addsimps [lappend_LNil_LNil, lappend_LNil_LCons, lappend_LCons]; | 
| 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 820 | |
| 969 | 821 | goal LList.thy "lappend LNil l = l"; | 
| 822 | by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 823 | by (ALLGOALS Simp_tac); | 
| 969 | 824 | qed "lappend_LNil"; | 
| 825 | ||
| 826 | goal LList.thy "lappend l LNil = l"; | |
| 827 | by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 828 | by (ALLGOALS Simp_tac); | 
| 969 | 829 | qed "lappend_LNil2"; | 
| 830 | ||
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 831 | Addsimps [lappend_LNil, lappend_LNil2]; | 
| 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 832 | |
| 969 | 833 | (*The infinite first argument blocks the second*) | 
| 834 | goal LList.thy "lappend (iterates f x) N = iterates f x"; | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 835 | by (res_inst_tac [("r", "range(%u.(lappend (iterates f u) N,iterates f u))")] 
 | 
| 969 | 836 | llist_equalityI 1); | 
| 837 | by (rtac rangeI 1); | |
| 1820 | 838 | by (safe_tac (!claset)); | 
| 969 | 839 | by (stac iterates 1); | 
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 840 | by (Simp_tac 1); | 
| 969 | 841 | qed "lappend_iterates"; | 
| 842 | ||
| 843 | (** Two proofs that lmap distributes over lappend **) | |
| 844 | ||
| 845 | (*Long proof requiring case analysis on both both arguments*) | |
| 846 | goal LList.thy "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"; | |
| 847 | by (res_inst_tac | |
| 848 |     [("r",  
 | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
969diff
changeset | 849 | "UN n. range(%l.(lmap f (lappend l n),lappend (lmap f l) (lmap f n)))")] | 
| 969 | 850 | llist_equalityI 1); | 
| 851 | by (rtac UN1_I 1); | |
| 852 | by (rtac rangeI 1); | |
| 1820 | 853 | by (safe_tac (!claset)); | 
| 969 | 854 | by (res_inst_tac [("l", "l")] llistE 1);
 | 
| 855 | by (res_inst_tac [("l", "n")] llistE 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 856 | by (ALLGOALS Asm_simp_tac); | 
| 969 | 857 | by (REPEAT_SOME (ares_tac [llistD_Fun_LCons_I, UN1_I RS UnI1, rangeI])); | 
| 858 | qed "lmap_lappend_distrib"; | |
| 859 | ||
| 860 | (*Shorter proof of theorem above using llist_equalityI as strong coinduction*) | |
| 861 | goal LList.thy "lmap f (lappend l n) = lappend (lmap f l) (lmap f n)"; | |
| 862 | by (res_inst_tac [("l","l")] llist_fun_equalityI 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 863 | by (Simp_tac 1); | 
| 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 864 | by (Simp_tac 1); | 
| 969 | 865 | qed "lmap_lappend_distrib"; | 
| 866 | ||
| 867 | (*Without strong coinduction, three case analyses might be needed*) | |
| 868 | goal LList.thy "lappend (lappend l1 l2) l3 = lappend l1 (lappend l2 l3)"; | |
| 869 | by (res_inst_tac [("l","l1")] llist_fun_equalityI 1);
 | |
| 2950 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 870 | by (Simp_tac 1); | 
| 
5d2e0865ecf3
Now puts basic rewrites for lappend & lmap into the simpset
 paulson parents: 
2911diff
changeset | 871 | by (Simp_tac 1); | 
| 969 | 872 | qed "lappend_assoc"; |