src/HOL/Probability/Independent_Family.thy
author hoelzl
Thu, 26 May 2011 14:11:58 +0200
changeset 43823 fa0ac7bee9ac
parent 43822 fe7f5a26e4c6
child 43824 685df9c0626d
permissions -rw-r--r--
add lemma kolmogorov_0_1_law
     1 (*  Title:      HOL/Probability/Independent_Family.thy
     2     Author:     Johannes Hölzl, TU München
     3 *)
     4 
     5 header {* Independent families of events, event sets, and random variables *}
     6 
     7 theory Independent_Family
     8   imports Probability_Measure
     9 begin
    10 
    11 definition (in prob_space)
    12   "indep_events A I \<longleftrightarrow> (A`I \<subseteq> sets M) \<and>
    13     (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j)))"
    14 
    15 definition (in prob_space)
    16   "indep_event A B \<longleftrightarrow> indep_events (bool_case A B) UNIV"
    17 
    18 definition (in prob_space)
    19   "indep_sets F I \<longleftrightarrow> (\<forall>i\<in>I. F i \<subseteq> sets M) \<and>
    20     (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> (\<forall>A\<in>Pi J F. prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))))"
    21 
    22 definition (in prob_space)
    23   "indep_set A B \<longleftrightarrow> indep_sets (bool_case A B) UNIV"
    24 
    25 definition (in prob_space)
    26   "indep_rv M' X I \<longleftrightarrow>
    27     (\<forall>i\<in>I. random_variable (M' i) (X i)) \<and>
    28     indep_sets (\<lambda>i. sigma_sets (space M) { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
    29 
    30 lemma (in prob_space) indep_sets_cong:
    31   "I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i = G i) \<Longrightarrow> indep_sets F I \<longleftrightarrow> indep_sets G J"
    32   by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+
    33 
    34 lemma (in prob_space) indep_events_finite_index_events:
    35   "indep_events F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_events F J)"
    36   by (auto simp: indep_events_def)
    37 
    38 lemma (in prob_space) indep_sets_finite_index_sets:
    39   "indep_sets F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_sets F J)"
    40 proof (intro iffI allI impI)
    41   assume *: "\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_sets F J"
    42   show "indep_sets F I" unfolding indep_sets_def
    43   proof (intro conjI ballI allI impI)
    44     fix i assume "i \<in> I"
    45     with *[THEN spec, of "{i}"] show "F i \<subseteq> events"
    46       by (auto simp: indep_sets_def)
    47   qed (insert *, auto simp: indep_sets_def)
    48 qed (auto simp: indep_sets_def)
    49 
    50 lemma (in prob_space) indep_sets_mono_index:
    51   "J \<subseteq> I \<Longrightarrow> indep_sets F I \<Longrightarrow> indep_sets F J"
    52   unfolding indep_sets_def by auto
    53 
    54 lemma (in prob_space) indep_sets_mono_sets:
    55   assumes indep: "indep_sets F I"
    56   assumes mono: "\<And>i. i\<in>I \<Longrightarrow> G i \<subseteq> F i"
    57   shows "indep_sets G I"
    58 proof -
    59   have "(\<forall>i\<in>I. F i \<subseteq> events) \<Longrightarrow> (\<forall>i\<in>I. G i \<subseteq> events)"
    60     using mono by auto
    61   moreover have "\<And>A J. J \<subseteq> I \<Longrightarrow> A \<in> (\<Pi> j\<in>J. G j) \<Longrightarrow> A \<in> (\<Pi> j\<in>J. F j)"
    62     using mono by (auto simp: Pi_iff)
    63   ultimately show ?thesis
    64     using indep by (auto simp: indep_sets_def)
    65 qed
    66 
    67 lemma (in prob_space) indep_setsI:
    68   assumes "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> events"
    69     and "\<And>A J. J \<noteq> {} \<Longrightarrow> J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> (\<forall>j\<in>J. A j \<in> F j) \<Longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
    70   shows "indep_sets F I"
    71   using assms unfolding indep_sets_def by (auto simp: Pi_iff)
    72 
    73 lemma (in prob_space) indep_setsD:
    74   assumes "indep_sets F I" and "J \<subseteq> I" "J \<noteq> {}" "finite J" "\<forall>j\<in>J. A j \<in> F j"
    75   shows "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
    76   using assms unfolding indep_sets_def by auto
    77 
    78 lemma (in prob_space) indep_setI:
    79   assumes ev: "A \<subseteq> events" "B \<subseteq> events"
    80     and indep: "\<And>a b. a \<in> A \<Longrightarrow> b \<in> B \<Longrightarrow> prob (a \<inter> b) = prob a * prob b"
    81   shows "indep_set A B"
    82   unfolding indep_set_def
    83 proof (rule indep_setsI)
    84   fix F J assume "J \<noteq> {}" "J \<subseteq> UNIV"
    85     and F: "\<forall>j\<in>J. F j \<in> (case j of True \<Rightarrow> A | False \<Rightarrow> B)"
    86   have "J \<in> Pow UNIV" by auto
    87   with F `J \<noteq> {}` indep[of "F True" "F False"]
    88   show "prob (\<Inter>j\<in>J. F j) = (\<Prod>j\<in>J. prob (F j))"
    89     unfolding UNIV_bool Pow_insert by (auto simp: ac_simps)
    90 qed (auto split: bool.split simp: ev)
    91 
    92 lemma (in prob_space) indep_setD:
    93   assumes indep: "indep_set A B" and ev: "a \<in> A" "b \<in> B"
    94   shows "prob (a \<inter> b) = prob a * prob b"
    95   using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "bool_case a b"] ev
    96   by (simp add: ac_simps UNIV_bool)
    97 
    98 lemma (in prob_space)
    99   assumes indep: "indep_set A B"
   100   shows indep_setD_ev1: "A \<subseteq> sets M"
   101     and indep_setD_ev2: "B \<subseteq> sets M"
   102   using indep unfolding indep_set_def indep_sets_def UNIV_bool by auto
   103 
   104 lemma dynkin_systemI':
   105   assumes 1: "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M"
   106   assumes empty: "{} \<in> sets M"
   107   assumes Diff: "\<And> A. A \<in> sets M \<Longrightarrow> space M - A \<in> sets M"
   108   assumes 2: "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> sets M
   109           \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
   110   shows "dynkin_system M"
   111 proof -
   112   from Diff[OF empty] have "space M \<in> sets M" by auto
   113   from 1 this Diff 2 show ?thesis
   114     by (intro dynkin_systemI) auto
   115 qed
   116 
   117 lemma (in prob_space) indep_sets_dynkin:
   118   assumes indep: "indep_sets F I"
   119   shows "indep_sets (\<lambda>i. sets (dynkin \<lparr> space = space M, sets = F i \<rparr>)) I"
   120     (is "indep_sets ?F I")
   121 proof (subst indep_sets_finite_index_sets, intro allI impI ballI)
   122   fix J assume "finite J" "J \<subseteq> I" "J \<noteq> {}"
   123   with indep have "indep_sets F J"
   124     by (subst (asm) indep_sets_finite_index_sets) auto
   125   { fix J K assume "indep_sets F K"
   126     let "?G S i" = "if i \<in> S then ?F i else F i"
   127     assume "finite J" "J \<subseteq> K"
   128     then have "indep_sets (?G J) K"
   129     proof induct
   130       case (insert j J)
   131       moreover def G \<equiv> "?G J"
   132       ultimately have G: "indep_sets G K" "\<And>i. i \<in> K \<Longrightarrow> G i \<subseteq> events" and "j \<in> K"
   133         by (auto simp: indep_sets_def)
   134       let ?D = "{E\<in>events. indep_sets (G(j := {E})) K }"
   135       { fix X assume X: "X \<in> events"
   136         assume indep: "\<And>J A. J \<noteq> {} \<Longrightarrow> J \<subseteq> K \<Longrightarrow> finite J \<Longrightarrow> j \<notin> J \<Longrightarrow> (\<forall>i\<in>J. A i \<in> G i)
   137           \<Longrightarrow> prob ((\<Inter>i\<in>J. A i) \<inter> X) = prob X * (\<Prod>i\<in>J. prob (A i))"
   138         have "indep_sets (G(j := {X})) K"
   139         proof (rule indep_setsI)
   140           fix i assume "i \<in> K" then show "(G(j:={X})) i \<subseteq> events"
   141             using G X by auto
   142         next
   143           fix A J assume J: "J \<noteq> {}" "J \<subseteq> K" "finite J" "\<forall>i\<in>J. A i \<in> (G(j := {X})) i"
   144           show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
   145           proof cases
   146             assume "j \<in> J"
   147             with J have "A j = X" by auto
   148             show ?thesis
   149             proof cases
   150               assume "J = {j}" then show ?thesis by simp
   151             next
   152               assume "J \<noteq> {j}"
   153               have "prob (\<Inter>i\<in>J. A i) = prob ((\<Inter>i\<in>J-{j}. A i) \<inter> X)"
   154                 using `j \<in> J` `A j = X` by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
   155               also have "\<dots> = prob X * (\<Prod>i\<in>J-{j}. prob (A i))"
   156               proof (rule indep)
   157                 show "J - {j} \<noteq> {}" "J - {j} \<subseteq> K" "finite (J - {j})" "j \<notin> J - {j}"
   158                   using J `J \<noteq> {j}` `j \<in> J` by auto
   159                 show "\<forall>i\<in>J - {j}. A i \<in> G i"
   160                   using J by auto
   161               qed
   162               also have "\<dots> = prob (A j) * (\<Prod>i\<in>J-{j}. prob (A i))"
   163                 using `A j = X` by simp
   164               also have "\<dots> = (\<Prod>i\<in>J. prob (A i))"
   165                 unfolding setprod.insert_remove[OF `finite J`, symmetric, of "\<lambda>i. prob  (A i)"]
   166                 using `j \<in> J` by (simp add: insert_absorb)
   167               finally show ?thesis .
   168             qed
   169           next
   170             assume "j \<notin> J"
   171             with J have "\<forall>i\<in>J. A i \<in> G i" by (auto split: split_if_asm)
   172             with J show ?thesis
   173               by (intro indep_setsD[OF G(1)]) auto
   174           qed
   175         qed }
   176       note indep_sets_insert = this
   177       have "dynkin_system \<lparr> space = space M, sets = ?D \<rparr>"
   178       proof (rule dynkin_systemI', simp_all, safe)
   179         show "indep_sets (G(j := {{}})) K"
   180           by (rule indep_sets_insert) auto
   181       next
   182         fix X assume X: "X \<in> events" and G': "indep_sets (G(j := {X})) K"
   183         show "indep_sets (G(j := {space M - X})) K"
   184         proof (rule indep_sets_insert)
   185           fix J A assume J: "J \<noteq> {}" "J \<subseteq> K" "finite J" "j \<notin> J" and A: "\<forall>i\<in>J. A i \<in> G i"
   186           then have A_sets: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> events"
   187             using G by auto
   188           have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) =
   189               prob ((\<Inter>j\<in>J. A j) - (\<Inter>i\<in>insert j J. (A(j := X)) i))"
   190             using A_sets sets_into_space X `J \<noteq> {}`
   191             by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
   192           also have "\<dots> = prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)"
   193             using J `J \<noteq> {}` `j \<notin> J` A_sets X sets_into_space
   194             by (auto intro!: finite_measure_Diff finite_INT split: split_if_asm)
   195           finally have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) =
   196               prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)" .
   197           moreover {
   198             have "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
   199               using J A `finite J` by (intro indep_setsD[OF G(1)]) auto
   200             then have "prob (\<Inter>j\<in>J. A j) = prob (space M) * (\<Prod>i\<in>J. prob (A i))"
   201               using prob_space by simp }
   202           moreover {
   203             have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = (\<Prod>i\<in>insert j J. prob ((A(j := X)) i))"
   204               using J A `j \<in> K` by (intro indep_setsD[OF G']) auto
   205             then have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = prob X * (\<Prod>i\<in>J. prob (A i))"
   206               using `finite J` `j \<notin> J` by (auto intro!: setprod_cong) }
   207           ultimately have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = (prob (space M) - prob X) * (\<Prod>i\<in>J. prob (A i))"
   208             by (simp add: field_simps)
   209           also have "\<dots> = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))"
   210             using X A by (simp add: finite_measure_compl)
   211           finally show "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))" .
   212         qed (insert X, auto)
   213       next
   214         fix F :: "nat \<Rightarrow> 'a set" assume disj: "disjoint_family F" and "range F \<subseteq> ?D"
   215         then have F: "\<And>i. F i \<in> events" "\<And>i. indep_sets (G(j:={F i})) K" by auto
   216         show "indep_sets (G(j := {\<Union>k. F k})) K"
   217         proof (rule indep_sets_insert)
   218           fix J A assume J: "j \<notin> J" "J \<noteq> {}" "J \<subseteq> K" "finite J" and A: "\<forall>i\<in>J. A i \<in> G i"
   219           then have A_sets: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> events"
   220             using G by auto
   221           have "prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)) = prob (\<Union>k. (\<Inter>i\<in>insert j J. (A(j := F k)) i))"
   222             using `J \<noteq> {}` `j \<notin> J` `j \<in> K` by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
   223           moreover have "(\<lambda>k. prob (\<Inter>i\<in>insert j J. (A(j := F k)) i)) sums prob (\<Union>k. (\<Inter>i\<in>insert j J. (A(j := F k)) i))"
   224           proof (rule finite_measure_UNION)
   225             show "disjoint_family (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i)"
   226               using disj by (rule disjoint_family_on_bisimulation) auto
   227             show "range (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i) \<subseteq> events"
   228               using A_sets F `finite J` `J \<noteq> {}` `j \<notin> J` by (auto intro!: Int)
   229           qed
   230           moreover { fix k
   231             from J A `j \<in> K` have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * (\<Prod>i\<in>J. prob (A i))"
   232               by (subst indep_setsD[OF F(2)]) (auto intro!: setprod_cong split: split_if_asm)
   233             also have "\<dots> = prob (F k) * prob (\<Inter>i\<in>J. A i)"
   234               using J A `j \<in> K` by (subst indep_setsD[OF G(1)]) auto
   235             finally have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * prob (\<Inter>i\<in>J. A i)" . }
   236           ultimately have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)))"
   237             by simp
   238           moreover
   239           have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob (\<Union>k. F k) * prob (\<Inter>i\<in>J. A i))"
   240             using disj F(1) by (intro finite_measure_UNION sums_mult2) auto
   241           then have "(\<lambda>k. prob (F k) * prob (\<Inter>i\<in>J. A i)) sums (prob (\<Union>k. F k) * (\<Prod>i\<in>J. prob (A i)))"
   242             using J A `j \<in> K` by (subst indep_setsD[OF G(1), symmetric]) auto
   243           ultimately
   244           show "prob ((\<Inter>j\<in>J. A j) \<inter> (\<Union>k. F k)) = prob (\<Union>k. F k) * (\<Prod>j\<in>J. prob (A j))"
   245             by (auto dest!: sums_unique)
   246         qed (insert F, auto)
   247       qed (insert sets_into_space, auto)
   248       then have mono: "sets (dynkin \<lparr>space = space M, sets = G j\<rparr>) \<subseteq>
   249         sets \<lparr>space = space M, sets = {E \<in> events. indep_sets (G(j := {E})) K}\<rparr>"
   250       proof (rule dynkin_system.dynkin_subset, simp_all, safe)
   251         fix X assume "X \<in> G j"
   252         then show "X \<in> events" using G `j \<in> K` by auto
   253         from `indep_sets G K`
   254         show "indep_sets (G(j := {X})) K"
   255           by (rule indep_sets_mono_sets) (insert `X \<in> G j`, auto)
   256       qed
   257       have "indep_sets (G(j:=?D)) K"
   258       proof (rule indep_setsI)
   259         fix i assume "i \<in> K" then show "(G(j := ?D)) i \<subseteq> events"
   260           using G(2) by auto
   261       next
   262         fix A J assume J: "J\<noteq>{}" "J \<subseteq> K" "finite J" and A: "\<forall>i\<in>J. A i \<in> (G(j := ?D)) i"
   263         show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
   264         proof cases
   265           assume "j \<in> J"
   266           with A have indep: "indep_sets (G(j := {A j})) K" by auto
   267           from J A show ?thesis
   268             by (intro indep_setsD[OF indep]) auto
   269         next
   270           assume "j \<notin> J"
   271           with J A have "\<forall>i\<in>J. A i \<in> G i" by (auto split: split_if_asm)
   272           with J show ?thesis
   273             by (intro indep_setsD[OF G(1)]) auto
   274         qed
   275       qed
   276       then have "indep_sets (G(j:=sets (dynkin \<lparr>space = space M, sets = G j\<rparr>))) K"
   277         by (rule indep_sets_mono_sets) (insert mono, auto)
   278       then show ?case
   279         by (rule indep_sets_mono_sets) (insert `j \<in> K` `j \<notin> J`, auto simp: G_def)
   280     qed (insert `indep_sets F K`, simp) }
   281   from this[OF `indep_sets F J` `finite J` subset_refl]
   282   show "indep_sets (\<lambda>i. sets (dynkin \<lparr> space = space M, sets = F i \<rparr>)) J"
   283     by (rule indep_sets_mono_sets) auto
   284 qed
   285 
   286 lemma (in prob_space) indep_sets_sigma:
   287   assumes indep: "indep_sets F I"
   288   assumes stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
   289   shows "indep_sets (\<lambda>i. sets (sigma \<lparr> space = space M, sets = F i \<rparr>)) I"
   290 proof -
   291   from indep_sets_dynkin[OF indep]
   292   show ?thesis
   293   proof (rule indep_sets_mono_sets, subst sigma_eq_dynkin, simp_all add: stable)
   294     fix i assume "i \<in> I"
   295     with indep have "F i \<subseteq> events" by (auto simp: indep_sets_def)
   296     with sets_into_space show "F i \<subseteq> Pow (space M)" by auto
   297   qed
   298 qed
   299 
   300 lemma (in prob_space) indep_sets_sigma_sets:
   301   assumes "indep_sets F I"
   302   assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
   303   shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I"
   304   using indep_sets_sigma[OF assms] by (simp add: sets_sigma)
   305 
   306 lemma (in prob_space) indep_sets2_eq:
   307   "indep_set A B \<longleftrightarrow> A \<subseteq> events \<and> B \<subseteq> events \<and> (\<forall>a\<in>A. \<forall>b\<in>B. prob (a \<inter> b) = prob a * prob b)"
   308   unfolding indep_set_def
   309 proof (intro iffI ballI conjI)
   310   assume indep: "indep_sets (bool_case A B) UNIV"
   311   { fix a b assume "a \<in> A" "b \<in> B"
   312     with indep_setsD[OF indep, of UNIV "bool_case a b"]
   313     show "prob (a \<inter> b) = prob a * prob b"
   314       unfolding UNIV_bool by (simp add: ac_simps) }
   315   from indep show "A \<subseteq> events" "B \<subseteq> events"
   316     unfolding indep_sets_def UNIV_bool by auto
   317 next
   318   assume *: "A \<subseteq> events \<and> B \<subseteq> events \<and> (\<forall>a\<in>A. \<forall>b\<in>B. prob (a \<inter> b) = prob a * prob b)"
   319   show "indep_sets (bool_case A B) UNIV"
   320   proof (rule indep_setsI)
   321     fix i show "(case i of True \<Rightarrow> A | False \<Rightarrow> B) \<subseteq> events"
   322       using * by (auto split: bool.split)
   323   next
   324     fix J X assume "J \<noteq> {}" "J \<subseteq> UNIV" and X: "\<forall>j\<in>J. X j \<in> (case j of True \<Rightarrow> A | False \<Rightarrow> B)"
   325     then have "J = {True} \<or> J = {False} \<or> J = {True,False}"
   326       by (auto simp: UNIV_bool)
   327     then show "prob (\<Inter>j\<in>J. X j) = (\<Prod>j\<in>J. prob (X j))"
   328       using X * by auto
   329   qed
   330 qed
   331 
   332 lemma (in prob_space) indep_set_sigma_sets:
   333   assumes "indep_set A B"
   334   assumes A: "Int_stable \<lparr> space = space M, sets = A \<rparr>"
   335   assumes B: "Int_stable \<lparr> space = space M, sets = B \<rparr>"
   336   shows "indep_set (sigma_sets (space M) A) (sigma_sets (space M) B)"
   337 proof -
   338   have "indep_sets (\<lambda>i. sigma_sets (space M) (case i of True \<Rightarrow> A | False \<Rightarrow> B)) UNIV"
   339   proof (rule indep_sets_sigma_sets)
   340     show "indep_sets (bool_case A B) UNIV"
   341       by (rule `indep_set A B`[unfolded indep_set_def])
   342     fix i show "Int_stable \<lparr>space = space M, sets = case i of True \<Rightarrow> A | False \<Rightarrow> B\<rparr>"
   343       using A B by (cases i) auto
   344   qed
   345   then show ?thesis
   346     unfolding indep_set_def
   347     by (rule indep_sets_mono_sets) (auto split: bool.split)
   348 qed
   349 
   350 lemma (in prob_space) indep_sets_collect_sigma:
   351   fixes I :: "'j \<Rightarrow> 'i set" and J :: "'j set" and E :: "'i \<Rightarrow> 'a set set"
   352   assumes indep: "indep_sets E (\<Union>j\<in>J. I j)"
   353   assumes Int_stable: "\<And>i j. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> Int_stable \<lparr>space = space M, sets = E i\<rparr>"
   354   assumes disjoint: "disjoint_family_on I J"
   355   shows "indep_sets (\<lambda>j. sigma_sets (space M) (\<Union>i\<in>I j. E i)) J"
   356 proof -
   357   let "?E j" = "{\<Inter>k\<in>K. E' k| E' K. finite K \<and> K \<noteq> {} \<and> K \<subseteq> I j \<and> (\<forall>k\<in>K. E' k \<in> E k) }"
   358 
   359   from indep have E: "\<And>j i. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> E i \<subseteq> sets M"
   360     unfolding indep_sets_def by auto
   361   { fix j
   362     let ?S = "sigma \<lparr> space = space M, sets = (\<Union>i\<in>I j. E i) \<rparr>"
   363     assume "j \<in> J"
   364     from E[OF this] interpret S: sigma_algebra ?S
   365       using sets_into_space by (intro sigma_algebra_sigma) auto
   366 
   367     have "sigma_sets (space M) (\<Union>i\<in>I j. E i) = sigma_sets (space M) (?E j)"
   368     proof (rule sigma_sets_eqI)
   369       fix A assume "A \<in> (\<Union>i\<in>I j. E i)"
   370       then guess i ..
   371       then show "A \<in> sigma_sets (space M) (?E j)"
   372         by (auto intro!: sigma_sets.intros exI[of _ "{i}"] exI[of _ "\<lambda>i. A"])
   373     next
   374       fix A assume "A \<in> ?E j"
   375       then obtain E' K where "finite K" "K \<noteq> {}" "K \<subseteq> I j" "\<And>k. k \<in> K \<Longrightarrow> E' k \<in> E k"
   376         and A: "A = (\<Inter>k\<in>K. E' k)"
   377         by auto
   378       then have "A \<in> sets ?S" unfolding A
   379         by (safe intro!: S.finite_INT)
   380            (auto simp: sets_sigma intro!: sigma_sets.Basic)
   381       then show "A \<in> sigma_sets (space M) (\<Union>i\<in>I j. E i)"
   382         by (simp add: sets_sigma)
   383     qed }
   384   moreover have "indep_sets (\<lambda>j. sigma_sets (space M) (?E j)) J"
   385   proof (rule indep_sets_sigma_sets)
   386     show "indep_sets ?E J"
   387     proof (intro indep_setsI)
   388       fix j assume "j \<in> J" with E show "?E j \<subseteq> events" by (force  intro!: finite_INT)
   389     next
   390       fix K A assume K: "K \<noteq> {}" "K \<subseteq> J" "finite K"
   391         and "\<forall>j\<in>K. A j \<in> ?E j"
   392       then have "\<forall>j\<in>K. \<exists>E' L. A j = (\<Inter>l\<in>L. E' l) \<and> finite L \<and> L \<noteq> {} \<and> L \<subseteq> I j \<and> (\<forall>l\<in>L. E' l \<in> E l)"
   393         by simp
   394       from bchoice[OF this] guess E' ..
   395       from bchoice[OF this] obtain L
   396         where A: "\<And>j. j\<in>K \<Longrightarrow> A j = (\<Inter>l\<in>L j. E' j l)"
   397         and L: "\<And>j. j\<in>K \<Longrightarrow> finite (L j)" "\<And>j. j\<in>K \<Longrightarrow> L j \<noteq> {}" "\<And>j. j\<in>K \<Longrightarrow> L j \<subseteq> I j"
   398         and E': "\<And>j l. j\<in>K \<Longrightarrow> l \<in> L j \<Longrightarrow> E' j l \<in> E l"
   399         by auto
   400 
   401       { fix k l j assume "k \<in> K" "j \<in> K" "l \<in> L j" "l \<in> L k"
   402         have "k = j"
   403         proof (rule ccontr)
   404           assume "k \<noteq> j"
   405           with disjoint `K \<subseteq> J` `k \<in> K` `j \<in> K` have "I k \<inter> I j = {}"
   406             unfolding disjoint_family_on_def by auto
   407           with L(2,3)[OF `j \<in> K`] L(2,3)[OF `k \<in> K`]
   408           show False using `l \<in> L k` `l \<in> L j` by auto
   409         qed }
   410       note L_inj = this
   411 
   412       def k \<equiv> "\<lambda>l. (SOME k. k \<in> K \<and> l \<in> L k)"
   413       { fix x j l assume *: "j \<in> K" "l \<in> L j"
   414         have "k l = j" unfolding k_def
   415         proof (rule some_equality)
   416           fix k assume "k \<in> K \<and> l \<in> L k"
   417           with * L_inj show "k = j" by auto
   418         qed (insert *, simp) }
   419       note k_simp[simp] = this
   420       let "?E' l" = "E' (k l) l"
   421       have "prob (\<Inter>j\<in>K. A j) = prob (\<Inter>l\<in>(\<Union>k\<in>K. L k). ?E' l)"
   422         by (auto simp: A intro!: arg_cong[where f=prob])
   423       also have "\<dots> = (\<Prod>l\<in>(\<Union>k\<in>K. L k). prob (?E' l))"
   424         using L K E' by (intro indep_setsD[OF indep]) (simp_all add: UN_mono)
   425       also have "\<dots> = (\<Prod>j\<in>K. \<Prod>l\<in>L j. prob (E' j l))"
   426         using K L L_inj by (subst setprod_UN_disjoint) auto
   427       also have "\<dots> = (\<Prod>j\<in>K. prob (A j))"
   428         using K L E' by (auto simp add: A intro!: setprod_cong indep_setsD[OF indep, symmetric]) blast
   429       finally show "prob (\<Inter>j\<in>K. A j) = (\<Prod>j\<in>K. prob (A j))" .
   430     qed
   431   next
   432     fix j assume "j \<in> J"
   433     show "Int_stable \<lparr> space = space M, sets = ?E j \<rparr>"
   434     proof (rule Int_stableI)
   435       fix a assume "a \<in> ?E j" then obtain Ka Ea
   436         where a: "a = (\<Inter>k\<in>Ka. Ea k)" "finite Ka" "Ka \<noteq> {}" "Ka \<subseteq> I j" "\<And>k. k\<in>Ka \<Longrightarrow> Ea k \<in> E k" by auto
   437       fix b assume "b \<in> ?E j" then obtain Kb Eb
   438         where b: "b = (\<Inter>k\<in>Kb. Eb k)" "finite Kb" "Kb \<noteq> {}" "Kb \<subseteq> I j" "\<And>k. k\<in>Kb \<Longrightarrow> Eb k \<in> E k" by auto
   439       let ?A = "\<lambda>k. (if k \<in> Ka \<inter> Kb then Ea k \<inter> Eb k else if k \<in> Kb then Eb k else if k \<in> Ka then Ea k else {})"
   440       have "a \<inter> b = INTER (Ka \<union> Kb) ?A"
   441         by (simp add: a b set_eq_iff) auto
   442       with a b `j \<in> J` Int_stableD[OF Int_stable] show "a \<inter> b \<in> ?E j"
   443         by (intro CollectI exI[of _ "Ka \<union> Kb"] exI[of _ ?A]) auto
   444     qed
   445   qed
   446   ultimately show ?thesis
   447     by (simp cong: indep_sets_cong)
   448 qed
   449 
   450 definition (in prob_space) terminal_events where
   451   "terminal_events A = (\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
   452 
   453 lemma (in prob_space) terminal_events_sets:
   454   assumes A: "\<And>i. A i \<subseteq> sets M"
   455   assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
   456   assumes X: "X \<in> terminal_events A"
   457   shows "X \<in> sets M"
   458 proof -
   459   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
   460   interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
   461   from X have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by (auto simp: terminal_events_def)
   462   from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
   463   then show "X \<in> sets M"
   464     by induct (insert A, auto)
   465 qed
   466 
   467 lemma (in prob_space) sigma_algebra_terminal_events:
   468   assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
   469   shows "sigma_algebra \<lparr> space = space M, sets = terminal_events A \<rparr>"
   470   unfolding terminal_events_def
   471 proof (simp add: sigma_algebra_iff2, safe)
   472   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
   473   interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
   474   { fix X x assume "X \<in> ?A" "x \<in> X" 
   475     then have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by auto
   476     from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
   477     then have "X \<subseteq> space M"
   478       by induct (insert A.sets_into_space, auto)
   479     with `x \<in> X` show "x \<in> space M" by auto }
   480   { fix F :: "nat \<Rightarrow> 'a set" and n assume "range F \<subseteq> ?A"
   481     then show "(UNION UNIV F) \<in> sigma_sets (space M) (UNION {n..} A)"
   482       by (intro sigma_sets.Union) auto }
   483 qed (auto intro!: sigma_sets.Compl sigma_sets.Empty)
   484 
   485 lemma (in prob_space) kolmogorov_0_1_law:
   486   fixes A :: "nat \<Rightarrow> 'a set set"
   487   assumes A: "\<And>i. A i \<subseteq> sets M"
   488   assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
   489   assumes indep: "indep_sets A UNIV"
   490   and X: "X \<in> terminal_events A"
   491   shows "prob X = 0 \<or> prob X = 1"
   492 proof -
   493   let ?D = "\<lparr> space = space M, sets = {D \<in> sets M. prob (X \<inter> D) = prob X * prob D} \<rparr>"
   494   interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
   495   interpret T: sigma_algebra "\<lparr> space = space M, sets = terminal_events A \<rparr>"
   496     by (rule sigma_algebra_terminal_events) fact
   497   have "X \<subseteq> space M" using T.space_closed X by auto
   498 
   499   have X_in: "X \<in> sets M"
   500     by (rule terminal_events_sets) fact+
   501 
   502   interpret D: dynkin_system ?D
   503   proof (rule dynkin_systemI)
   504     fix D assume "D \<in> sets ?D" then show "D \<subseteq> space ?D"
   505       using sets_into_space by auto
   506   next
   507     show "space ?D \<in> sets ?D"
   508       using prob_space `X \<subseteq> space M` by (simp add: Int_absorb2)
   509   next
   510     fix A assume A: "A \<in> sets ?D"
   511     have "prob (X \<inter> (space M - A)) = prob (X - (X \<inter> A))"
   512       using `X \<subseteq> space M` by (auto intro!: arg_cong[where f=prob])
   513     also have "\<dots> = prob X - prob (X \<inter> A)"
   514       using X_in A by (intro finite_measure_Diff) auto
   515     also have "\<dots> = prob X * prob (space M) - prob X * prob A"
   516       using A prob_space by auto
   517     also have "\<dots> = prob X * prob (space M - A)"
   518       using X_in A sets_into_space
   519       by (subst finite_measure_Diff) (auto simp: field_simps)
   520     finally show "space ?D - A \<in> sets ?D"
   521       using A `X \<subseteq> space M` by auto
   522   next
   523     fix F :: "nat \<Rightarrow> 'a set" assume dis: "disjoint_family F" and "range F \<subseteq> sets ?D"
   524     then have F: "range F \<subseteq> events" "\<And>i. prob (X \<inter> F i) = prob X * prob (F i)"
   525       by auto
   526     have "(\<lambda>i. prob (X \<inter> F i)) sums prob (\<Union>i. X \<inter> F i)"
   527     proof (rule finite_measure_UNION)
   528       show "range (\<lambda>i. X \<inter> F i) \<subseteq> events"
   529         using F X_in by auto
   530       show "disjoint_family (\<lambda>i. X \<inter> F i)"
   531         using dis by (rule disjoint_family_on_bisimulation) auto
   532     qed
   533     with F have "(\<lambda>i. prob X * prob (F i)) sums prob (X \<inter> (\<Union>i. F i))"
   534       by simp
   535     moreover have "(\<lambda>i. prob X * prob (F i)) sums (prob X * prob (\<Union>i. F i))"
   536       by (intro mult_right.sums finite_measure_UNION F dis)
   537     ultimately have "prob (X \<inter> (\<Union>i. F i)) = prob X * prob (\<Union>i. F i)"
   538       by (auto dest!: sums_unique)
   539     with F show "(\<Union>i. F i) \<in> sets ?D"
   540       by auto
   541   qed
   542 
   543   { fix n
   544     have "indep_sets (\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) UNIV"
   545     proof (rule indep_sets_collect_sigma)
   546       have *: "(\<Union>b. case b of True \<Rightarrow> {..n} | False \<Rightarrow> {Suc n..}) = UNIV" (is "?U = _")
   547         by (simp split: bool.split add: set_eq_iff) (metis not_less_eq_eq)
   548       with indep show "indep_sets A ?U" by simp
   549       show "disjoint_family (bool_case {..n} {Suc n..})"
   550         unfolding disjoint_family_on_def by (auto split: bool.split)
   551       fix m
   552       show "Int_stable \<lparr>space = space M, sets = A m\<rparr>"
   553         unfolding Int_stable_def using A.Int by auto
   554     qed
   555     also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) = 
   556       bool_case (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
   557       by (auto intro!: ext split: bool.split)
   558     finally have indep: "indep_set (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
   559       unfolding indep_set_def by simp
   560 
   561     have "sigma_sets (space M) (\<Union>m\<in>{..n}. A m) \<subseteq> sets ?D"
   562     proof (simp add: subset_eq, rule)
   563       fix D assume D: "D \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
   564       have "X \<in> sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m)"
   565         using X unfolding terminal_events_def by simp
   566       from indep_setD[OF indep D this] indep_setD_ev1[OF indep] D
   567       show "D \<in> events \<and> prob (X \<inter> D) = prob X * prob D"
   568         by (auto simp add: ac_simps)
   569     qed }
   570   then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> sets ?D" (is "?A \<subseteq> _")
   571     by auto
   572 
   573   have "sigma \<lparr> space = space M, sets = ?A \<rparr> =
   574     dynkin \<lparr> space = space M, sets = ?A \<rparr>" (is "sigma ?UA = dynkin ?UA")
   575   proof (rule sigma_eq_dynkin)
   576     { fix B n assume "B \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
   577       then have "B \<subseteq> space M"
   578         by induct (insert A sets_into_space, auto) }
   579     then show "sets ?UA \<subseteq> Pow (space ?UA)" by auto
   580     show "Int_stable ?UA"
   581     proof (rule Int_stableI)
   582       fix a assume "a \<in> ?A" then guess n .. note a = this
   583       fix b assume "b \<in> ?A" then guess m .. note b = this
   584       interpret Amn: sigma_algebra "sigma \<lparr>space = space M, sets = (\<Union>i\<in>{..max m n}. A i)\<rparr>"
   585         using A sets_into_space by (intro sigma_algebra_sigma) auto
   586       have "sigma_sets (space M) (\<Union>i\<in>{..n}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
   587         by (intro sigma_sets_subseteq UN_mono) auto
   588       with a have "a \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
   589       moreover
   590       have "sigma_sets (space M) (\<Union>i\<in>{..m}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
   591         by (intro sigma_sets_subseteq UN_mono) auto
   592       with b have "b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
   593       ultimately have "a \<inter> b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
   594         using Amn.Int[of a b] by (simp add: sets_sigma)
   595       then show "a \<inter> b \<in> (\<Union>n. sigma_sets (space M) (\<Union>i\<in>{..n}. A i))" by auto
   596     qed
   597   qed
   598   moreover have "sets (dynkin ?UA) \<subseteq> sets ?D"
   599   proof (rule D.dynkin_subset)
   600     show "sets ?UA \<subseteq> sets ?D" using `?A \<subseteq> sets ?D` by auto
   601   qed simp
   602   ultimately have "sets (sigma ?UA) \<subseteq> sets ?D" by simp
   603   moreover
   604   have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> sigma_sets (space M) ?A"
   605     by (intro sigma_sets_subseteq UN_mono) (auto intro: sigma_sets.Basic)
   606   then have "terminal_events A \<subseteq> sets (sigma ?UA)"
   607     unfolding sets_sigma terminal_events_def by auto
   608   moreover note `X \<in> terminal_events A`
   609   ultimately have "X \<in> sets ?D" by auto
   610   then show ?thesis by auto
   611 qed
   612 
   613 end