| author | wenzelm | 
| Thu, 26 May 1994 16:43:48 +0200 | |
| changeset 400 | 3c2c40c87112 | 
| parent 290 | 37d580c16af5 | 
| child 757 | 2ca12511676d | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: CCL/ex/stream | 
| 2 | ID: $Id$ | |
| 3 | Author: Martin Coen, Cambridge University Computer Laboratory | |
| 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 | |
| 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);
 | |
| 20 | by (fast_tac (ccl_cs addSIs prems) 1); | |
| 21 | by (safe_tac type_cs); | |
| 22 | be (XH_to_E ListsXH) 1; | |
| 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); | 
| 26 | val map_comp = result(); | |
| 27 | ||
| 28 | (*** Mapping the identity function leaves a list unchanged ***) | |
| 29 | ||
| 30 | val prems = goal Stream.thy "l:Lists(A) ==> map(%x.x,l) = l"; | |
| 31 | by (eq_coinduct3_tac | |
| 32 |        "{p. EX x y.p=<x,y> & (EX l:Lists(A).x=map(%x.x,l) & y=l)}"  1);
 | |
| 33 | by (fast_tac (ccl_cs addSIs prems) 1); | |
| 34 | by (safe_tac type_cs); | |
| 35 | be (XH_to_E ListsXH) 1; | |
| 36 | by (EQgen_tac list_ss [] 1); | |
| 37 | by (fast_tac ccl_cs 1); | |
| 38 | val map_id = result(); | |
| 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)"; | |
| 44 | by (eq_coinduct3_tac "{p. EX x y.p=<x,y> & (EX l:Lists(A).EX m:Lists(A). \
 | |
| 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); | |
| 48 | be (XH_to_E ListsXH) 1; | |
| 49 | by (EQgen_tac list_ss [] 1); | |
| 50 | be (XH_to_E ListsXH) 1; | |
| 51 | by (EQgen_tac list_ss [] 1); | |
| 52 | by (fast_tac ccl_cs 1); | |
| 53 | val map_append = result(); | |
| 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"; | |
| 59 | by (eq_coinduct3_tac "{p. EX x y.p=<x,y> & (EX k:Lists(A).EX l:Lists(A).EX m:Lists(A). \
 | |
| 60 | \ x=k @ l @ m & y=(k @ l) @ m)}" 1); | |
| 61 | by (fast_tac (ccl_cs addSIs prems) 1); | |
| 62 | by (safe_tac type_cs); | |
| 63 | be (XH_to_E ListsXH) 1; | |
| 64 | by (EQgen_tac list_ss [] 1); | |
| 65 | be (XH_to_E ListsXH) 1;back(); | |
| 66 | by (EQgen_tac list_ss [] 1); | |
| 67 | be (XH_to_E ListsXH) 1; | |
| 68 | by (EQgen_tac list_ss [] 1); | |
| 69 | by (fast_tac ccl_cs 1); | |
| 70 | val append_assoc = result(); | |
| 71 | ||
| 72 | (*** Appending anything to an infinite list doesn't alter it ****) | |
| 73 | ||
| 74 | val prems = goal Stream.thy "l:ILists(A) ==> l @ m = l"; | |
| 75 | by (eq_coinduct3_tac "{p. EX x y.p=<x,y> & (EX l:ILists(A).EX m.x=l@m & y=l)}" 1);
 | |
| 76 | by (fast_tac (ccl_cs addSIs prems) 1); | |
| 77 | by (safe_tac set_cs); | |
| 78 | be (XH_to_E IListsXH) 1; | |
| 79 | by (EQgen_tac list_ss [] 1); | |
| 80 | by (fast_tac ccl_cs 1); | |
| 81 | val ilist_append = result(); | |
| 82 | ||
| 83 | (*** The equivalance of two versions of an iteration function ***) | |
| 84 | (* *) | |
| 290 
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
 clasohm parents: 
216diff
changeset | 85 | (* fun iter1(f,a) = a$iter1(f,f(a)) *) | 
| 
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
 clasohm parents: 
216diff
changeset | 86 | (* fun iter2(f,a) = a$map(f,iter2(f,a)) *) | 
| 0 | 87 | |
| 290 
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
 clasohm parents: 
216diff
changeset | 88 | goalw Stream.thy [iter1_def] "iter1(f,a) = a$iter1(f,f(a))"; | 
| 0 | 89 | br (letrecB RS trans) 1; | 
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 90 | by (simp_tac term_ss 1); | 
| 0 | 91 | val iter1B = result(); | 
| 92 | ||
| 290 
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
 clasohm parents: 
216diff
changeset | 93 | goalw Stream.thy [iter2_def] "iter2(f,a) = a $ map(f,iter2(f,a))"; | 
| 0 | 94 | br (letrecB RS trans) 1; | 
| 95 | br refl 1; | |
| 96 | val iter2B = result(); | |
| 97 | ||
| 98 | val [prem] =goal Stream.thy | |
| 290 
37d580c16af5
changed "." to "$" and added parentheses to eliminate ambiguity
 clasohm parents: 
216diff
changeset | 99 | "n:Nat ==> map(f) ^ n ` iter2(f,a) = (f ^ n ` a) $ (map(f) ^ n ` map(f,iter2(f,a)))"; | 
| 0 | 100 | br (iter2B RS ssubst) 1;back();back(); | 
| 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); | 
| 0 | 102 | val iter2Blemma = result(); | 
| 103 | ||
| 104 | goal Stream.thy "iter1(f,a) = iter2(f,a)"; | |
| 105 | by (eq_coinduct3_tac | |
| 8 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 106 |     "{p. EX x y.p=<x,y> & (EX n:Nat.x=iter1(f,f^n`a) & y=map(f)^n`iter2(f,a))}"
 | 
| 
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, | 
| 
c3d2c6dcf3f0
Installation of new simplfier.  Previously appeared to set up the old
 lcp parents: 
0diff
changeset | 109 | napplyBzero RS sym RS arg_cong]) 1); | 
| 0 | 110 | by (EQgen_tac list_ss [iter1B,iter2Blemma] 1); | 
| 111 | by (rtac (napply_f RS ssubst) 1 THEN atac 1); | |
| 112 | by (res_inst_tac [("f1","f")] (napplyBsucc RS subst) 1);
 | |
| 113 | by (fast_tac type_cs 1); | |
| 114 | val iter1_iter2_eq = result(); |