1(*  Title:      CCL/ex/Stream.thy
2    Author:     Martin Coen, Cambridge University Computer Laboratory
3    Copyright   1993  University of Cambridge
4*)
5
6section \<open>Programs defined over streams\<close>
7
8theory Stream
9imports List
10begin
11
12definition iter1 :: "[i\<Rightarrow>i,i]\<Rightarrow>i"
13  where "iter1(f,a) == letrec iter x be x$iter(f(x)) in iter(a)"
14
15definition iter2 :: "[i\<Rightarrow>i,i]\<Rightarrow>i"
16  where "iter2(f,a) == letrec iter x be x$map(f,iter(x)) in iter(a)"
17
18(*
19Proving 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
25subsection \<open>Map of composition is composition of maps\<close>
26
27lemma map_comp:
28  assumes 1: "l:Lists(A)"
29  shows "map(f \<circ> g,l) = map(f,map(g,l))"
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)))}")
31   apply (blast intro: 1)
32  apply safe
33  apply (drule ListsXH [THEN iffD1])
34  apply EQgen
35   apply fastforce
36  done
37
38(*** Mapping the identity function leaves a list unchanged ***)
39
40lemma map_id:
41  assumes 1: "l:Lists(A)"
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) }")
44  apply (blast intro: 1)
45  apply safe
46  apply (drule ListsXH [THEN iffD1])
47  apply EQgen
48  apply blast
49  done
50
51
52subsection \<open>Mapping distributes over append\<close>
53
54lemma map_append:
55  assumes "l:Lists(A)"
56    and "m:Lists(A)"
57  shows "map(f,l@m) = map(f,l) @ map(f,m)"
58  apply (eq_coinduct3
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))}")
60  apply (blast intro: assms)
61  apply safe
62  apply (drule ListsXH [THEN iffD1])
63  apply EQgen
64  apply (drule ListsXH [THEN iffD1])
65  apply EQgen
66  apply blast
67  done
68
69
70subsection \<open>Append is associative\<close>
71
72lemma 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"
77  apply (eq_coinduct3
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) }")
79  apply (blast intro: assms)
80  apply safe
81  apply (drule ListsXH [THEN iffD1])
82  apply EQgen
83   prefer 2
84   apply blast
85  apply (tactic \<open>DEPTH_SOLVE (eresolve_tac \<^context> [XH_to_E @{thm ListsXH}] 1
86    THEN EQgen_tac \<^context> [] 1)\<close>)
87  done
88
89
90subsection \<open>Appending anything to an infinite list doesn't alter it\<close>
91
92lemma ilist_append:
93  assumes "l:ILists(A)"
94  shows "l @ m = l"
95  apply (eq_coinduct3 "{p. EX x y. p=<x,y> \<and> (EX l:ILists (A) .EX m. x=l@m \<and> y=l)}")
96  apply (blast intro: assms)
97  apply safe
98  apply (drule IListsXH [THEN iffD1])
99  apply EQgen
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
108lemma 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
114lemma 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
120lemma iter2Blemma:
121  "n:Nat \<Longrightarrow>  
122    map(f) ^ n ` iter2(f,a) = (f ^ n ` a) $ (map(f) ^ n ` map(f,iter2(f,a)))"
123  apply (rule_tac P = "\<lambda>x. lhs(x) = rhs" for lhs rhs in iter2B [THEN ssubst])
124  apply (simp add: nmapBcons)
125  done
126
127lemma iter1_iter2_eq: "iter1(f,a) = iter2(f,a)"
128  apply (eq_coinduct3
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))}")
130  apply (fast intro!: napplyBzero [symmetric] napplyBzero [symmetric, THEN arg_cong])
131  apply (EQgen iter1B iter2Blemma)
132  apply (subst napply_f, assumption)
133  apply (rule_tac f1 = f in napplyBsucc [THEN subst])
134  apply blast
135  done
136
137end
138