diff -r f6ffe53387ef -r 0ccec59194af src/HOL/BNF/Examples/Stream.thy --- a/src/HOL/BNF/Examples/Stream.thy Mon Nov 18 12:26:00 2013 +0100 +++ b/src/HOL/BNF/Examples/Stream.thy Mon Nov 18 15:46:52 2013 +0100 @@ -87,10 +87,10 @@ by (induct xs) auto -subsection {* set of streams with elements in some fixes set *} +subsection {* set of streams with elements in some fixed set *} coinductive_set - streams :: "'a set => 'a stream set" + streams :: "'a set \ 'a stream set" for A :: "'a set" where Stream[intro!, simp, no_atp]: "\a \ A; s \ streams A\ \ a ## s \ streams A" @@ -98,6 +98,15 @@ lemma shift_streams: "\w \ lists A; s \ streams A\ \ w @- s \ streams A" by (induct w) auto +lemma streams_Stream: "x ## s \ streams A \ x \ A \ s \ streams A" + by (auto elim: streams.cases) + +lemma streams_stl: "s \ streams A \ stl s \ streams A" + by (cases s) (auto simp: streams_Stream) + +lemma streams_shd: "s \ streams A \ shd s \ A" + by (cases s) (auto simp: streams_Stream) + lemma sset_streams: assumes "sset s \ A" shows "s \ streams A" @@ -105,6 +114,28 @@ case streams then show ?case by (cases s) simp qed +lemma streams_sset: + assumes "s \ streams A" + shows "sset s \ A" +proof + fix x assume "x \ sset s" from this `s \ streams A` show "x \ A" + by (induct s) (auto intro: streams_shd streams_stl) +qed + +lemma streams_iff_sset: "s \ streams A \ sset s \ A" + by (metis sset_streams streams_sset) + +lemma streams_mono: "s \ streams A \ A \ B \ s \ streams B" + unfolding streams_iff_sset by auto + +lemma smap_streams: "s \ streams A \ (\x. x \ A \ f x \ B) \ smap f s \ streams B" + unfolding streams_iff_sset stream.set_map by auto + +lemma streams_empty: "streams {} = {}" + by (auto elim: streams.cases) + +lemma streams_UNIV[simp]: "streams UNIV = UNIV" + by (auto simp: streams_iff_sset) subsection {* nth, take, drop for streams *} @@ -234,6 +265,9 @@ lemma stream_all_shift[simp]: "stream_all P (xs @- s) = (list_all P xs \ stream_all P s)" unfolding stream_all_iff list_all_iff by auto +lemma stream_all_Stream: "stream_all P (x ## X) \ P x \ stream_all P X" + by simp + subsection {* recurring stream out of a list *} @@ -336,6 +370,9 @@ fix m assume "n \ m" thus "m \ ?L" by (metis le_iff_add snth_fromN snth_sset) qed +lemma fromN_Suc_eq_smap: "fromN (Suc n) = smap Suc (fromN n)" + by (coinduction arbitrary: n) auto + abbreviation "nats \ fromN 0"