| author | paulson | 
| Wed, 05 Jul 2000 17:42:06 +0200 | |
| changeset 9249 | c71db8c28727 | 
| parent 5062 | fbdb0b541314 | 
| child 17456 | bcf7544875b2 | 
| permissions | -rw-r--r-- | 
| 1459 | 1 | (* Title: CCL/ex/stream | 
| 0 | 2 | ID: $Id$ | 
| 1459 | 3 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 0 | 4 | Copyright 1993 University of Cambridge | 
| 5 | ||
| 6 | For stream.thy. | |
| 7 | ||
| 8 | Proving properties about infinite lists using coinduction: | |
| 9 | Lists(A) is the set of all finite and infinite lists of elements of A. | |
| 10 | ILists(A) is the set of infinite lists of elements of A. | |
| 11 | *) | |
| 12 | ||
| 13 | open Stream; | |
| 14 | ||
| 15 | (*** Map of composition is composition of maps ***) | |
| 16 | ||
| 17 | val prems = goal Stream.thy "l:Lists(A) ==> map(f o g,l) = map(f,map(g,l))"; | |
| 18 | by (eq_coinduct3_tac | |
| 3837 | 19 |        "{p. EX x y. p=<x,y> & (EX l:Lists(A).x=map(f o g,l) & y=map(f,map(g,l)))}"  1);
 | 
| 0 | 20 | by (fast_tac (ccl_cs addSIs prems) 1); | 
| 21 | by (safe_tac type_cs); | |
| 1459 | 22 | by (etac (XH_to_E ListsXH) 1); | 
| 0 | 23 | by (EQgen_tac list_ss [] 1); | 
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 24 | by (simp_tac list_ss 1); | 
| 0 | 25 | by (fast_tac ccl_cs 1); | 
| 757 | 26 | qed "map_comp"; | 
| 0 | 27 | |
| 28 | (*** Mapping the identity function leaves a list unchanged ***) | |
| 29 | ||
| 3837 | 30 | val prems = goal Stream.thy "l:Lists(A) ==> map(%x. x,l) = l"; | 
| 0 | 31 | by (eq_coinduct3_tac | 
| 3837 | 32 |        "{p. EX x y. p=<x,y> & (EX l:Lists(A).x=map(%x. x,l) & y=l)}"  1);
 | 
