# HG changeset patch # User traytel # Date 1355409547 -3600 # Node ID d4fdda801e190cc56967a53dc5737a1882eab6a8 # Parent 8f6c111038204c892729a7f8564d7ccec8b6a121 short library for streams diff -r 8f6c11103820 -r d4fdda801e19 src/HOL/BNF/Examples/Koenig.thy --- a/src/HOL/BNF/Examples/Koenig.thy Thu Dec 13 15:36:08 2012 +0100 +++ b/src/HOL/BNF/Examples/Koenig.thy Thu Dec 13 15:39:07 2012 +0100 @@ -9,11 +9,9 @@ header {* Koenig's lemma *} theory Koenig -imports TreeFI +imports TreeFI Stream begin -codata 'a stream = Stream (shd: 'a) (stl: "'a stream") - (* selectors for streams *) lemma shd_def': "shd as = fst (stream_dtor as)" unfolding shd_def stream_case_def fst_def by (rule refl) diff -r 8f6c11103820 -r d4fdda801e19 src/HOL/BNF/Examples/Stream.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/BNF/Examples/Stream.thy Thu Dec 13 15:39:07 2012 +0100 @@ -0,0 +1,192 @@ +(* Title: HOL/BNF/Examples/Stream.thy + Author: Dmitriy Traytel, TU Muenchen + Author: Andrei Popescu, TU Muenchen + Copyright 2012 + +Infinite streams. +*) + +header {* Infinite Streams *} + +theory Stream +imports "../BNF" +begin + +codata 'a stream = Stream (shd: 'a) (stl: "'a stream") + +(* TODO: Provide by the package*) +theorem stream_set_induct: + "\\s. P (shd s) s; \s y. \y \ stream_set (stl s); P y (stl s)\ \ P y s\ \ + \y \ stream_set s. P y s" +by (rule stream.dtor_set_induct) + (auto simp add: shd_def stl_def stream_case_def fsts_def snds_def split_beta) + +theorem shd_stream_set: "shd s \ stream_set s" +by (auto simp add: shd_def stl_def stream_case_def fsts_def snds_def split_beta) + (metis UnCI fsts_def insertI1 stream.dtor_set) + +theorem stl_stream_set: "y \ stream_set (stl s) \ y \ stream_set s" +by (auto simp add: shd_def stl_def stream_case_def fsts_def snds_def split_beta) + (metis insertI1 set_mp snds_def stream.dtor_set_set_incl) + +(* only for the non-mutual case: *) +theorem stream_set_induct1[consumes 1, case_names shd stl, induct set: "stream_set"]: + assumes "y \ stream_set s" and "\s. P (shd s) s" + and "\s y. \y \ stream_set (stl s); P y (stl s)\ \ P y s" + shows "P y s" +using assms stream_set_induct by blast +(* end TODO *) + + +subsection {* prepend list to stream *} + +primrec shift :: "'a list \ 'a stream \ 'a stream" (infixr "@-" 65) where + "shift [] s = s" +| "shift (x # xs) s = Stream x (shift xs s)" + +lemma shift_append[simp]: "(xs @ ys) @- s = xs @- ys @- s" +by (induct xs) auto + +lemma shift_simps[simp]: + "shd (xs @- s) = (if xs = [] then shd s else hd xs)" + "stl (xs @- s) = (if xs = [] then stl s else tl xs @- s)" +by (induct xs) auto + + +subsection {* recurring stream out of a list *} + +definition cycle :: "'a list \ 'a stream" where + "cycle = stream_unfold hd (\xs. tl xs @ [hd xs])" + +lemma cycle_simps[simp]: + "shd (cycle u) = hd u" + "stl (cycle u) = cycle (tl u @ [hd u])" +by (auto simp: cycle_def) + + +lemma cycle_decomp: "u \ [] \ cycle u = u @- cycle u" +proof (coinduct rule: stream.coinduct[of "\s1 s2. \u. s1 = cycle u \ s2 = u @- cycle u \ u \ []"]) + case (2 s1 s2) + then obtain u where "s1 = cycle u \ s2 = u @- cycle u \ u \ []" by blast + thus ?case using stream.unfold[of hd "\xs. tl xs @ [hd xs]" u] by (auto simp: cycle_def) +qed auto + +lemma cycle_Cons: "cycle (x # xs) = Stream x (cycle (xs @ [x]))" +proof (coinduct rule: stream.coinduct[of "\s1 s2. \x xs. s1 = cycle (x # xs) \ s2 = Stream x (cycle (xs @ [x]))"]) + case (2 s1 s2) + then obtain x xs where "s1 = cycle (x # xs) \ s2 = Stream x (cycle (xs @ [x]))" by blast + thus ?case + by (auto simp: cycle_def intro!: exI[of _ "hd (xs @ [x])"] exI[of _ "tl (xs @ [x])"] stream.unfold) +qed auto + +coinductive_set + streams :: "'a set => 'a stream set" + for A :: "'a set" +where + Stream[intro!, simp, no_atp]: "\a \ A; s \ streams A\ \ Stream a s \ streams A" + +lemma shift_streams: "\w \ lists A; s \ streams A\ \ w @- s \ streams A" +by (induct w) auto + +lemma stream_set_streams: + assumes "stream_set s \ A" + shows "s \ streams A" +proof (coinduct rule: streams.coinduct[of "\s'. \a s. s' = Stream a s \ a \ A \ stream_set s \ A"]) + case streams from assms show ?case by (cases s) auto +next + fix s' assume "\a s. s' = Stream a s \ a \ A \ stream_set s \ A" + then guess a s by (elim exE) + with assms show "\a l. s' = Stream a l \ + a \ A \ ((\a s. l = Stream a s \ a \ A \ stream_set s \ A) \ l \ streams A)" + by (cases s) auto +qed + + +subsection {* flatten a stream of lists *} + +definition flat where + "flat \ stream_unfold (hd o shd) (\s. if tl (shd s) = [] then stl s else Stream (tl (shd s)) (stl s))" + +lemma flat_simps[simp]: + "shd (flat ws) = hd (shd ws)" + "stl (flat ws) = flat (if tl (shd ws) = [] then stl ws else Stream (tl (shd ws)) (stl ws))" +unfolding flat_def by auto + +lemma flat_Cons[simp]: "flat (Stream (x#xs) w) = Stream x (flat (if xs = [] then w else Stream xs w))" +unfolding flat_def using stream.unfold[of "hd o shd" _ "Stream (x#xs) w"] by auto + +lemma flat_Stream[simp]: "xs \ [] \ flat (Stream xs ws) = xs @- flat ws" +by (induct xs) auto + +lemma flat_unfold: "shd ws \ [] \ flat ws = shd ws @- flat (stl ws)" +by (cases ws) auto + + +subsection {* take, drop, nth for streams *} + +primrec stake :: "nat \ 'a stream \ 'a list" where + "stake 0 s = []" +| "stake (Suc n) s = shd s # stake n (stl s)" + +primrec sdrop :: "nat \ 'a stream \ 'a stream" where + "sdrop 0 s = s" +| "sdrop (Suc n) s = sdrop n (stl s)" + +primrec snth :: "nat \ 'a stream \ 'a" where + "snth 0 s = shd s" +| "snth (Suc n) s = snth n (stl s)" + +lemma stake_sdrop: "stake n s @- sdrop n s = s" +by (induct n arbitrary: s) auto + +lemma stake_empty: "stake n s = [] \ n = 0" +by (cases n) auto + +lemma sdrop_shift: "\s = w @- s'; length w = n\ \ sdrop n s = s'" +by (induct n arbitrary: w s) auto + +lemma stake_shift: "\s = w @- s'; length w = n\ \ stake n s = w" +by (induct n arbitrary: w s) auto + +lemma stake_add[simp]: "stake m s @ stake n (sdrop m s) = stake (m + n) s" +by (induct m arbitrary: s) auto + +lemma sdrop_add[simp]: "sdrop n (sdrop m s) = sdrop (m + n) s" +by (induct m arbitrary: s) auto + +lemma cycle_rotated: "\v \ []; cycle u = v @- s\ \ cycle (tl u @ [hd u]) = tl v @- s" +by (auto dest: arg_cong[of _ _ stl]) + +lemma stake_append: "stake n (u @- s) = take (min (length u) n) u @ stake (n - length u) s" +proof (induct n arbitrary: u) + case (Suc n) thus ?case by (cases u) auto +qed auto + +lemma stake_cycle_le[simp]: + assumes "u \ []" "n < length u" + shows "stake n (cycle u) = take n u" +using min_absorb2[OF less_imp_le_nat[OF assms(2)]] +by (subst cycle_decomp[OF assms(1)], subst stake_append) auto + +lemma stake_cycle_eq[simp]: "u \ [] \ stake (length u) (cycle u) = u" +by (metis cycle_decomp stake_shift) + +lemma sdrop_cycle_eq[simp]: "u \ [] \ sdrop (length u) (cycle u) = cycle u" +by (metis cycle_decomp sdrop_shift) + +lemma stake_cycle_eq_mod_0[simp]: "\u \ []; n mod length u = 0\ \ + stake n (cycle u) = concat (replicate (n div length u) u)" +by (induct "n div length u" arbitrary: n u) (auto simp: stake_add[symmetric]) + +lemma sdrop_cycle_eq_mod_0[simp]: "\u \ []; n mod length u = 0\ \ + sdrop n (cycle u) = cycle u" +by (induct "n div length u" arbitrary: n u) (auto simp: sdrop_add[symmetric]) + +lemma stake_cycle: "u \ [] \ + stake n (cycle u) = concat (replicate (n div length u) u) @ take (n mod length u) u" +by (subst mod_div_equality[of n "length u", symmetric], unfold stake_add[symmetric]) auto + +lemma sdrop_cycle: "u \ [] \ sdrop n (cycle u) = cycle (rotate (n mod length u) u)" +by (induct n arbitrary: u) (auto simp: rotate1_rotate_swap rotate1_hd_tl rotate_conv_mod[symmetric]) + +end