| author | haftmann | 
| Thu, 03 Aug 2017 12:49:58 +0200 | |
| changeset 66326 | 9eb8a2d07852 | 
| parent 65366 | 10ca63a18e56 | 
| child 67091 | 1393c2340eec | 
| permissions | -rw-r--r-- | 
| 
58607
 
1f90ea1b4010
move Stream theory from Datatype_Examples to Library
 
hoelzl 
parents: 
58309 
diff
changeset
 | 
1  | 
(* Title: HOL/Library/Stream.thy  | 
| 50518 | 2  | 
Author: Dmitriy Traytel, TU Muenchen  | 
3  | 
Author: Andrei Popescu, TU Muenchen  | 
|
| 51778 | 4  | 
Copyright 2012, 2013  | 
| 50518 | 5  | 
|
6  | 
Infinite streams.  | 
|
7  | 
*)  | 
|
8  | 
||
| 60500 | 9  | 
section \<open>Infinite Streams\<close>  | 
| 50518 | 10  | 
|
11  | 
theory Stream  | 
|
| 65366 | 12  | 
imports Nat_Bijection  | 
| 50518 | 13  | 
begin  | 
14  | 
||
| 
57206
 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 
blanchet 
parents: 
57175 
diff
changeset
 | 
15  | 
codatatype (sset: 'a) stream =  | 
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
16  | 
SCons (shd: 'a) (stl: "'a stream") (infixr "##" 65)  | 
| 
57206
 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 
blanchet 
parents: 
57175 
diff
changeset
 | 
17  | 
for  | 
| 
 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 
blanchet 
parents: 
57175 
diff
changeset
 | 
18  | 
map: smap  | 
| 
 
d9be905d6283
changed syntax of map: and rel: arguments to BNF-based datatypes
 
blanchet 
parents: 
57175 
diff
changeset
 | 
19  | 
rel: stream_all2  | 
| 51409 | 20  | 
|
| 60011 | 21  | 
context  | 
22  | 
begin  | 
|
23  | 
||
| 51462 | 24  | 
(*for code generation only*)  | 
| 60011 | 25  | 
qualified definition smember :: "'a \<Rightarrow> 'a stream \<Rightarrow> bool" where  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
26  | 
[code_abbrev]: "smember x s \<longleftrightarrow> x \<in> sset s"  | 
| 51462 | 27  | 
|
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
28  | 
lemma smember_code[code, simp]: "smember x (y ## s) = (if x = y then True else smember x s)"  | 
| 51462 | 29  | 
unfolding smember_def by auto  | 
30  | 
||
| 60011 | 31  | 
end  | 
| 51462 | 32  | 
|
| 
57983
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57206 
diff
changeset
 | 
33  | 
lemmas smap_simps[simp] = stream.map_sel  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57206 
diff
changeset
 | 
34  | 
lemmas shd_sset = stream.set_sel(1)  | 
| 
 
6edc3529bb4e
reordered some (co)datatype property names for more consistency
 
blanchet 
parents: 
57206 
diff
changeset
 | 
35  | 
lemmas stl_sset = stream.set_sel(2)  | 
| 50518 | 36  | 
|
| 57986 | 37  | 
theorem sset_induct[consumes 1, case_names shd stl, induct set: sset]:  | 
38  | 
assumes "y \<in> sset s" and "\<And>s. P (shd s) s" and "\<And>s y. \<lbrakk>y \<in> sset (stl s); P y (stl s)\<rbrakk> \<Longrightarrow> P y s"  | 
|
| 50518 | 39  | 
shows "P y s"  | 
| 57986 | 40  | 
using assms by induct (metis stream.sel(1), auto)  | 
| 50518 | 41  | 
|
| 59000 | 42  | 
lemma smap_ctr: "smap f s = x ## s' \<longleftrightarrow> f (shd s) = x \<and> smap f (stl s) = s'"  | 
43  | 
by (cases s) simp  | 
|
| 50518 | 44  | 
|
| 60500 | 45  | 
subsection \<open>prepend list to stream\<close>  | 
| 50518 | 46  | 
|
47  | 
primrec shift :: "'a list \<Rightarrow> 'a stream \<Rightarrow> 'a stream" (infixr "@-" 65) where  | 
|
48  | 
"shift [] s = s"  | 
|
| 
51023
 
550f265864e3
infix syntax for streams (reflecting the one for lists)
 
traytel 
parents: 
50518 
diff
changeset
 | 
49  | 
| "shift (x # xs) s = x ## shift xs s"  | 
| 50518 | 50  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
51  | 
lemma smap_shift[simp]: "smap f (xs @- s) = map f xs @- smap f s"  | 
| 51353 | 52  | 
by (induct xs) auto  | 
53  | 
||
| 50518 | 54  | 
lemma shift_append[simp]: "(xs @ ys) @- s = xs @- ys @- s"  | 
| 51141 | 55  | 
by (induct xs) auto  | 
| 50518 | 56  | 
|
57  | 
lemma shift_simps[simp]:  | 
|
58  | 
"shd (xs @- s) = (if xs = [] then shd s else hd xs)"  | 
|
59  | 
"stl (xs @- s) = (if xs = [] then stl s else tl xs @- s)"  | 
|
| 51141 | 60  | 
by (induct xs) auto  | 
| 50518 | 61  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
62  | 
lemma sset_shift[simp]: "sset (xs @- s) = set xs \<union> sset s"  | 
| 51141 | 63  | 
by (induct xs) auto  | 
| 50518 | 64  | 
|
| 51352 | 65  | 
lemma shift_left_inj[simp]: "xs @- s1 = xs @- s2 \<longleftrightarrow> s1 = s2"  | 
66  | 
by (induct xs) auto  | 
|
67  | 
||
| 50518 | 68  | 
|
| 60500 | 69  | 
subsection \<open>set of streams with elements in some fixed set\<close>  | 
| 50518 | 70  | 
|
| 
61681
 
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
 
wenzelm 
parents: 
61424 
diff
changeset
 | 
71  | 
context  | 
| 62093 | 72  | 
notes [[inductive_internals]]  | 
| 
61681
 
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
 
wenzelm 
parents: 
61424 
diff
changeset
 | 
73  | 
begin  | 
| 
 
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
 
wenzelm 
parents: 
61424 
diff
changeset
 | 
74  | 
|
| 50518 | 75  | 
coinductive_set  | 
| 54469 | 76  | 
streams :: "'a set \<Rightarrow> 'a stream set"  | 
| 50518 | 77  | 
for A :: "'a set"  | 
78  | 
where  | 
|
| 
51023
 
550f265864e3
infix syntax for streams (reflecting the one for lists)
 
traytel 
parents: 
50518 
diff
changeset
 | 
79  | 
Stream[intro!, simp, no_atp]: "\<lbrakk>a \<in> A; s \<in> streams A\<rbrakk> \<Longrightarrow> a ## s \<in> streams A"  | 
| 50518 | 80  | 
|
| 
61681
 
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
 
wenzelm 
parents: 
61424 
diff
changeset
 | 
81  | 
end  | 
| 
 
ca53150406c9
option "inductive_defs" controls exposure of def and mono facts;
 
wenzelm 
parents: 
61424 
diff
changeset
 | 
82  | 
|
| 59000 | 83  | 
lemma in_streams: "stl s \<in> streams S \<Longrightarrow> shd s \<in> S \<Longrightarrow> s \<in> streams S"  | 
84  | 
by (cases s) auto  | 
|
85  | 
||
86  | 
lemma streamsE: "s \<in> streams A \<Longrightarrow> (shd s \<in> A \<Longrightarrow> stl s \<in> streams A \<Longrightarrow> P) \<Longrightarrow> P"  | 
|
87  | 
by (erule streams.cases) simp_all  | 
|
88  | 
||
89  | 
lemma Stream_image: "x ## y \<in> (op ## x') ` Y \<longleftrightarrow> x = x' \<and> y \<in> Y"  | 
|
90  | 
by auto  | 
|
91  | 
||
| 50518 | 92  | 
lemma shift_streams: "\<lbrakk>w \<in> lists A; s \<in> streams A\<rbrakk> \<Longrightarrow> w @- s \<in> streams A"  | 
| 51141 | 93  | 
by (induct w) auto  | 
| 50518 | 94  | 
|
| 54469 | 95  | 
lemma streams_Stream: "x ## s \<in> streams A \<longleftrightarrow> x \<in> A \<and> s \<in> streams A"  | 
96  | 
by (auto elim: streams.cases)  | 
|
97  | 
||
98  | 
lemma streams_stl: "s \<in> streams A \<Longrightarrow> stl s \<in> streams A"  | 
|
99  | 
by (cases s) (auto simp: streams_Stream)  | 
|
100  | 
||
101  | 
lemma streams_shd: "s \<in> streams A \<Longrightarrow> shd s \<in> A"  | 
|
102  | 
by (cases s) (auto simp: streams_Stream)  | 
|
103  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
104  | 
lemma sset_streams:  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
105  | 
assumes "sset s \<subseteq> A"  | 
| 50518 | 106  | 
shows "s \<in> streams A"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
107  | 
using assms proof (coinduction arbitrary: s)  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
108  | 
case streams then show ?case by (cases s) simp  | 
| 50518 | 109  | 
qed  | 
110  | 
||
| 54469 | 111  | 
lemma streams_sset:  | 
112  | 
assumes "s \<in> streams A"  | 
|
113  | 
shows "sset s \<subseteq> A"  | 
|
114  | 
proof  | 
|
| 60500 | 115  | 
fix x assume "x \<in> sset s" from this \<open>s \<in> streams A\<close> show "x \<in> A"  | 
| 54469 | 116  | 
by (induct s) (auto intro: streams_shd streams_stl)  | 
117  | 
qed  | 
|
118  | 
||
119  | 
lemma streams_iff_sset: "s \<in> streams A \<longleftrightarrow> sset s \<subseteq> A"  | 
|
120  | 
by (metis sset_streams streams_sset)  | 
|
121  | 
||
122  | 
lemma streams_mono: "s \<in> streams A \<Longrightarrow> A \<subseteq> B \<Longrightarrow> s \<in> streams B"  | 
|
123  | 
unfolding streams_iff_sset by auto  | 
|
124  | 
||
| 59000 | 125  | 
lemma streams_mono2: "S \<subseteq> T \<Longrightarrow> streams S \<subseteq> streams T"  | 
126  | 
by (auto intro: streams_mono)  | 
|
127  | 
||
| 54469 | 128  | 
lemma smap_streams: "s \<in> streams A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> smap f s \<in> streams B"  | 
129  | 
unfolding streams_iff_sset stream.set_map by auto  | 
|
130  | 
||
131  | 
lemma streams_empty: "streams {} = {}"
 | 
|
132  | 
by (auto elim: streams.cases)  | 
|
133  | 
||
134  | 
lemma streams_UNIV[simp]: "streams UNIV = UNIV"  | 
|
135  | 
by (auto simp: streams_iff_sset)  | 
|
| 50518 | 136  | 
|
| 60500 | 137  | 
subsection \<open>nth, take, drop for streams\<close>  | 
| 51141 | 138  | 
|
139  | 
primrec snth :: "'a stream \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 100) where  | 
|
140  | 
"s !! 0 = shd s"  | 
|
141  | 
| "s !! Suc n = stl s !! n"  | 
|
142  | 
||
| 59000 | 143  | 
lemma snth_Stream: "(x ## s) !! Suc i = s !! i"  | 
144  | 
by simp  | 
|
145  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
146  | 
lemma snth_smap[simp]: "smap f s !! n = f (s !! n)"  | 
| 51141 | 147  | 
by (induct n arbitrary: s) auto  | 
148  | 
||
149  | 
lemma shift_snth_less[simp]: "p < length xs \<Longrightarrow> (xs @- s) !! p = xs ! p"  | 
|
150  | 
by (induct p arbitrary: xs) (auto simp: hd_conv_nth nth_tl)  | 
|
151  | 
||
152  | 
lemma shift_snth_ge[simp]: "p \<ge> length xs \<Longrightarrow> (xs @- s) !! p = s !! (p - length xs)"  | 
|
153  | 
by (induct p arbitrary: xs) (auto simp: Suc_diff_eq_diff_pred)  | 
|
154  | 
||
| 57175 | 155  | 
lemma shift_snth: "(xs @- s) !! n = (if n < length xs then xs ! n else s !! (n - length xs))"  | 
156  | 
by auto  | 
|
157  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
158  | 
lemma snth_sset[simp]: "s !! n \<in> sset s"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
159  | 
by (induct n arbitrary: s) (auto intro: shd_sset stl_sset)  | 
| 51141 | 160  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
161  | 
lemma sset_range: "sset s = range (snth s)"  | 
| 51141 | 162  | 
proof (intro equalityI subsetI)  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
163  | 
fix x assume "x \<in> sset s"  | 
| 51141 | 164  | 
thus "x \<in> range (snth s)"  | 
165  | 
proof (induct s)  | 
|
166  | 
case (stl s x)  | 
|
167  | 
then obtain n where "x = stl s !! n" by auto  | 
|
168  | 
thus ?case by (auto intro: range_eqI[of _ _ "Suc n"])  | 
|
169  | 
qed (auto intro: range_eqI[of _ _ 0])  | 
|
170  | 
qed auto  | 
|
| 50518 | 171  | 
|
| 59000 | 172  | 
lemma streams_iff_snth: "s \<in> streams X \<longleftrightarrow> (\<forall>n. s !! n \<in> X)"  | 
173  | 
by (force simp: streams_iff_sset sset_range)  | 
|
174  | 
||
175  | 
lemma snth_in: "s \<in> streams X \<Longrightarrow> s !! n \<in> X"  | 
|
176  | 
by (simp add: streams_iff_snth)  | 
|
177  | 
||
| 50518 | 178  | 
primrec stake :: "nat \<Rightarrow> 'a stream \<Rightarrow> 'a list" where  | 
179  | 
"stake 0 s = []"  | 
|
180  | 
| "stake (Suc n) s = shd s # stake n (stl s)"  | 
|
181  | 
||
| 51141 | 182  | 
lemma length_stake[simp]: "length (stake n s) = n"  | 
183  | 
by (induct n arbitrary: s) auto  | 
|
184  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
185  | 
lemma stake_smap[simp]: "stake n (smap f s) = map f (stake n s)"  | 
| 51141 | 186  | 
by (induct n arbitrary: s) auto  | 
187  | 
||
| 57175 | 188  | 
lemma take_stake: "take n (stake m s) = stake (min n m) s"  | 
189  | 
proof (induct m arbitrary: s n)  | 
|
190  | 
case (Suc m) thus ?case by (cases n) auto  | 
|
191  | 
qed simp  | 
|
192  | 
||
| 50518 | 193  | 
primrec sdrop :: "nat \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where  | 
194  | 
"sdrop 0 s = s"  | 
|
195  | 
| "sdrop (Suc n) s = sdrop n (stl s)"  | 
|
196  | 
||
| 51141 | 197  | 
lemma sdrop_simps[simp]:  | 
198  | 
"shd (sdrop n s) = s !! n" "stl (sdrop n s) = sdrop (Suc n) s"  | 
|
199  | 
by (induct n arbitrary: s) auto  | 
|
200  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
201  | 
lemma sdrop_smap[simp]: "sdrop n (smap f s) = smap f (sdrop n s)"  | 
| 51141 | 202  | 
by (induct n arbitrary: s) auto  | 
| 50518 | 203  | 
|
| 51352 | 204  | 
lemma sdrop_stl: "sdrop n (stl s) = stl (sdrop n s)"  | 
205  | 
by (induct n) auto  | 
|
206  | 
||
| 57175 | 207  | 
lemma drop_stake: "drop n (stake m s) = stake (m - n) (sdrop n s)"  | 
208  | 
proof (induct m arbitrary: s n)  | 
|
209  | 
case (Suc m) thus ?case by (cases n) auto  | 
|
210  | 
qed simp  | 
|
211  | 
||
| 50518 | 212  | 
lemma stake_sdrop: "stake n s @- sdrop n s = s"  | 
| 51141 | 213  | 
by (induct n arbitrary: s) auto  | 
214  | 
||
215  | 
lemma id_stake_snth_sdrop:  | 
|
216  | 
"s = stake i s @- s !! i ## sdrop (Suc i) s"  | 
|
217  | 
by (subst stake_sdrop[symmetric, of _ i]) (metis sdrop_simps stream.collapse)  | 
|
| 50518 | 218  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
219  | 
lemma smap_alt: "smap f s = s' \<longleftrightarrow> (\<forall>n. f (s !! n) = s' !! n)" (is "?L = ?R")  | 
| 51141 | 220  | 
proof  | 
221  | 
assume ?R  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
222  | 
then have "\<And>n. smap f (sdrop n s) = sdrop n s'"  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
223  | 
by coinduction (auto intro: exI[of _ 0] simp del: sdrop.simps(2))  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
224  | 
then show ?L using sdrop.simps(1) by metis  | 
| 51141 | 225  | 
qed auto  | 
226  | 
||
227  | 
lemma stake_invert_Nil[iff]: "stake n s = [] \<longleftrightarrow> n = 0"  | 
|
228  | 
by (induct n) auto  | 
|
| 50518 | 229  | 
|
| 57175 | 230  | 
lemma sdrop_shift: "sdrop i (w @- s) = drop i w @- sdrop (i - length w) s"  | 
231  | 
by (induct i arbitrary: w s) (auto simp: drop_tl drop_Suc neq_Nil_conv)  | 
|
| 50518 | 232  | 
|
| 57175 | 233  | 
lemma stake_shift: "stake i (w @- s) = take i w @ stake (i - length w) s"  | 
234  | 
by (induct i arbitrary: w s) (auto simp: neq_Nil_conv)  | 
|
| 50518 | 235  | 
|
236  | 
lemma stake_add[simp]: "stake m s @ stake n (sdrop m s) = stake (m + n) s"  | 
|
| 51141 | 237  | 
by (induct m arbitrary: s) auto  | 
| 50518 | 238  | 
|
239  | 
lemma sdrop_add[simp]: "sdrop n (sdrop m s) = sdrop (m + n) s"  | 
|
| 51141 | 240  | 
by (induct m arbitrary: s) auto  | 
241  | 
||
| 57175 | 242  | 
lemma sdrop_snth: "sdrop n s !! m = s !! (n + m)"  | 
243  | 
by (induct n arbitrary: m s) auto  | 
|
244  | 
||
| 
64320
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
245  | 
partial_function (tailrec) sdrop_while :: "('a \<Rightarrow> bool) \<Rightarrow> 'a stream \<Rightarrow> 'a stream" where
 | 
| 51430 | 246  | 
"sdrop_while P s = (if P (shd s) then sdrop_while P (stl s) else s)"  | 
247  | 
||
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
248  | 
lemma sdrop_while_SCons[code]:  | 
| 
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
249  | 
"sdrop_while P (a ## s) = (if P a then sdrop_while P s else a ## s)"  | 
| 51430 | 250  | 
by (subst sdrop_while.simps) simp  | 
251  | 
||
252  | 
lemma sdrop_while_sdrop_LEAST:  | 
|
253  | 
assumes "\<exists>n. P (s !! n)"  | 
|
254  | 
shows "sdrop_while (Not o P) s = sdrop (LEAST n. P (s !! n)) s"  | 
|
255  | 
proof -  | 
|
256  | 
from assms obtain m where "P (s !! m)" "\<And>n. P (s !! n) \<Longrightarrow> m \<le> n"  | 
|
257  | 
and *: "(LEAST n. P (s !! n)) = m" by atomize_elim (auto intro: LeastI Least_le)  | 
|
258  | 
thus ?thesis unfolding *  | 
|
259  | 
proof (induct m arbitrary: s)  | 
|
260  | 
case (Suc m)  | 
|
261  | 
hence "sdrop_while (Not \<circ> P) (stl s) = sdrop m (stl s)"  | 
|
262  | 
by (metis (full_types) not_less_eq_eq snth.simps(2))  | 
|
263  | 
moreover from Suc(3) have "\<not> (P (s !! 0))" by blast  | 
|
264  | 
ultimately show ?case by (subst sdrop_while.simps) simp  | 
|
265  | 
qed (metis comp_apply sdrop.simps(1) sdrop_while.simps snth.simps(1))  | 
|
266  | 
qed  | 
|
267  | 
||
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
268  | 
primcorec sfilter where  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
269  | 
"shd (sfilter P s) = shd (sdrop_while (Not o P) s)"  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
270  | 
| "stl (sfilter P s) = sfilter P (stl (sdrop_while (Not o P) s))"  | 
| 52905 | 271  | 
|
272  | 
lemma sfilter_Stream: "sfilter P (x ## s) = (if P x then x ## sfilter P s else sfilter P s)"  | 
|
273  | 
proof (cases "P x")  | 
|
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
274  | 
case True thus ?thesis by (subst sfilter.ctr) (simp add: sdrop_while_SCons)  | 
| 52905 | 275  | 
next  | 
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
276  | 
case False thus ?thesis by (subst (1 2) sfilter.ctr) (simp add: sdrop_while_SCons)  | 
| 52905 | 277  | 
qed  | 
278  | 
||
| 51141 | 279  | 
|
| 60500 | 280  | 
subsection \<open>unary predicates lifted to streams\<close>  | 
| 51141 | 281  | 
|
282  | 
definition "stream_all P s = (\<forall>p. P (s !! p))"  | 
|
283  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
284  | 
lemma stream_all_iff[iff]: "stream_all P s \<longleftrightarrow> Ball (sset s) P"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
285  | 
unfolding stream_all_def sset_range by auto  | 
| 51141 | 286  | 
|
287  | 
lemma stream_all_shift[simp]: "stream_all P (xs @- s) = (list_all P xs \<and> stream_all P s)"  | 
|
288  | 
unfolding stream_all_iff list_all_iff by auto  | 
|
289  | 
||
| 54469 | 290  | 
lemma stream_all_Stream: "stream_all P (x ## X) \<longleftrightarrow> P x \<and> stream_all P X"  | 
291  | 
by simp  | 
|
292  | 
||
| 51141 | 293  | 
|
| 60500 | 294  | 
subsection \<open>recurring stream out of a list\<close>  | 
| 51141 | 295  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
296  | 
primcorec cycle :: "'a list \<Rightarrow> 'a stream" where  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
297  | 
"shd (cycle xs) = hd xs"  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
298  | 
| "stl (cycle xs) = cycle (tl xs @ [hd xs])"  | 
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
299  | 
|
| 51141 | 300  | 
lemma cycle_decomp: "u \<noteq> [] \<Longrightarrow> cycle u = u @- cycle u"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
301  | 
proof (coinduction arbitrary: u)  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
302  | 
case Eq_stream then show ?case using stream.collapse[of "cycle u"]  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
303  | 
by (auto intro!: exI[of _ "tl u @ [hd u]"])  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
304  | 
qed  | 
| 51141 | 305  | 
|
| 51409 | 306  | 
lemma cycle_Cons[code]: "cycle (x # xs) = x ## cycle (xs @ [x])"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
307  | 
by (subst cycle.ctr) simp  | 
| 50518 | 308  | 
|
309  | 
lemma cycle_rotated: "\<lbrakk>v \<noteq> []; cycle u = v @- s\<rbrakk> \<Longrightarrow> cycle (tl u @ [hd u]) = tl v @- s"  | 
|
| 51141 | 310  | 
by (auto dest: arg_cong[of _ _ stl])  | 
| 50518 | 311  | 
|
312  | 
lemma stake_append: "stake n (u @- s) = take (min (length u) n) u @ stake (n - length u) s"  | 
|
313  | 
proof (induct n arbitrary: u)  | 
|
314  | 
case (Suc n) thus ?case by (cases u) auto  | 
|
315  | 
qed auto  | 
|
316  | 
||
317  | 
lemma stake_cycle_le[simp]:  | 
|
318  | 
assumes "u \<noteq> []" "n < length u"  | 
|
319  | 
shows "stake n (cycle u) = take n u"  | 
|
320  | 
using min_absorb2[OF less_imp_le_nat[OF assms(2)]]  | 
|
| 51141 | 321  | 
by (subst cycle_decomp[OF assms(1)], subst stake_append) auto  | 
| 50518 | 322  | 
|
323  | 
lemma stake_cycle_eq[simp]: "u \<noteq> [] \<Longrightarrow> stake (length u) (cycle u) = u"  | 
|
| 57175 | 324  | 
by (subst cycle_decomp) (auto simp: stake_shift)  | 
| 50518 | 325  | 
|
326  | 
lemma sdrop_cycle_eq[simp]: "u \<noteq> [] \<Longrightarrow> sdrop (length u) (cycle u) = cycle u"  | 
|
| 57175 | 327  | 
by (subst cycle_decomp) (auto simp: sdrop_shift)  | 
| 50518 | 328  | 
|
329  | 
lemma stake_cycle_eq_mod_0[simp]: "\<lbrakk>u \<noteq> []; n mod length u = 0\<rbrakk> \<Longrightarrow>  | 
|
330  | 
stake n (cycle u) = concat (replicate (n div length u) u)"  | 
|
| 51141 | 331  | 
by (induct "n div length u" arbitrary: n u) (auto simp: stake_add[symmetric])  | 
| 50518 | 332  | 
|
333  | 
lemma sdrop_cycle_eq_mod_0[simp]: "\<lbrakk>u \<noteq> []; n mod length u = 0\<rbrakk> \<Longrightarrow>  | 
|
334  | 
sdrop n (cycle u) = cycle u"  | 
|
| 51141 | 335  | 
by (induct "n div length u" arbitrary: n u) (auto simp: sdrop_add[symmetric])  | 
| 50518 | 336  | 
|
337  | 
lemma stake_cycle: "u \<noteq> [] \<Longrightarrow>  | 
|
338  | 
stake n (cycle u) = concat (replicate (n div length u) u) @ take (n mod length u) u"  | 
|
| 
64242
 
93c6f0da5c70
more standardized theorem names for facts involving the div and mod identity
 
haftmann 
parents: 
63649 
diff
changeset
 | 
339  | 
by (subst div_mult_mod_eq[of n "length u", symmetric], unfold stake_add[symmetric]) auto  | 
| 50518 | 340  | 
|
341  | 
lemma sdrop_cycle: "u \<noteq> [] \<Longrightarrow> sdrop n (cycle u) = cycle (rotate (n mod length u) u)"  | 
|
| 51141 | 342  | 
by (induct n arbitrary: u) (auto simp: rotate1_rotate_swap rotate1_hd_tl rotate_conv_mod[symmetric])  | 
343  | 
||
| 63192 | 344  | 
lemma sset_cycle[simp]:  | 
| 
64320
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
345  | 
assumes "xs \<noteq> []"  | 
| 63192 | 346  | 
shows "sset (cycle xs) = set xs"  | 
347  | 
proof (intro set_eqI iffI)  | 
|
348  | 
fix x  | 
|
349  | 
assume "x \<in> sset (cycle xs)"  | 
|
350  | 
then show "x \<in> set xs" using assms  | 
|
351  | 
by (induction "cycle xs" arbitrary: xs rule: sset_induct) (fastforce simp: neq_Nil_conv)+  | 
|
352  | 
qed (metis assms UnI1 cycle_decomp sset_shift)  | 
|
353  | 
||
| 51141 | 354  | 
|
| 60500 | 355  | 
subsection \<open>iterated application of a function\<close>  | 
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
356  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
357  | 
primcorec siterate where  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
358  | 
"shd (siterate f x) = x"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
359  | 
| "stl (siterate f x) = siterate f (f x)"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
360  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
361  | 
lemma stake_Suc: "stake (Suc n) s = stake n s @ [s !! n]"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
362  | 
by (induct n arbitrary: s) auto  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
363  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
364  | 
lemma snth_siterate[simp]: "siterate f x !! n = (f^^n) x"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
365  | 
by (induct n arbitrary: x) (auto simp: funpow_swap1)  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
366  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
367  | 
lemma sdrop_siterate[simp]: "sdrop n (siterate f x) = siterate f ((f^^n) x)"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
368  | 
by (induct n arbitrary: x) (auto simp: funpow_swap1)  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
369  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
370  | 
lemma stake_siterate[simp]: "stake n (siterate f x) = map (\<lambda>n. (f^^n) x) [0 ..< n]"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
371  | 
by (induct n arbitrary: x) (auto simp del: stake.simps(2) simp: stake_Suc)  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
372  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
373  | 
lemma sset_siterate: "sset (siterate f x) = {(f^^n) x | n. True}"
 | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
374  | 
by (auto simp: sset_range)  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
375  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
376  | 
lemma smap_siterate: "smap f (siterate f x) = siterate f (f x)"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
377  | 
by (coinduction arbitrary: x) auto  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
378  | 
|
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
379  | 
|
| 60500 | 380  | 
subsection \<open>stream repeating a single element\<close>  | 
| 51141 | 381  | 
|
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
382  | 
abbreviation "sconst \<equiv> siterate id"  | 
| 51141 | 383  | 
|
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
384  | 
lemma shift_replicate_sconst[simp]: "replicate n x @- sconst x = sconst x"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
385  | 
by (subst (3) stake_sdrop[symmetric]) (simp add: map_replicate_trivial)  | 
| 51141 | 386  | 
|
| 57175 | 387  | 
lemma sset_sconst[simp]: "sset (sconst x) = {x}"
 | 
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
388  | 
by (simp add: sset_siterate)  | 
| 51141 | 389  | 
|
| 57175 | 390  | 
lemma sconst_alt: "s = sconst x \<longleftrightarrow> sset s = {x}"
 | 
391  | 
proof  | 
|
392  | 
  assume "sset s = {x}"
 | 
|
393  | 
then show "s = sconst x"  | 
|
394  | 
proof (coinduction arbitrary: s)  | 
|
395  | 
case Eq_stream  | 
|
| 63649 | 396  | 
    then have "shd s = x" "sset (stl s) \<subseteq> {x}" by (cases s; auto)+
 | 
| 57175 | 397  | 
    then have "sset (stl s) = {x}" by (cases "stl s") auto
 | 
| 60500 | 398  | 
with \<open>shd s = x\<close> show ?case by auto  | 
| 57175 | 399  | 
qed  | 
400  | 
qed simp  | 
|
401  | 
||
| 59016 | 402  | 
lemma sconst_cycle: "sconst x = cycle [x]"  | 
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
403  | 
by coinduction auto  | 
| 51141 | 404  | 
|
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
405  | 
lemma smap_sconst: "smap f (sconst x) = sconst (f x)"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
406  | 
by coinduction auto  | 
| 51141 | 407  | 
|
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
408  | 
lemma sconst_streams: "x \<in> A \<Longrightarrow> sconst x \<in> streams A"  | 
| 
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
409  | 
by (simp add: streams_iff_sset)  | 
| 51141 | 410  | 
|
| 
64320
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
411  | 
lemma streams_empty_iff: "streams S = {} \<longleftrightarrow> S = {}"
 | 
| 
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
412  | 
proof safe  | 
| 
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
413  | 
  fix x assume "x \<in> S" "streams S = {}"
 | 
| 
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
414  | 
then have "sconst x \<in> streams S"  | 
| 
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
415  | 
by (intro sconst_streams)  | 
| 
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
416  | 
  then show "x \<in> {}"
 | 
| 
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
417  | 
    unfolding \<open>streams S = {}\<close> by simp
 | 
| 
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
418  | 
qed (auto simp: streams_empty)  | 
| 51141 | 419  | 
|
| 60500 | 420  | 
subsection \<open>stream of natural numbers\<close>  | 
| 51141 | 421  | 
|
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
422  | 
abbreviation "fromN \<equiv> siterate Suc"  | 
| 54469 | 423  | 
|
| 51141 | 424  | 
abbreviation "nats \<equiv> fromN 0"  | 
425  | 
||
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
426  | 
lemma sset_fromN[simp]: "sset (fromN n) = {n ..}"
 | 
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
427  | 
by (auto simp add: sset_siterate le_iff_add)  | 
| 
54497
 
c76dec4df4d7
BNF/Examples/Stream: rename same to sconst; define same, fromN in terms of siterate
 
hoelzl 
parents: 
54469 
diff
changeset
 | 
428  | 
|
| 57175 | 429  | 
lemma stream_smap_fromN: "s = smap (\<lambda>j. let i = j - n in s !! i) (fromN n)"  | 
430  | 
by (coinduction arbitrary: s n)  | 
|
431  | 
(force simp: neq_Nil_conv Let_def snth.simps(2)[symmetric] Suc_diff_Suc  | 
|
432  | 
intro: stream.map_cong split: if_splits simp del: snth.simps(2))  | 
|
433  | 
||
434  | 
lemma stream_smap_nats: "s = smap (snth s) nats"  | 
|
435  | 
using stream_smap_fromN[where n = 0] by simp  | 
|
436  | 
||
| 51141 | 437  | 
|
| 60500 | 438  | 
subsection \<open>flatten a stream of lists\<close>  | 
| 51462 | 439  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
440  | 
primcorec flat where  | 
| 51462 | 441  | 
"shd (flat ws) = hd (shd ws)"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
442  | 
| "stl (flat ws) = flat (if tl (shd ws) = [] then stl ws else tl (shd ws) ## stl ws)"  | 
| 51462 | 443  | 
|
444  | 
lemma flat_Cons[simp, code]: "flat ((x # xs) ## ws) = x ## flat (if xs = [] then ws else xs ## ws)"  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
445  | 
by (subst flat.ctr) simp  | 
| 51462 | 446  | 
|
447  | 
lemma flat_Stream[simp]: "xs \<noteq> [] \<Longrightarrow> flat (xs ## ws) = xs @- flat ws"  | 
|
448  | 
by (induct xs) auto  | 
|
449  | 
||
450  | 
lemma flat_unfold: "shd ws \<noteq> [] \<Longrightarrow> flat ws = shd ws @- flat (stl ws)"  | 
|
451  | 
by (cases ws) auto  | 
|
452  | 
||
| 
64320
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
453  | 
lemma flat_snth: "\<forall>xs \<in> sset s. xs \<noteq> [] \<Longrightarrow> flat s !! n = (if n < length (shd s) then  | 
| 51462 | 454  | 
shd s ! n else flat (stl s) !! (n - length (shd s)))"  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
455  | 
by (metis flat_unfold not_less shd_sset shift_snth_ge shift_snth_less)  | 
| 51462 | 456  | 
|
| 
64320
 
ba194424b895
HOL-Probability: move stopping time from AFP/Markov_Models
 
hoelzl 
parents: 
64242 
diff
changeset
 | 
457  | 
lemma sset_flat[simp]: "\<forall>xs \<in> sset s. xs \<noteq> [] \<Longrightarrow>  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
458  | 
sset (flat s) = (\<Union>xs \<in> sset s. set xs)" (is "?P \<Longrightarrow> ?L = ?R")  | 
| 51462 | 459  | 
proof safe  | 
460  | 
fix x assume ?P "x : ?L"  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
461  | 
then obtain m where "x = flat s !! m" by (metis image_iff sset_range)  | 
| 60500 | 462  | 
with \<open>?P\<close> obtain n m' where "x = s !! n ! m'" "m' < length (s !! n)"  | 
| 51462 | 463  | 
proof (atomize_elim, induct m arbitrary: s rule: less_induct)  | 
464  | 
case (less y)  | 
|
465  | 
thus ?case  | 
|
466  | 
proof (cases "y < length (shd s)")  | 
|
467  | 
case True thus ?thesis by (metis flat_snth less(2,3) snth.simps(1))  | 
|
468  | 
next  | 
|
469  | 
case False  | 
|
470  | 
hence "x = flat (stl s) !! (y - length (shd s))" by (metis less(2,3) flat_snth)  | 
|
471  | 
moreover  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53290 
diff
changeset
 | 
472  | 
      { from less(2) have *: "length (shd s) > 0" by (cases s) simp_all
 | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53290 
diff
changeset
 | 
473  | 
with False have "y > 0" by (cases y) simp_all  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53290 
diff
changeset
 | 
474  | 
with * have "y - length (shd s) < y" by simp  | 
| 51462 | 475  | 
}  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
476  | 
moreover have "\<forall>xs \<in> sset (stl s). xs \<noteq> []" using less(2) by (cases s) auto  | 
| 51462 | 477  | 
ultimately have "\<exists>n m'. x = stl s !! n ! m' \<and> m' < length (stl s !! n)" by (intro less(1)) auto  | 
478  | 
thus ?thesis by (metis snth.simps(2))  | 
|
479  | 
qed  | 
|
480  | 
qed  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
481  | 
thus "x \<in> ?R" by (auto simp: sset_range dest!: nth_mem)  | 
| 51462 | 482  | 
next  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
483  | 
fix x xs assume "xs \<in> sset s" ?P "x \<in> set xs" thus "x \<in> ?L"  | 
| 57986 | 484  | 
by (induct rule: sset_induct)  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
485  | 
(metis UnI1 flat_unfold shift.simps(1) sset_shift,  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
486  | 
metis UnI2 flat_unfold shd_sset stl_sset sset_shift)  | 
| 51462 | 487  | 
qed  | 
488  | 
||
489  | 
||
| 60500 | 490  | 
subsection \<open>merge a stream of streams\<close>  | 
| 51462 | 491  | 
|
492  | 
definition smerge :: "'a stream stream \<Rightarrow> 'a stream" where  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
493  | 
"smerge ss = flat (smap (\<lambda>n. map (\<lambda>s. s !! n) (stake (Suc n) ss) @ stake n (ss !! n)) nats)"  | 
| 51462 | 494  | 
|
495  | 
lemma stake_nth[simp]: "m < n \<Longrightarrow> stake n s ! m = s !! m"  | 
|
496  | 
by (induct n arbitrary: s m) (auto simp: nth_Cons', metis Suc_pred snth.simps(2))  | 
|
497  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
498  | 
lemma snth_sset_smerge: "ss !! n !! m \<in> sset (smerge ss)"  | 
| 51462 | 499  | 
proof (cases "n \<le> m")  | 
500  | 
case False thus ?thesis unfolding smerge_def  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
501  | 
by (subst sset_flat)  | 
| 53290 | 502  | 
(auto simp: stream.set_map in_set_conv_nth simp del: stake.simps  | 
| 51462 | 503  | 
intro!: exI[of _ n, OF disjI2] exI[of _ m, OF mp])  | 
504  | 
next  | 
|
505  | 
case True thus ?thesis unfolding smerge_def  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
506  | 
by (subst sset_flat)  | 
| 53290 | 507  | 
(auto simp: stream.set_map in_set_conv_nth image_iff simp del: stake.simps snth.simps  | 
| 51462 | 508  | 
intro!: exI[of _ m, OF disjI1] bexI[of _ "ss !! n"] exI[of _ n, OF mp])  | 
509  | 
qed  | 
|
510  | 
||
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
511  | 
lemma sset_smerge: "sset (smerge ss) = UNION (sset ss) sset"  | 
| 51462 | 512  | 
proof safe  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
513  | 
fix x assume "x \<in> sset (smerge ss)"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
514  | 
thus "x \<in> UNION (sset ss) sset"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
515  | 
unfolding smerge_def by (subst (asm) sset_flat)  | 
| 53290 | 516  | 
(auto simp: stream.set_map in_set_conv_nth sset_range simp del: stake.simps, fast+)  | 
| 51462 | 517  | 
next  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
518  | 
fix s x assume "s \<in> sset ss" "x \<in> sset s"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
519  | 
thus "x \<in> sset (smerge ss)" using snth_sset_smerge by (auto simp: sset_range)  | 
| 51462 | 520  | 
qed  | 
521  | 
||
522  | 
||
| 60500 | 523  | 
subsection \<open>product of two streams\<close>  | 
| 51462 | 524  | 
|
525  | 
definition sproduct :: "'a stream \<Rightarrow> 'b stream \<Rightarrow> ('a \<times> 'b) stream" where
 | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
526  | 
"sproduct s1 s2 = smerge (smap (\<lambda>x. smap (Pair x) s2) s1)"  | 
| 51462 | 527  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
528  | 
lemma sset_sproduct: "sset (sproduct s1 s2) = sset s1 \<times> sset s2"  | 
| 53290 | 529  | 
unfolding sproduct_def sset_smerge by (auto simp: stream.set_map)  | 
| 51462 | 530  | 
|
531  | 
||
| 60500 | 532  | 
subsection \<open>interleave two streams\<close>  | 
| 51462 | 533  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
534  | 
primcorec sinterleave where  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
535  | 
"shd (sinterleave s1 s2) = shd s1"  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
536  | 
| "stl (sinterleave s1 s2) = sinterleave s2 (stl s1)"  | 
| 51462 | 537  | 
|
538  | 
lemma sinterleave_code[code]:  | 
|
539  | 
"sinterleave (x ## s1) s2 = x ## sinterleave s2 s1"  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
540  | 
by (subst sinterleave.ctr) simp  | 
| 51462 | 541  | 
|
542  | 
lemma sinterleave_snth[simp]:  | 
|
543  | 
"even n \<Longrightarrow> sinterleave s1 s2 !! n = s1 !! (n div 2)"  | 
|
| 
58710
 
7216a10d69ba
augmented and tuned facts on even/odd and division
 
haftmann 
parents: 
58607 
diff
changeset
 | 
544  | 
"odd n \<Longrightarrow> sinterleave s1 s2 !! n = s2 !! (n div 2)"  | 
| 
 
7216a10d69ba
augmented and tuned facts on even/odd and division
 
haftmann 
parents: 
58607 
diff
changeset
 | 
545  | 
by (induct n arbitrary: s1 s2) simp_all  | 
| 51462 | 546  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
547  | 
lemma sset_sinterleave: "sset (sinterleave s1 s2) = sset s1 \<union> sset s2"  | 
| 51462 | 548  | 
proof (intro equalityI subsetI)  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
549  | 
fix x assume "x \<in> sset (sinterleave s1 s2)"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
550  | 
then obtain n where "x = sinterleave s1 s2 !! n" unfolding sset_range by blast  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
551  | 
thus "x \<in> sset s1 \<union> sset s2" by (cases "even n") auto  | 
| 51462 | 552  | 
next  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
553  | 
fix x assume "x \<in> sset s1 \<union> sset s2"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
554  | 
thus "x \<in> sset (sinterleave s1 s2)"  | 
| 51462 | 555  | 
proof  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
556  | 
assume "x \<in> sset s1"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
557  | 
then obtain n where "x = s1 !! n" unfolding sset_range by blast  | 
| 51462 | 558  | 
hence "sinterleave s1 s2 !! (2 * n) = x" by simp  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
559  | 
thus ?thesis unfolding sset_range by blast  | 
| 51462 | 560  | 
next  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
561  | 
assume "x \<in> sset s2"  | 
| 
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
562  | 
then obtain n where "x = s2 !! n" unfolding sset_range by blast  | 
| 51462 | 563  | 
hence "sinterleave s1 s2 !! (2 * n + 1) = x" by simp  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
564  | 
thus ?thesis unfolding sset_range by blast  | 
| 51462 | 565  | 
qed  | 
566  | 
qed  | 
|
567  | 
||
568  | 
||
| 60500 | 569  | 
subsection \<open>zip\<close>  | 
| 51141 | 570  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
571  | 
primcorec szip where  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
572  | 
"shd (szip s1 s2) = (shd s1, shd s2)"  | 
| 
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
573  | 
| "stl (szip s1 s2) = szip (stl s1) (stl s2)"  | 
| 51141 | 574  | 
|
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
575  | 
lemma szip_unfold[code]: "szip (a ## s1) (b ## s2) = (a, b) ## (szip s1 s2)"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
576  | 
by (subst szip.ctr) simp  | 
| 51409 | 577  | 
|
| 51141 | 578  | 
lemma snth_szip[simp]: "szip s1 s2 !! n = (s1 !! n, s2 !! n)"  | 
579  | 
by (induct n arbitrary: s1 s2) auto  | 
|
580  | 
||
| 57175 | 581  | 
lemma stake_szip[simp]:  | 
582  | 
"stake n (szip s1 s2) = zip (stake n s1) (stake n s2)"  | 
|
583  | 
by (induct n arbitrary: s1 s2) auto  | 
|
584  | 
||
585  | 
lemma sdrop_szip[simp]: "sdrop n (szip s1 s2) = szip (sdrop n s1) (sdrop n s2)"  | 
|
586  | 
by (induct n arbitrary: s1 s2) auto  | 
|
587  | 
||
588  | 
lemma smap_szip_fst:  | 
|
589  | 
"smap (\<lambda>x. f (fst x)) (szip s1 s2) = smap f s1"  | 
|
590  | 
by (coinduction arbitrary: s1 s2) auto  | 
|
591  | 
||
592  | 
lemma smap_szip_snd:  | 
|
593  | 
"smap (\<lambda>x. g (snd x)) (szip s1 s2) = smap g s2"  | 
|
594  | 
by (coinduction arbitrary: s1 s2) auto  | 
|
595  | 
||
| 51141 | 596  | 
|
| 60500 | 597  | 
subsection \<open>zip via function\<close>  | 
| 51141 | 598  | 
|
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
599  | 
primcorec smap2 where  | 
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
600  | 
"shd (smap2 f s1 s2) = f (shd s1) (shd s2)"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
601  | 
| "stl (smap2 f s1 s2) = smap2 f (stl s1) (stl s2)"  | 
| 51141 | 602  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
603  | 
lemma smap2_unfold[code]:  | 
| 
54720
 
0a9920e46b3a
removed obsolete codegen setup; Stream -> SCons; tuned
 
traytel 
parents: 
54498 
diff
changeset
 | 
604  | 
"smap2 f (a ## s1) (b ## s2) = f a b ## (smap2 f s1 s2)"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
605  | 
by (subst smap2.ctr) simp  | 
| 51409 | 606  | 
|
| 
51772
 
d2b265ebc1fa
specify nicer names for map, set and rel in the stream library
 
traytel 
parents: 
51766 
diff
changeset
 | 
607  | 
lemma smap2_szip:  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60500 
diff
changeset
 | 
608  | 
"smap2 f s1 s2 = smap (case_prod f) (szip s1 s2)"  | 
| 
54027
 
e5853a648b59
use new coinduction method and primcorec in examples
 
traytel 
parents: 
53808 
diff
changeset
 | 
609  | 
by (coinduction arbitrary: s1 s2) auto  | 
| 50518 | 610  | 
|
| 57175 | 611  | 
lemma smap_smap2[simp]:  | 
612  | 
"smap f (smap2 g s1 s2) = smap2 (\<lambda>x y. f (g x y)) s1 s2"  | 
|
613  | 
unfolding smap2_szip stream.map_comp o_def split_def ..  | 
|
614  | 
||
615  | 
lemma smap2_alt:  | 
|
616  | 
"(smap2 f s1 s2 = s) = (\<forall>n. f (s1 !! n) (s2 !! n) = s !! n)"  | 
|
617  | 
unfolding smap2_szip smap_alt by auto  | 
|
618  | 
||
619  | 
lemma snth_smap2[simp]:  | 
|
620  | 
"smap2 f s1 s2 !! n = f (s1 !! n) (s2 !! n)"  | 
|
621  | 
by (induct n arbitrary: s1 s2) auto  | 
|
622  | 
||
623  | 
lemma stake_smap2[simp]:  | 
|
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
60500 
diff
changeset
 | 
624  | 
"stake n (smap2 f s1 s2) = map (case_prod f) (zip (stake n s1) (stake n s2))"  | 
| 57175 | 625  | 
by (induct n arbitrary: s1 s2) auto  | 
626  | 
||
627  | 
lemma sdrop_smap2[simp]:  | 
|
628  | 
"sdrop n (smap2 f s1 s2) = smap2 f (sdrop n s1) (sdrop n s2)"  | 
|
629  | 
by (induct n arbitrary: s1 s2) auto  | 
|
630  | 
||
| 50518 | 631  | 
end  |