| 0 | 33 | by (fast_tac (ccl_cs addSIs prems) 1); | 
| 34 | by (safe_tac type_cs); | |
| 1459 | 35 | by (etac (XH_to_E ListsXH) 1); | 
| 0 | 36 | by (EQgen_tac list_ss [] 1); | 
| 37 | by (fast_tac ccl_cs 1); | |
| 757 | 38 | qed "map_id"; | 
| 0 | 39 | |
| 40 | (*** Mapping distributes over append ***) | |
| 41 | ||
| 42 | val prems = goal Stream.thy | |
| 43 | "[| l:Lists(A); m:Lists(A) |] ==> map(f,l@m) = map(f,l) @ map(f,m)"; | |
| 3837 | 44 | by (eq_coinduct3_tac "{p. EX x y. p=<x,y> & (EX l:Lists(A).EX m:Lists(A). \
 | 
| 0 | 45 | \ x=map(f,l@m) & y=map(f,l) @ map(f,m))}" 1); | 
| 46 | by (fast_tac (ccl_cs addSIs prems) 1); | |
| 47 | by (safe_tac type_cs); | |
| 1459 | 48 | by (etac (XH_to_E ListsXH) 1); | 
| 0 | 49 | by (EQgen_tac list_ss [] 1); | 
| 1459 | 50 | by (etac (XH_to_E ListsXH) 1); | 
| 0 | 51 | by (EQgen_tac list_ss [] 1); | 
| 52 | by (fast_tac ccl_cs 1); | |
| 757 | 53 | qed "map_append"; | 
| 0 | 54 | |
| 55 | (*** Append is associative ***) | |
| 56 | ||
| 57 | val prems = goal Stream.thy | |
| 58 | "[| k:Lists(A); l:Lists(A); m:Lists(A) |] ==> k @ l @ m = (k @ l) @ m"; | |
| 1001 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 59 | by (eq_coinduct3_tac | 
| 3837 | 60 |     "{p. EX x y. p=<x,y> & (EX k:Lists(A).EX l:Lists(A).EX m:Lists(A). \
 | 
| 1001 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 61 | \ x=k @ l @ m & y=(k @ l) @ m)}" 1); | 
| 0 | 62 | by (fast_tac (ccl_cs addSIs prems) 1); | 
| 63 | by (safe_tac type_cs); | |
| 1459 | 64 | by (etac (XH_to_E ListsXH) 1); | 
| 0 | 65 | by (EQgen_tac list_ss [] 1); | 
| 1001 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 66 | by (fast_tac ccl_cs 2); | 
| 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 67 | by (DEPTH_SOLVE (etac (XH_to_E ListsXH) 1 THEN EQgen_tac list_ss [] 1)); | 
| 757 | 68 | qed "append_assoc"; | 
| 0 | 69 | |
| 70 | (*** Appending anything to an infinite list doesn't alter it ****) | |
| 71 | ||
| 72 | val prems = goal Stream.thy "l:ILists(A) ==> l @ m = l"; | |
| 1001 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 73 | by (eq_coinduct3_tac | 
| 3837 | 74 |     "{p. EX x y. p=<x,y> & (EX l:ILists(A).EX m. x=l@m & y=l)}" 1);
 | 
| 0 | 75 | by (fast_tac (ccl_cs addSIs prems) 1); | 
| 76 | by (safe_tac set_cs); | |
| 1459 | 77 | by (etac (XH_to_E IListsXH) 1); | 
| 0 | 78 | by (EQgen_tac list_ss [] 1); | 
| 79 | by (fast_tac ccl_cs 1); | |
| 757 | 80 | qed "ilist_append"; | 
| 0 | 81 | |
| 82 | (*** The equivalance of two versions of an iteration function ***) | |
| 83 | (* *) | |
| 290 
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
 clasohm parents: 
216diff
changeset | 84 | (* fun iter1(f,a) = a$iter1(f,f(a)) *) | 
| 
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
 clasohm parents: 
216diff
changeset | 85 | (* fun iter2(f,a) = a$map(f,iter2(f,a)) *) | 
| 0 | 86 | |
| 5062 | 87 | Goalw [iter1_def] "iter1(f,a) = a$iter1(f,f(a))"; | 
| 1459 | 88 | by (rtac (letrecB RS trans) 1); | 
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 89 | by (simp_tac term_ss 1); | 
| 757 | 90 | qed "iter1B"; | 
| 0 | 91 | |
| 5062 | 92 | Goalw [iter2_def] "iter2(f,a) = a $ map(f,iter2(f,a))"; | 
| 1459 | 93 | by (rtac (letrecB RS trans) 1); | 
| 94 | by (rtac refl 1); | |
| 757 | 95 | qed "iter2B"; | 
| 0 | 96 | |
| 97 | val [prem] =goal Stream.thy | |
| 1001 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 98 | "n:Nat ==> \ | 
| 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 99 | \ map(f) ^ n ` iter2(f,a) = (f ^ n ` a) $ (map(f) ^ n ` map(f,iter2(f,a)))"; | 
| 
1f416fb5de91
Simplified some proofs and made them work for new hyp_subst_tac.
 lcp parents: 
757diff
changeset | 100 | by (res_inst_tac [("P", "%x. ?lhs(x) = ?rhs")] (iter2B RS ssubst) 1);
 | 
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 101 | by (simp_tac (list_ss addsimps [prem RS nmapBcons]) 1); | 
| 757 | 102 | qed "iter2Blemma"; | 
| 0 | 103 | |
| 5062 | 104 | Goal "iter1(f,a) = iter2(f,a)"; | 
| 0 | 105 | by (eq_coinduct3_tac | 
| 3837 | 106 |     "{p. EX x y. p=<x,y> & (EX n:Nat. x=iter1(f,f^n`a) & y=map(f)^n`iter2(f,a))}"
 | 
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 107 | 1); | 
| 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 108 | by (fast_tac (type_cs addSIs [napplyBzero RS sym, | 
| 1459 | 109 | napplyBzero RS sym RS arg_cong]) 1); | 
| 0 | 110 | by (EQgen_tac list_ss [iter1B,iter2Blemma] 1); | 
| 2035 | 111 | by (stac napply_f 1 THEN atac 1); | 
| 0 | 112 | by (res_inst_tac [("f1","f")] (napplyBsucc RS subst) 1);
 | 
| 113 | by (fast_tac type_cs 1); | |
| 757 | 114 | qed "iter1_iter2_eq"; |