| author | wenzelm | 
| Wed, 30 Mar 2016 16:36:23 +0200 | |
| changeset 62762 | ac039c4981b6 | 
| parent 60770 | 240563fbf41d | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 17456 | 1 | (* Title: CCL/ex/Stream.thy | 
| 1474 | 2 | Author: Martin Coen, Cambridge University Computer Laboratory | 
| 0 | 3 | Copyright 1993 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 60770 | 6 | section \<open>Programs defined over streams\<close> | 
| 17456 | 7 | |
| 8 | theory Stream | |
| 9 | imports List | |
| 10 | begin | |
| 0 | 11 | |
| 58977 | 12 | definition iter1 :: "[i\<Rightarrow>i,i]\<Rightarrow>i" | 
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
41526diff
changeset | 13 | where "iter1(f,a) == letrec iter x be x$iter(f(x)) in iter(a)" | 
| 0 | 14 | |
| 58977 | 15 | definition iter2 :: "[i\<Rightarrow>i,i]\<Rightarrow>i" | 
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
41526diff
changeset | 16 | where "iter2(f,a) == letrec iter x be x$map(f,iter(x)) in iter(a)" | 
| 20140 | 17 | |
| 18 | (* | |
| 19 | Proving properties about infinite lists using coinduction: | |
| 20 | Lists(A) is the set of all finite and infinite lists of elements of A. | |
| 21 | ILists(A) is the set of infinite lists of elements of A. | |
| 22 | *) | |
| 23 | ||
| 24 | ||
| 60770 | 25 | subsection \<open>Map of composition is composition of maps\<close> | 
| 20140 | 26 | |
| 27 | lemma map_comp: | |
| 28 | assumes 1: "l:Lists(A)" | |
| 42155 
ffe99b07c9c0
modernized specifications -- some attempts to avoid wild axiomatizations;
 wenzelm parents: 
41526diff
changeset | 29 | shows "map(f \<circ> g,l) = map(f,map(g,l))" | 
| 58977 | 30 |   apply (eq_coinduct3 "{p. EX x y. p=<x,y> \<and> (EX l:Lists (A) .x=map (f \<circ> g,l) \<and> y=map (f,map (g,l)))}")
 | 
| 20140 | 31 | apply (blast intro: 1) | 
| 32 | apply safe | |
| 33 | apply (drule ListsXH [THEN iffD1]) | |
| 58971 | 34 | apply EQgen | 
| 47966 | 35 | apply fastforce | 
| 20140 | 36 | done | 
| 37 | ||
| 38 | (*** Mapping the identity function leaves a list unchanged ***) | |
| 39 | ||
| 40 | lemma map_id: | |
| 41 | assumes 1: "l:Lists(A)" | |
| 58977 | 42 | shows "map(\<lambda>x. x, l) = l" | 
| 43 |   apply (eq_coinduct3 "{p. EX x y. p=<x,y> \<and> (EX l:Lists (A) .x=map (\<lambda>x. x,l) \<and> y=l) }")
 | |
| 20140 | 44 | apply (blast intro: 1) | 
| 45 | apply safe | |
| 46 | apply (drule ListsXH [THEN iffD1]) | |
| 58971 | 47 | apply EQgen | 
| 20140 | 48 | apply blast | 
| 49 | done | |
| 50 | ||
| 51 | ||
| 60770 | 52 | subsection \<open>Mapping distributes over append\<close> | 
| 20140 | 53 | |
| 54 | lemma map_append: | |
| 55 | assumes "l:Lists(A)" | |
| 56 | and "m:Lists(A)" | |
| 57 | shows "map(f,l@m) = map(f,l) @ map(f,m)" | |
| 58971 | 58 | apply (eq_coinduct3 | 
| 58977 | 59 |     "{p. EX x y. p=<x,y> \<and> (EX l:Lists (A). EX m:Lists (A). x=map (f,l@m) \<and> y=map (f,l) @ map (f,m))}")
 | 
| 41526 | 60 | apply (blast intro: assms) | 
| 20140 | 61 | apply safe | 
| 62 | apply (drule ListsXH [THEN iffD1]) | |
| 58971 | 63 | apply EQgen | 
| 20140 | 64 | apply (drule ListsXH [THEN iffD1]) | 
| 58971 | 65 | apply EQgen | 
| 20140 | 66 | apply blast | 
| 67 | done | |
| 68 | ||
| 69 | ||
| 60770 | 70 | subsection \<open>Append is associative\<close> | 
| 20140 | 71 | |
| 72 | lemma append_assoc: | |
| 73 | assumes "k:Lists(A)" | |
| 74 | and "l:Lists(A)" | |
| 75 | and "m:Lists(A)" | |
| 76 | shows "k @ l @ m = (k @ l) @ m" | |
| 58971 | 77 | apply (eq_coinduct3 | 
| 58977 | 78 |     "{p. EX x y. p=<x,y> \<and> (EX k:Lists (A). EX l:Lists (A). EX m:Lists (A). x=k @ l @ m \<and> y= (k @ l) @ m) }")
 | 
| 41526 | 79 | apply (blast intro: assms) | 
| 20140 | 80 | apply safe | 
| 81 | apply (drule ListsXH [THEN iffD1]) | |
| 58971 | 82 | apply EQgen | 
| 20140 | 83 | prefer 2 | 
| 84 | apply blast | |
| 60770 | 85 |   apply (tactic \<open>DEPTH_SOLVE (eresolve_tac @{context} [XH_to_E @{thm ListsXH}] 1
 | 
| 86 |     THEN EQgen_tac @{context} [] 1)\<close>)
 | |
| 20140 | 87 | done | 
| 88 | ||
| 89 | ||
| 60770 | 90 | subsection \<open>Appending anything to an infinite list doesn't alter it\<close> | 
| 20140 | 91 | |
| 92 | lemma ilist_append: | |
| 93 | assumes "l:ILists(A)" | |
| 94 | shows "l @ m = l" | |
| 58977 | 95 |   apply (eq_coinduct3 "{p. EX x y. p=<x,y> \<and> (EX l:ILists (A) .EX m. x=l@m \<and> y=l)}")
 | 
| 41526 | 96 | apply (blast intro: assms) | 
| 20140 | 97 | apply safe | 
| 98 | apply (drule IListsXH [THEN iffD1]) | |
| 58971 | 99 | apply EQgen | 
| 20140 | 100 | apply blast | 
| 101 | done | |
| 102 | ||
| 103 | (*** The equivalance of two versions of an iteration function ***) | |
| 104 | (* *) | |
| 105 | (* fun iter1(f,a) = a$iter1(f,f(a)) *) | |
| 106 | (* fun iter2(f,a) = a$map(f,iter2(f,a)) *) | |
| 107 | ||
| 108 | lemma iter1B: "iter1(f,a) = a$iter1(f,f(a))" | |
| 109 | apply (unfold iter1_def) | |
| 110 | apply (rule letrecB [THEN trans]) | |
| 111 | apply simp | |
| 112 | done | |
| 113 | ||
| 114 | lemma iter2B: "iter2(f,a) = a $ map(f,iter2(f,a))" | |
| 115 | apply (unfold iter2_def) | |
| 116 | apply (rule letrecB [THEN trans]) | |
| 117 | apply (rule refl) | |
| 118 | done | |
| 119 | ||
| 120 | lemma iter2Blemma: | |
| 58977 | 121 | "n:Nat \<Longrightarrow> | 
| 20140 | 122 | map(f) ^ n ` iter2(f,a) = (f ^ n ` a) $ (map(f) ^ n ` map(f,iter2(f,a)))" | 
| 59807 | 123 | apply (rule_tac P = "\<lambda>x. lhs(x) = rhs" for lhs rhs in iter2B [THEN ssubst]) | 
| 20140 | 124 | apply (simp add: nmapBcons) | 
| 125 | done | |
| 126 | ||
| 127 | lemma iter1_iter2_eq: "iter1(f,a) = iter2(f,a)" | |
| 58971 | 128 | apply (eq_coinduct3 | 
| 58977 | 129 |     "{p. EX x y. p=<x,y> \<and> (EX n:Nat. x=iter1 (f,f^n`a) \<and> y=map (f) ^n`iter2 (f,a))}")
 | 
| 20140 | 130 | apply (fast intro!: napplyBzero [symmetric] napplyBzero [symmetric, THEN arg_cong]) | 
| 58971 | 131 | apply (EQgen iter1B iter2Blemma) | 
| 20140 | 132 | apply (subst napply_f, assumption) | 
| 133 | apply (rule_tac f1 = f in napplyBsucc [THEN subst]) | |
| 134 | apply blast | |
| 135 | done | |
| 0 | 136 | |
| 137 | end |