# HG changeset patch # User hoelzl # Date 1476981719 -7200 # Node ID ba194424b8959c7f54afd43447b3ba2267df70fb # Parent a33bbac4335952191a6eee7754ef5702d5208f5c HOL-Probability: move stopping time from AFP/Markov_Models diff -r a33bbac43359 -r ba194424b895 src/HOL/Analysis/Borel_Space.thy --- a/src/HOL/Analysis/Borel_Space.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Analysis/Borel_Space.thy Thu Oct 20 18:41:59 2016 +0200 @@ -345,6 +345,14 @@ "A \ sets borel \ insert x A \ sets (borel :: 'a::t1_space measure)" unfolding insert_def by (rule sets.Un) auto +lemma sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV" +proof - + have "(\a\A. {a}) \ sets borel" for A :: "'a set" + by (intro sets.countable_UN') auto + then show ?thesis + by auto +qed + lemma borel_comp[measurable]: "A \ sets borel \ - A \ sets borel" unfolding Compl_eq_Diff_UNIV by simp diff -r a33bbac43359 -r ba194424b895 src/HOL/Analysis/Extended_Real_Limits.thy --- a/src/HOL/Analysis/Extended_Real_Limits.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Analysis/Extended_Real_Limits.thy Thu Oct 20 18:41:59 2016 +0200 @@ -56,6 +56,15 @@ by simp qed +instance enat :: second_countable_topology +proof + show "\B::enat set set. countable B \ open = generate_topology B" + proof (intro exI conjI) + show "countable (range lessThan \ range greaterThan::enat set set)" + by auto + qed (simp add: open_enat_def) +qed + instance ereal :: second_countable_topology proof (standard, intro exI conjI) let ?B = "(\r\\. {{..< r}, {r <..}} :: ereal set set)" diff -r a33bbac43359 -r ba194424b895 src/HOL/Analysis/Measurable.thy --- a/src/HOL/Analysis/Measurable.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Analysis/Measurable.thy Thu Oct 20 18:41:59 2016 +0200 @@ -646,6 +646,14 @@ shows "liminf A \ sets M" by (subst liminf_SUP_INF, auto) +lemma measurable_case_enat[measurable (raw)]: + assumes f: "f \ M \\<^sub>M count_space UNIV" and g: "\i. g i \ M \\<^sub>M N" and h: "h \ M \\<^sub>M N" + shows "(\x. case f x of enat i \ g i x | \ \ h x) \ M \\<^sub>M N" + apply (rule measurable_compose_countable[OF _ f]) + subgoal for i + by (cases i) (auto intro: g h) + done + hide_const (open) pred end diff -r a33bbac43359 -r ba194424b895 src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Analysis/Measure_Space.thy Thu Oct 20 18:41:59 2016 +0200 @@ -1469,6 +1469,9 @@ lemma distr_id[simp]: "distr N N (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) +lemma distr_id2: "sets M = sets N \ distr N M (\x. x) = N" + by (rule measure_eqI) (auto simp: emeasure_distr) + lemma measure_distr: "f \ measurable M N \ S \ sets N \ measure (distr M N f) S = measure M (f -` S \ space M)" by (simp add: emeasure_distr measure_def) @@ -3516,6 +3519,11 @@ finally show ?thesis . qed +lemma measurable_SUP2: + "I \ {} \ (\i. i \ I \ f \ measurable N (M i)) \ + (\i j. i \ I \ j \ I \ space (M i) = space (M j)) \ f \ measurable N (SUP i:I. M i)" + by (auto intro!: measurable_Sup2) + lemma sets_Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "sets (SUP m:M. sigma \ m) = sets (sigma \ (\M))" diff -r a33bbac43359 -r ba194424b895 src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy Thu Oct 20 18:41:59 2016 +0200 @@ -458,6 +458,11 @@ (fastforce simp: topological_space_class.topological_basis_def)+ qed +instance nat :: second_countable_topology +proof + show "\B::nat set set. countable B \ open = generate_topology B" + by (intro exI[of _ "range lessThan \ range greaterThan"]) (auto simp: open_nat_def) +qed lemma countable_separating_set_linorder1: shows "\B::('a::{linorder_topology, second_countable_topology} set). countable B \ (\x y. x < y \ (\b \ B. x < b \ b \ y))" diff -r a33bbac43359 -r ba194424b895 src/HOL/Library/Extended_Nonnegative_Real.thy --- a/src/HOL/Library/Extended_Nonnegative_Real.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Library/Extended_Nonnegative_Real.thy Thu Oct 20 18:41:59 2016 +0200 @@ -1131,6 +1131,9 @@ lemma enn2real_positive_iff: "0 < enn2real x \ (0 < x \ x < top)" by (cases x rule: ennreal_cases) auto +lemma enn2real_eq_1_iff[simp]: "enn2real x = 1 \ x = 1" + by (cases x) auto + subsection \Coercion from @{typ enat} to @{typ ennreal}\ diff -r a33bbac43359 -r ba194424b895 src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy --- a/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Library/Linear_Temporal_Logic_on_Streams.thy Thu Oct 20 18:41:59 2016 +0200 @@ -662,7 +662,7 @@ lemma sfilter_not_P[simp]: "\ P (shd s) \ sfilter P s = sfilter P (stl s)" using sfilter_Stream[of P "shd s" "stl s"] by simp -lemma sfilter_eq: +lemma sfilter_eq: assumes "ev (holds P) s" shows "sfilter P s = x ## s' \ P x \ (not (holds P) suntil (HLD {x} aand nxt (\s. sfilter P s = s'))) s" @@ -685,7 +685,7 @@ proof assume "alw Q (sfilter P s)" with * show "alw (\x. Q (sfilter P x)) s" proof (coinduction arbitrary: s rule: alw_coinduct) - case (stl s) + case (stl s) then have "ev (holds P) s" by blast from this stl show ?case @@ -694,7 +694,7 @@ next assume "alw (\x. Q (sfilter P x)) s" with * show "alw Q (sfilter P s)" proof (coinduction arbitrary: s rule: alw_coinduct) - case (stl s) + case (stl s) then have "ev (holds P) s" by blast from this stl show ?case @@ -767,4 +767,22 @@ lemma hld_smap': "HLD x (smap f s) = HLD (f -` x) s" by (simp add: HLD_def) +lemma pigeonhole_stream: + assumes "alw (HLD s) \" + assumes "finite s" + shows "\x\s. alw (ev (HLD {x})) \" +proof - + have "\i\UNIV. \x\s. \ !! i = x" + using `alw (HLD s) \` by (simp add: alw_iff_sdrop HLD_iff) + from pigeonhole_infinite_rel[OF infinite_UNIV_nat `finite s` this] + show ?thesis + by (simp add: HLD_iff infinite_iff_alw_ev[symmetric]) +qed + +lemma ev_eq_suntil: "ev P \ \ (not P suntil P) \" +proof + assume "ev P \" then show "((\xs. \ P xs) suntil P) \" + by (induction rule: ev_induct_strong) (auto intro: suntil.intros) +qed (auto simp: ev_suntil) + end diff -r a33bbac43359 -r ba194424b895 src/HOL/Library/Stream.thy --- a/src/HOL/Library/Stream.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Library/Stream.thy Thu Oct 20 18:41:59 2016 +0200 @@ -242,7 +242,7 @@ lemma sdrop_snth: "sdrop n s !! m = s !! (n + m)" by (induct n arbitrary: m s) auto -partial_function (tailrec) sdrop_while :: "('a \ bool) \ 'a stream \ 'a stream" where +partial_function (tailrec) sdrop_while :: "('a \ bool) \ 'a stream \ 'a stream" where "sdrop_while P s = (if P (shd s) then sdrop_while P (stl s) else s)" lemma sdrop_while_SCons[code]: @@ -342,7 +342,7 @@ by (induct n arbitrary: u) (auto simp: rotate1_rotate_swap rotate1_hd_tl rotate_conv_mod[symmetric]) lemma sset_cycle[simp]: - assumes "xs \ []" + assumes "xs \ []" shows "sset (cycle xs) = set xs" proof (intro set_eqI iffI) fix x @@ -408,6 +408,14 @@ lemma sconst_streams: "x \ A \ sconst x \ streams A" by (simp add: streams_iff_sset) +lemma streams_empty_iff: "streams S = {} \ S = {}" +proof safe + fix x assume "x \ S" "streams S = {}" + then have "sconst x \ streams S" + by (intro sconst_streams) + then show "x \ {}" + unfolding \streams S = {}\ by simp +qed (auto simp: streams_empty) subsection \stream of natural numbers\ @@ -442,11 +450,11 @@ lemma flat_unfold: "shd ws \ [] \ flat ws = shd ws @- flat (stl ws)" by (cases ws) auto -lemma flat_snth: "\xs \ sset s. xs \ [] \ flat s !! n = (if n < length (shd s) then +lemma flat_snth: "\xs \ sset s. xs \ [] \ flat s !! n = (if n < length (shd s) then shd s ! n else flat (stl s) !! (n - length (shd s)))" by (metis flat_unfold not_less shd_sset shift_snth_ge shift_snth_less) -lemma sset_flat[simp]: "\xs \ sset s. xs \ [] \ +lemma sset_flat[simp]: "\xs \ sset s. xs \ [] \ sset (flat s) = (\xs \ sset s. set xs)" (is "?P \ ?L = ?R") proof safe fix x assume ?P "x : ?L" diff -r a33bbac43359 -r ba194424b895 src/HOL/Probability/Giry_Monad.thy --- a/src/HOL/Probability/Giry_Monad.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Probability/Giry_Monad.thy Thu Oct 20 18:41:59 2016 +0200 @@ -1778,4 +1778,10 @@ shows "space (M \ f) = space B" using M by (intro space_bind[OF sets_kernel[OF f]]) auto +lemma bind_distr_return: + "f \ M \\<^sub>M N \ g \ N \\<^sub>M L \ space M \ {} \ + distr M N f \ (\x. return L (g x)) = distr M L (\x. g (f x))" + by (subst bind_distr[OF _ measurable_compose[OF _ return_measurable]]) + (auto intro!: bind_return_distr') + end diff -r a33bbac43359 -r ba194424b895 src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Probability/Probability.thy Thu Oct 20 18:41:59 2016 +0200 @@ -13,6 +13,7 @@ Stream_Space Conditional_Expectation Essential_Supremum + Stopping_Time begin end diff -r a33bbac43359 -r ba194424b895 src/HOL/Probability/Stopping_Time.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Stopping_Time.thy Thu Oct 20 18:41:59 2016 +0200 @@ -0,0 +1,262 @@ +(* Author: Johannes Hölzl *) + +section {* Stopping times *} + +theory Stopping_Time + imports "../Analysis/Analysis" +begin + +subsection \Stopping Time\ + +text \This is also called strong stopping time. Then stopping time is T with alternative is + \T x < t\ measurable.\ + +definition stopping_time :: "('t::linorder \ 'a measure) \ ('a \ 't) \ bool" +where + "stopping_time F T = (\t. Measurable.pred (F t) (\x. T x \ t))" + +lemma stopping_time_cong: "(\t x. x \ space (F t) \ T x = S x) \ stopping_time F T = stopping_time F S" + unfolding stopping_time_def by (intro arg_cong[where f=All] ext measurable_cong) simp + +lemma stopping_timeD: "stopping_time F T \ Measurable.pred (F t) (\x. T x \ t)" + by (auto simp: stopping_time_def) + +lemma stopping_timeD2: "stopping_time F T \ Measurable.pred (F t) (\x. t < T x)" + unfolding not_le[symmetric] by (auto intro: stopping_timeD Measurable.pred_intros_logic) + +lemma stopping_timeI[intro?]: "(\t. Measurable.pred (F t) (\x. T x \ t)) \ stopping_time F T" + by (auto simp: stopping_time_def) + +lemma measurable_stopping_time: + fixes T :: "'a \ 't::{linorder_topology, second_countable_topology}" + assumes T: "stopping_time F T" + and M: "\t. sets (F t) \ sets M" "\t. space (F t) = space M" + shows "T \ M \\<^sub>M borel" +proof (rule borel_measurableI_le) + show "{x \ space M. T x \ t} \ sets M" for t + using stopping_timeD[OF T] M by (auto simp: Measurable.pred_def) +qed + +lemma stopping_time_const: "stopping_time F (\x. c)" + by (auto simp: stopping_time_def) + +lemma stopping_time_min: + "stopping_time F T \ stopping_time F S \ stopping_time F (\x. min (T x) (S x))" + by (auto simp: stopping_time_def min_le_iff_disj intro!: pred_intros_logic) + +lemma stopping_time_max: + "stopping_time F T \ stopping_time F S \ stopping_time F (\x. max (T x) (S x))" + by (auto simp: stopping_time_def intro!: pred_intros_logic) + +section \Filtration\ + +locale filtration = + fixes \ :: "'a set" and F :: "'t::{linorder_topology, second_countable_topology} \ 'a measure" + assumes space_F: "\i. space (F i) = \" + assumes sets_F_mono: "\i j. i \ j \ sets (F i) \ sets (F j)" +begin + +subsection \$\sigma$-algebra of a Stopping Time\ + +definition pre_sigma :: "('a \ 't) \ 'a measure" +where + "pre_sigma T = sigma \ {A. \t. {\\A. T \ \ t} \ sets (F t)}" + +lemma space_pre_sigma: "space (pre_sigma T) = \" + unfolding pre_sigma_def using sets.space_closed[of "F _"] + by (intro space_measure_of) (auto simp: space_F) + +lemma measure_pre_sigma[simp]: "emeasure (pre_sigma T) = (\_. 0)" + by (simp add: pre_sigma_def emeasure_sigma) + +lemma sigma_algebra_pre_sigma: + assumes T: "stopping_time F T" + shows "sigma_algebra \ {A. \t. {\\A. T \ \ t} \ sets (F t)}" + unfolding sigma_algebra_iff2 +proof (intro sigma_algebra_iff2[THEN iffD2] conjI ballI allI impI CollectI) + show "{A. \t. {\ \ A. T \ \ t} \ sets (F t)} \ Pow \" + using sets.space_closed[of "F _"] by (auto simp: space_F) +next + fix A t assume "A \ {A. \t. {\ \ A. T \ \ t} \ sets (F t)}" + then have "{\ \ space (F t). T \ \ t} - {\ \ A. T \ \ t} \ sets (F t)" + using T stopping_timeD[measurable] by auto + also have "{\ \ space (F t). T \ \ t} - {\ \ A. T \ \ t} = {\ \ \ - A. T \ \ t}" + by (auto simp: space_F) + finally show "{\ \ \ - A. T \ \ t} \ sets (F t)" . +next + fix AA :: "nat \ 'a set" and t assume "range AA \ {A. \t. {\ \ A. T \ \ t} \ sets (F t)}" + then have "(\i. {\ \ AA i. T \ \ t}) \ sets (F t)" for t + by auto + also have "(\i. {\ \ AA i. T \ \ t}) = {\ \ UNION UNIV AA. T \ \ t}" + by auto + finally show "{\ \ UNION UNIV AA. T \ \ t} \ sets (F t)" . +qed auto + +lemma sets_pre_sigma: "stopping_time F T \ sets (pre_sigma T) = {A. \t. {\\A. T \ \ t} \ sets (F t)}" + unfolding pre_sigma_def by (rule sigma_algebra.sets_measure_of_eq[OF sigma_algebra_pre_sigma]) + +lemma sets_pre_sigmaI: "stopping_time F T \ (\t. {\\A. T \ \ t} \ sets (F t)) \ A \ sets (pre_sigma T)" + unfolding sets_pre_sigma by auto + +lemma pred_pre_sigmaI: + assumes T: "stopping_time F T" + shows "(\t. Measurable.pred (F t) (\\. P \ \ T \ \ t)) \ Measurable.pred (pre_sigma T) P" + unfolding pred_def space_F space_pre_sigma by (intro sets_pre_sigmaI[OF T]) simp + +lemma sets_pre_sigmaD: "stopping_time F T \ A \ sets (pre_sigma T) \ {\\A. T \ \ t} \ sets (F t)" + unfolding sets_pre_sigma by auto + +lemma stopping_time_le_const: "stopping_time F T \ s \ t \ Measurable.pred (F t) (\\. T \ \ s)" + using stopping_timeD[of F T] sets_F_mono[of _ t] by (auto simp: pred_def space_F) + +lemma measurable_stopping_time_pre_sigma: + assumes T: "stopping_time F T" shows "T \ pre_sigma T \\<^sub>M borel" +proof (intro borel_measurableI_le sets_pre_sigmaI[OF T]) + fix t t' + have "{\\space (F (min t' t)). T \ \ min t' t} \ sets (F (min t' t))" + using T unfolding pred_def[symmetric] by (rule stopping_timeD) + also have "\ \ sets (F t)" + by (rule sets_F_mono) simp + finally show "{\ \ {x \ space (pre_sigma T). T x \ t'}. T \ \ t} \ sets (F t)" + by (simp add: space_pre_sigma space_F) +qed + +lemma mono_pre_sigma: + assumes T: "stopping_time F T" and S: "stopping_time F S" + and le: "\\. \ \ \ \ T \ \ S \" + shows "sets (pre_sigma T) \ sets (pre_sigma S)" + unfolding sets_pre_sigma[OF S] sets_pre_sigma[OF T] +proof safe + interpret sigma_algebra \ "{A. \t. {\\A. T \ \ t} \ sets (F t)}" + using T by (rule sigma_algebra_pre_sigma) + fix A t assume A: "\t. {\\A. T \ \ t} \ sets (F t)" + then have "A \ \" + using sets_into_space by auto + from A have "{\\A. T \ \ t} \ {\\space (F t). S \ \ t} \ sets (F t)" + using stopping_timeD[OF S] by (auto simp: pred_def) + also have "{\\A. T \ \ t} \ {\\space (F t). S \ \ t} = {\\A. S \ \ t}" + using \A \ \\ sets_into_space[of A] le by (auto simp: space_F intro: order_trans) + finally show "{\\A. S \ \ t} \ sets (F t)" + by auto +qed + +lemma stopping_time_less_const: + assumes T: "stopping_time F T" shows "Measurable.pred (F t) (\\. T \ < t)" +proof - + guess D :: "'t set" by (rule countable_dense_setE) + note D = this + show ?thesis + proof cases + assume *: "\t't''. t' < t'' \ t'' < t" + { fix t' assume "t' < t" + with * have "{t' <..< t} \ {}" + by fastforce + with D(2)[OF _ this] + have "\d\D. t'< d \ d < t" + by auto } + note ** = this + + show ?thesis + proof (rule measurable_cong[THEN iffD2]) + show "T \ < t \ (\r\{r\D. r < t}. T \ \ r)" for \ + by (auto dest: ** intro: less_imp_le) + show "Measurable.pred (F t) (\w. \r\{r \ D. r < t}. T w \ r)" + by (intro measurable_pred_countable stopping_time_le_const[OF T] countable_Collect D) auto + qed + next + assume "\ (\t't''. t' < t'' \ t'' < t)" + then obtain t' where t': "t' < t" "\t''. t'' < t \ t'' \ t'" + by (auto simp: not_less[symmetric]) + show ?thesis + proof (rule measurable_cong[THEN iffD2]) + show "T \ < t \ T \ \ t'" for \ + using t' by auto + show "Measurable.pred (F t) (\w. T w \ t')" + using \t' by (intro stopping_time_le_const[OF T]) auto + qed + qed +qed + +lemma stopping_time_eq_const: "stopping_time F T \ Measurable.pred (F t) (\\. T \ = t)" + unfolding eq_iff using stopping_time_less_const[of T t] + by (intro pred_intros_logic stopping_time_le_const) (auto simp: not_less[symmetric] ) + +lemma stopping_time_less: + assumes T: "stopping_time F T" and S: "stopping_time F S" + shows "Measurable.pred (pre_sigma T) (\\. T \ < S \)" +proof (rule pred_pre_sigmaI[OF T]) + fix t + obtain D :: "'t set" + where [simp]: "countable D" and semidense_D: "\x y. x < y \ (\b\D. x \ b \ b < y)" + using countable_separating_set_linorder2 by auto + show "Measurable.pred (F t) (\\. T \ < S \ \ T \ \ t)" + proof (rule measurable_cong[THEN iffD2]) + let ?f = "\\. if T \ = t then \ S \ \ t else \s\{s\D. s \ t}. T \ \ s \ \ (S \ \ s)" + { fix \ assume "T \ \ t" "T \ \ t" "T \ < S \" + then have "T \ < min t (S \)" + by auto + then obtain r where "r \ D" "T \ \ r" "r < min t (S \)" + by (metis semidense_D) + then have "\s\{s\D. s \ t}. T \ \ s \ s < S \" + by auto } + then show "(T \ < S \ \ T \ \ t) = ?f \" for \ + by (auto simp: not_le) + show "Measurable.pred (F t) ?f" + by (intro pred_intros_logic measurable_If measurable_pred_countable countable_Collect + stopping_time_le_const predE stopping_time_eq_const T S) + auto + qed +qed + +end + +lemma stopping_time_SUP_enat: + fixes T :: "nat \ ('a \ enat)" + shows "(\i. stopping_time F (T i)) \ stopping_time F (SUP i. T i)" + unfolding stopping_time_def SUP_apply SUP_le_iff by (auto intro!: pred_intros_countable) + +lemma less_eSuc_iff: "a < eSuc b \ (a \ b \ a \ \)" + by (cases a) auto + +lemma stopping_time_Inf_enat: + fixes F :: "enat \ 'a measure" + assumes F: "filtration \ F" + assumes P: "\i. Measurable.pred (F i) (P i)" + shows "stopping_time F (\\. Inf {i. P i \})" +proof (rule stopping_timeI, cases) + fix t :: enat assume "t = \" then show "Measurable.pred (F t) (\\. Inf {i. P i \} \ t)" + by auto +next + fix t :: enat assume "t \ \" + moreover + { fix i \ assume "Inf {i. P i \} \ t" + with \t \ \\ have "(\i\t. P i \)" + unfolding Inf_le_iff by (cases t) (auto elim!: allE[of _ "eSuc t"] simp: less_eSuc_iff) } + ultimately have *: "\\. Inf {i. P i \} \ t \ (\i\{..t}. P i \)" + by (auto intro!: Inf_lower2) + show "Measurable.pred (F t) (\\. Inf {i. P i \} \ t)" + unfolding * using filtration.sets_F_mono[OF F, of _ t] P + by (intro pred_intros_countable_bounded) (auto simp: pred_def filtration.space_F[OF F]) +qed + +lemma stopping_time_Inf_nat: + fixes F :: "nat \ 'a measure" + assumes F: "filtration \ F" + assumes P: "\i. Measurable.pred (F i) (P i)" and wf: "\i \. \ \ \ \ \n. P n \" + shows "stopping_time F (\\. Inf {i. P i \})" + unfolding stopping_time_def +proof (intro allI, subst measurable_cong) + fix t \ assume "\ \ space (F t)" + then have "\ \ \" + using filtration.space_F[OF F] by auto + from wf[OF this] have "((LEAST n. P n \) \ t) = (\i\t. P i \)" + by (rule LeastI2_wellorder_ex) auto + then show "(Inf {i. P i \} \ t) = (\i\{..t}. P i \)" + by (simp add: Inf_nat_def Bex_def) +next + fix t from P show "Measurable.pred (F t) (\w. \i\{..t}. P i w)" + using filtration.sets_F_mono[OF F, of _ t] + by (intro pred_intros_countable_bounded) (auto simp: pred_def filtration.space_F[OF F]) +qed + +end diff -r a33bbac43359 -r ba194424b895 src/HOL/Probability/Stream_Space.thy --- a/src/HOL/Probability/Stream_Space.thy Thu Oct 20 18:41:58 2016 +0200 +++ b/src/HOL/Probability/Stream_Space.thy Thu Oct 20 18:41:59 2016 +0200 @@ -446,6 +446,17 @@ by (cases "xs = []") (auto simp: * space_stream_space del: in_listsD) qed (auto simp: * ae sets_M del: in_listsD intro!: streams_sets) +lemma sets_sstart[measurable]: "sstart \ xs \ sets (stream_space (count_space UNIV))" +proof (induction xs) + case (Cons x xs) + note this[measurable] + have "sstart \ (x # xs) = {\\space (stream_space (count_space UNIV)). \ \ sstart \ (x # xs)}" + by (auto simp: space_stream_space) + also have "\ \ sets (stream_space (count_space UNIV))" + unfolding in_sstart by measurable + finally show ?case . +qed (auto intro!: streams_sets) + primrec scylinder :: "'a set \ 'a set list \ 'a stream set" where "scylinder S [] = streams S"