src/HOL/Probability/Independent_Family.thy
author noschinl
Mon, 19 Dec 2011 14:41:08 +0100
changeset 45931 99cf6e470816
parent 45777 c36637603821
child 46731 5302e932d1e5
permissions -rw-r--r--
weaken preconditions on lemmas
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     1
(*  Title:      HOL/Probability/Independent_Family.thy
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     2
    Author:     Johannes Hölzl, TU München
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     3
*)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     4
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     5
header {* Independent families of events, event sets, and random variables *}
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     6
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     7
theory Independent_Family
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     8
  imports Probability_Measure
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
     9
begin
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    10
42985
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    11
lemma INT_decseq_offset:
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    12
  assumes "decseq F"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    13
  shows "(\<Inter>i. F i) = (\<Inter>i\<in>{n..}. F i)"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    14
proof safe
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    15
  fix x i assume x: "x \<in> (\<Inter>i\<in>{n..}. F i)"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    16
  show "x \<in> F i"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    17
  proof cases
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    18
    from x have "x \<in> F n" by auto
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    19
    also assume "i \<le> n" with `decseq F` have "F n \<subseteq> F i"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    20
      unfolding decseq_def by simp
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    21
    finally show ?thesis .
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    22
  qed (insert x, simp)
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    23
qed auto
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    24
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    25
definition (in prob_space)
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
    26
  "indep_events A I \<longleftrightarrow> (A`I \<subseteq> events) \<and>
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    27
    (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j)))"
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    28
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    29
definition (in prob_space)
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    30
  "indep_event A B \<longleftrightarrow> indep_events (bool_case A B) UNIV"
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    31
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    32
definition (in prob_space)
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
    33
  "indep_sets F I \<longleftrightarrow> (\<forall>i\<in>I. F i \<subseteq> events) \<and>
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    34
    (\<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))))"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    35
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    36
definition (in prob_space)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    37
  "indep_set A B \<longleftrightarrow> indep_sets (bool_case A B) UNIV"
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    38
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    39
definition (in prob_space)
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
    40
  "indep_vars M' X I \<longleftrightarrow>
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    41
    (\<forall>i\<in>I. random_variable (M' i) (X i)) \<and>
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    42
    indep_sets (\<lambda>i. sigma_sets (space M) { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    43
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
    44
definition (in prob_space)
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
    45
  "indep_var Ma A Mb B \<longleftrightarrow> indep_vars (bool_case Ma Mb) (bool_case A B) UNIV"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
    46
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
    47
lemma (in prob_space) indep_sets_cong[cong]:
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    48
  "I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i = G i) \<Longrightarrow> indep_sets F I \<longleftrightarrow> indep_sets G J"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    49
  by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    50
42985
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    51
lemma (in prob_space) indep_sets_singleton_iff_indep_events:
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    52
  "indep_sets (\<lambda>i. {F i}) I \<longleftrightarrow> indep_events F I"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    53
  unfolding indep_sets_def indep_events_def
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    54
  by (simp, intro conj_cong ball_cong all_cong imp_cong) (auto simp: Pi_iff)
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
    55
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    56
lemma (in prob_space) indep_events_finite_index_events:
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    57
  "indep_events F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_events F J)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    58
  by (auto simp: indep_events_def)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
    59
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    60
lemma (in prob_space) indep_sets_finite_index_sets:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    61
  "indep_sets F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_sets F J)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    62
proof (intro iffI allI impI)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    63
  assume *: "\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_sets F J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    64
  show "indep_sets F I" unfolding indep_sets_def
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    65
  proof (intro conjI ballI allI impI)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    66
    fix i assume "i \<in> I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    67
    with *[THEN spec, of "{i}"] show "F i \<subseteq> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    68
      by (auto simp: indep_sets_def)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    69
  qed (insert *, auto simp: indep_sets_def)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    70
qed (auto simp: indep_sets_def)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    71
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    72
lemma (in prob_space) indep_sets_mono_index:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    73
  "J \<subseteq> I \<Longrightarrow> indep_sets F I \<Longrightarrow> indep_sets F J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    74
  unfolding indep_sets_def by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    75
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    76
lemma (in prob_space) indep_sets_mono_sets:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    77
  assumes indep: "indep_sets F I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    78
  assumes mono: "\<And>i. i\<in>I \<Longrightarrow> G i \<subseteq> F i"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    79
  shows "indep_sets G I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    80
proof -
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    81
  have "(\<forall>i\<in>I. F i \<subseteq> events) \<Longrightarrow> (\<forall>i\<in>I. G i \<subseteq> events)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    82
    using mono by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    83
  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)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    84
    using mono by (auto simp: Pi_iff)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    85
  ultimately show ?thesis
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    86
    using indep by (auto simp: indep_sets_def)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    87
qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    88
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    89
lemma (in prob_space) indep_setsI:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    90
  assumes "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    91
    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))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    92
  shows "indep_sets F I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    93
  using assms unfolding indep_sets_def by (auto simp: Pi_iff)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    94
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    95
lemma (in prob_space) indep_setsD:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    96
  assumes "indep_sets F I" and "J \<subseteq> I" "J \<noteq> {}" "finite J" "\<forall>j\<in>J. A j \<in> F j"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    97
  shows "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    98
  using assms unfolding indep_sets_def by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
    99
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   100
lemma (in prob_space) indep_setI:
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   101
  assumes ev: "A \<subseteq> events" "B \<subseteq> events"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   102
    and indep: "\<And>a b. a \<in> A \<Longrightarrow> b \<in> B \<Longrightarrow> prob (a \<inter> b) = prob a * prob b"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   103
  shows "indep_set A B"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   104
  unfolding indep_set_def
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   105
proof (rule indep_setsI)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   106
  fix F J assume "J \<noteq> {}" "J \<subseteq> UNIV"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   107
    and F: "\<forall>j\<in>J. F j \<in> (case j of True \<Rightarrow> A | False \<Rightarrow> B)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   108
  have "J \<in> Pow UNIV" by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   109
  with F `J \<noteq> {}` indep[of "F True" "F False"]
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   110
  show "prob (\<Inter>j\<in>J. F j) = (\<Prod>j\<in>J. prob (F j))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   111
    unfolding UNIV_bool Pow_insert by (auto simp: ac_simps)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   112
qed (auto split: bool.split simp: ev)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   113
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   114
lemma (in prob_space) indep_setD:
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   115
  assumes indep: "indep_set A B" and ev: "a \<in> A" "b \<in> B"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   116
  shows "prob (a \<inter> b) = prob a * prob b"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   117
  using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "bool_case a b"] ev
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   118
  by (simp add: ac_simps UNIV_bool)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   119
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   120
lemma (in prob_space) indep_var_eq:
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   121
  "indep_var S X T Y \<longleftrightarrow>
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   122
    (random_variable S X \<and> random_variable T Y) \<and>
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   123
    indep_set
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   124
      (sigma_sets (space M) { X -` A \<inter> space M | A. A \<in> sets S})
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   125
      (sigma_sets (space M) { Y -` A \<inter> space M | A. A \<in> sets T})"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   126
  unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   127
  by (intro arg_cong2[where f="op \<and>"] arg_cong2[where f=indep_sets] ext)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   128
     (auto split: bool.split)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   129
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   130
lemma (in prob_space)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   131
  assumes indep: "indep_set A B"
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   132
  shows indep_setD_ev1: "A \<subseteq> events"
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   133
    and indep_setD_ev2: "B \<subseteq> events"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   134
  using indep unfolding indep_set_def indep_sets_def UNIV_bool by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   135
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   136
lemma (in prob_space) indep_sets_dynkin:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   137
  assumes indep: "indep_sets F I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   138
  shows "indep_sets (\<lambda>i. sets (dynkin \<lparr> space = space M, sets = F i \<rparr>)) I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   139
    (is "indep_sets ?F I")
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   140
proof (subst indep_sets_finite_index_sets, intro allI impI ballI)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   141
  fix J assume "finite J" "J \<subseteq> I" "J \<noteq> {}"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   142
  with indep have "indep_sets F J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   143
    by (subst (asm) indep_sets_finite_index_sets) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   144
  { fix J K assume "indep_sets F K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   145
    let "?G S i" = "if i \<in> S then ?F i else F i"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   146
    assume "finite J" "J \<subseteq> K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   147
    then have "indep_sets (?G J) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   148
    proof induct
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   149
      case (insert j J)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   150
      moreover def G \<equiv> "?G J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   151
      ultimately have G: "indep_sets G K" "\<And>i. i \<in> K \<Longrightarrow> G i \<subseteq> events" and "j \<in> K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   152
        by (auto simp: indep_sets_def)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   153
      let ?D = "{E\<in>events. indep_sets (G(j := {E})) K }"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   154
      { fix X assume X: "X \<in> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   155
        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)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   156
          \<Longrightarrow> prob ((\<Inter>i\<in>J. A i) \<inter> X) = prob X * (\<Prod>i\<in>J. prob (A i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   157
        have "indep_sets (G(j := {X})) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   158
        proof (rule indep_setsI)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   159
          fix i assume "i \<in> K" then show "(G(j:={X})) i \<subseteq> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   160
            using G X by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   161
        next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   162
          fix A J assume J: "J \<noteq> {}" "J \<subseteq> K" "finite J" "\<forall>i\<in>J. A i \<in> (G(j := {X})) i"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   163
          show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   164
          proof cases
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   165
            assume "j \<in> J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   166
            with J have "A j = X" by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   167
            show ?thesis
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   168
            proof cases
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   169
              assume "J = {j}" then show ?thesis by simp
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   170
            next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   171
              assume "J \<noteq> {j}"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   172
              have "prob (\<Inter>i\<in>J. A i) = prob ((\<Inter>i\<in>J-{j}. A i) \<inter> X)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   173
                using `j \<in> J` `A j = X` by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   174
              also have "\<dots> = prob X * (\<Prod>i\<in>J-{j}. prob (A i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   175
              proof (rule indep)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   176
                show "J - {j} \<noteq> {}" "J - {j} \<subseteq> K" "finite (J - {j})" "j \<notin> J - {j}"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   177
                  using J `J \<noteq> {j}` `j \<in> J` by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   178
                show "\<forall>i\<in>J - {j}. A i \<in> G i"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   179
                  using J by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   180
              qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   181
              also have "\<dots> = prob (A j) * (\<Prod>i\<in>J-{j}. prob (A i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   182
                using `A j = X` by simp
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   183
              also have "\<dots> = (\<Prod>i\<in>J. prob (A i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   184
                unfolding setprod.insert_remove[OF `finite J`, symmetric, of "\<lambda>i. prob  (A i)"]
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   185
                using `j \<in> J` by (simp add: insert_absorb)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   186
              finally show ?thesis .
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   187
            qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   188
          next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   189
            assume "j \<notin> J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   190
            with J have "\<forall>i\<in>J. A i \<in> G i" by (auto split: split_if_asm)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   191
            with J show ?thesis
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   192
              by (intro indep_setsD[OF G(1)]) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   193
          qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   194
        qed }
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   195
      note indep_sets_insert = this
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   196
      have "dynkin_system \<lparr> space = space M, sets = ?D \<rparr>"
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   197
      proof (rule dynkin_systemI', simp_all cong del: indep_sets_cong, safe)
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   198
        show "indep_sets (G(j := {{}})) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   199
          by (rule indep_sets_insert) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   200
      next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   201
        fix X assume X: "X \<in> events" and G': "indep_sets (G(j := {X})) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   202
        show "indep_sets (G(j := {space M - X})) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   203
        proof (rule indep_sets_insert)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   204
          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"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   205
          then have A_sets: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   206
            using G by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   207
          have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) =
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   208
              prob ((\<Inter>j\<in>J. A j) - (\<Inter>i\<in>insert j J. (A(j := X)) i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   209
            using A_sets sets_into_space X `J \<noteq> {}`
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   210
            by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   211
          also have "\<dots> = prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   212
            using J `J \<noteq> {}` `j \<notin> J` A_sets X sets_into_space
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   213
            by (auto intro!: finite_measure_Diff finite_INT split: split_if_asm)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   214
          finally have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) =
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   215
              prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)" .
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   216
          moreover {
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   217
            have "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   218
              using J A `finite J` by (intro indep_setsD[OF G(1)]) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   219
            then have "prob (\<Inter>j\<in>J. A j) = prob (space M) * (\<Prod>i\<in>J. prob (A i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   220
              using prob_space by simp }
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   221
          moreover {
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   222
            have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = (\<Prod>i\<in>insert j J. prob ((A(j := X)) i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   223
              using J A `j \<in> K` by (intro indep_setsD[OF G']) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   224
            then have "prob (\<Inter>i\<in>insert j J. (A(j := X)) i) = prob X * (\<Prod>i\<in>J. prob (A i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   225
              using `finite J` `j \<notin> J` by (auto intro!: setprod_cong) }
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   226
          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))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   227
            by (simp add: field_simps)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   228
          also have "\<dots> = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   229
            using X A by (simp add: finite_measure_compl)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   230
          finally show "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) = prob (space M - X) * (\<Prod>i\<in>J. prob (A i))" .
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   231
        qed (insert X, auto)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   232
      next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   233
        fix F :: "nat \<Rightarrow> 'a set" assume disj: "disjoint_family F" and "range F \<subseteq> ?D"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   234
        then have F: "\<And>i. F i \<in> events" "\<And>i. indep_sets (G(j:={F i})) K" by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   235
        show "indep_sets (G(j := {\<Union>k. F k})) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   236
        proof (rule indep_sets_insert)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   237
          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"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   238
          then have A_sets: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   239
            using G by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   240
          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))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   241
            using `J \<noteq> {}` `j \<notin> J` `j \<in> K` by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   242
          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))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   243
          proof (rule finite_measure_UNION)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   244
            show "disjoint_family (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   245
              using disj by (rule disjoint_family_on_bisimulation) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   246
            show "range (\<lambda>k. \<Inter>i\<in>insert j J. (A(j := F k)) i) \<subseteq> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   247
              using A_sets F `finite J` `J \<noteq> {}` `j \<notin> J` by (auto intro!: Int)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   248
          qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   249
          moreover { fix k
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   250
            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))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   251
              by (subst indep_setsD[OF F(2)]) (auto intro!: setprod_cong split: split_if_asm)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   252
            also have "\<dots> = prob (F k) * prob (\<Inter>i\<in>J. A i)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   253
              using J A `j \<in> K` by (subst indep_setsD[OF G(1)]) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   254
            finally have "prob (\<Inter>i\<in>insert j J. (A(j := F k)) i) = prob (F k) * prob (\<Inter>i\<in>J. A i)" . }
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   255
          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)))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   256
            by simp
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   257
          moreover
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   258
          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))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   259
            using disj F(1) by (intro finite_measure_UNION sums_mult2) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   260
          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)))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   261
            using J A `j \<in> K` by (subst indep_setsD[OF G(1), symmetric]) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   262
          ultimately
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   263
          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))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   264
            by (auto dest!: sums_unique)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   265
        qed (insert F, auto)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   266
      qed (insert sets_into_space, auto)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   267
      then have mono: "sets (dynkin \<lparr>space = space M, sets = G j\<rparr>) \<subseteq>
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   268
        sets \<lparr>space = space M, sets = {E \<in> events. indep_sets (G(j := {E})) K}\<rparr>"
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   269
      proof (rule dynkin_system.dynkin_subset, simp_all cong del: indep_sets_cong, safe)
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   270
        fix X assume "X \<in> G j"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   271
        then show "X \<in> events" using G `j \<in> K` by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   272
        from `indep_sets G K`
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   273
        show "indep_sets (G(j := {X})) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   274
          by (rule indep_sets_mono_sets) (insert `X \<in> G j`, auto)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   275
      qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   276
      have "indep_sets (G(j:=?D)) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   277
      proof (rule indep_setsI)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   278
        fix i assume "i \<in> K" then show "(G(j := ?D)) i \<subseteq> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   279
          using G(2) by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   280
      next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   281
        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"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   282
        show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   283
        proof cases
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   284
          assume "j \<in> J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   285
          with A have indep: "indep_sets (G(j := {A j})) K" by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   286
          from J A show ?thesis
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   287
            by (intro indep_setsD[OF indep]) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   288
        next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   289
          assume "j \<notin> J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   290
          with J A have "\<forall>i\<in>J. A i \<in> G i" by (auto split: split_if_asm)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   291
          with J show ?thesis
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   292
            by (intro indep_setsD[OF G(1)]) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   293
        qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   294
      qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   295
      then have "indep_sets (G(j:=sets (dynkin \<lparr>space = space M, sets = G j\<rparr>))) K"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   296
        by (rule indep_sets_mono_sets) (insert mono, auto)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   297
      then show ?case
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   298
        by (rule indep_sets_mono_sets) (insert `j \<in> K` `j \<notin> J`, auto simp: G_def)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   299
    qed (insert `indep_sets F K`, simp) }
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   300
  from this[OF `indep_sets F J` `finite J` subset_refl]
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   301
  show "indep_sets (\<lambda>i. sets (dynkin \<lparr> space = space M, sets = F i \<rparr>)) J"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   302
    by (rule indep_sets_mono_sets) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   303
qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   304
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   305
lemma (in prob_space) indep_sets_sigma:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   306
  assumes indep: "indep_sets F I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   307
  assumes stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   308
  shows "indep_sets (\<lambda>i. sets (sigma \<lparr> space = space M, sets = F i \<rparr>)) I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   309
proof -
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   310
  from indep_sets_dynkin[OF indep]
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   311
  show ?thesis
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   312
  proof (rule indep_sets_mono_sets, subst sigma_eq_dynkin, simp_all add: stable)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   313
    fix i assume "i \<in> I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   314
    with indep have "F i \<subseteq> events" by (auto simp: indep_sets_def)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   315
    with sets_into_space show "F i \<subseteq> Pow (space M)" by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   316
  qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   317
qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   318
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   319
lemma (in prob_space) indep_sets_sigma_sets:
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   320
  assumes "indep_sets F I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   321
  assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   322
  shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   323
  using indep_sets_sigma[OF assms] by (simp add: sets_sigma)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   324
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   325
lemma (in prob_space) indep_sets_sigma_sets_iff:
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   326
  assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   327
  shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I \<longleftrightarrow> indep_sets F I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   328
proof
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   329
  assume "indep_sets F I" then show "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   330
    by (rule indep_sets_sigma_sets) fact
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   331
next
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   332
  assume "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I" then show "indep_sets F I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   333
    by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic)
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   334
qed
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   335
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   336
lemma (in prob_space) indep_sets2_eq:
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   337
  "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)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   338
  unfolding indep_set_def
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   339
proof (intro iffI ballI conjI)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   340
  assume indep: "indep_sets (bool_case A B) UNIV"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   341
  { fix a b assume "a \<in> A" "b \<in> B"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   342
    with indep_setsD[OF indep, of UNIV "bool_case a b"]
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   343
    show "prob (a \<inter> b) = prob a * prob b"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   344
      unfolding UNIV_bool by (simp add: ac_simps) }
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   345
  from indep show "A \<subseteq> events" "B \<subseteq> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   346
    unfolding indep_sets_def UNIV_bool by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   347
next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   348
  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)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   349
  show "indep_sets (bool_case A B) UNIV"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   350
  proof (rule indep_setsI)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   351
    fix i show "(case i of True \<Rightarrow> A | False \<Rightarrow> B) \<subseteq> events"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   352
      using * by (auto split: bool.split)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   353
  next
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   354
    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)"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   355
    then have "J = {True} \<or> J = {False} \<or> J = {True,False}"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   356
      by (auto simp: UNIV_bool)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   357
    then show "prob (\<Inter>j\<in>J. X j) = (\<Prod>j\<in>J. prob (X j))"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   358
      using X * by auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   359
  qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   360
qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   361
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   362
lemma (in prob_space) indep_set_sigma_sets:
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   363
  assumes "indep_set A B"
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   364
  assumes A: "Int_stable \<lparr> space = space M, sets = A \<rparr>"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   365
  assumes B: "Int_stable \<lparr> space = space M, sets = B \<rparr>"
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   366
  shows "indep_set (sigma_sets (space M) A) (sigma_sets (space M) B)"
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   367
proof -
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   368
  have "indep_sets (\<lambda>i. sigma_sets (space M) (case i of True \<Rightarrow> A | False \<Rightarrow> B)) UNIV"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   369
  proof (rule indep_sets_sigma_sets)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   370
    show "indep_sets (bool_case A B) UNIV"
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   371
      by (rule `indep_set A B`[unfolded indep_set_def])
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   372
    fix i show "Int_stable \<lparr>space = space M, sets = case i of True \<Rightarrow> A | False \<Rightarrow> B\<rparr>"
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   373
      using A B by (cases i) auto
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   374
  qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   375
  then show ?thesis
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   376
    unfolding indep_set_def
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   377
    by (rule indep_sets_mono_sets) (auto split: bool.split)
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   378
qed
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   379
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   380
lemma (in prob_space) indep_sets_collect_sigma:
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   381
  fixes I :: "'j \<Rightarrow> 'i set" and J :: "'j set" and E :: "'i \<Rightarrow> 'a set set"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   382
  assumes indep: "indep_sets E (\<Union>j\<in>J. I j)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   383
  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>"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   384
  assumes disjoint: "disjoint_family_on I J"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   385
  shows "indep_sets (\<lambda>j. sigma_sets (space M) (\<Union>i\<in>I j. E i)) J"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   386
proof -
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   387
  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) }"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   388
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   389
  from indep have E: "\<And>j i. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> E i \<subseteq> events"
42981
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   390
    unfolding indep_sets_def by auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   391
  { fix j
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   392
    let ?S = "sigma \<lparr> space = space M, sets = (\<Union>i\<in>I j. E i) \<rparr>"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   393
    assume "j \<in> J"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   394
    from E[OF this] interpret S: sigma_algebra ?S
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   395
      using sets_into_space by (intro sigma_algebra_sigma) auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   396
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   397
    have "sigma_sets (space M) (\<Union>i\<in>I j. E i) = sigma_sets (space M) (?E j)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   398
    proof (rule sigma_sets_eqI)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   399
      fix A assume "A \<in> (\<Union>i\<in>I j. E i)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   400
      then guess i ..
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   401
      then show "A \<in> sigma_sets (space M) (?E j)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   402
        by (auto intro!: sigma_sets.intros exI[of _ "{i}"] exI[of _ "\<lambda>i. A"])
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   403
    next
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   404
      fix A assume "A \<in> ?E j"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   405
      then obtain E' K where "finite K" "K \<noteq> {}" "K \<subseteq> I j" "\<And>k. k \<in> K \<Longrightarrow> E' k \<in> E k"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   406
        and A: "A = (\<Inter>k\<in>K. E' k)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   407
        by auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   408
      then have "A \<in> sets ?S" unfolding A
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   409
        by (safe intro!: S.finite_INT)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   410
           (auto simp: sets_sigma intro!: sigma_sets.Basic)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   411
      then show "A \<in> sigma_sets (space M) (\<Union>i\<in>I j. E i)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   412
        by (simp add: sets_sigma)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   413
    qed }
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   414
  moreover have "indep_sets (\<lambda>j. sigma_sets (space M) (?E j)) J"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   415
  proof (rule indep_sets_sigma_sets)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   416
    show "indep_sets ?E J"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   417
    proof (intro indep_setsI)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   418
      fix j assume "j \<in> J" with E show "?E j \<subseteq> events" by (force  intro!: finite_INT)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   419
    next
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   420
      fix K A assume K: "K \<noteq> {}" "K \<subseteq> J" "finite K"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   421
        and "\<forall>j\<in>K. A j \<in> ?E j"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   422
      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)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   423
        by simp
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   424
      from bchoice[OF this] guess E' ..
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   425
      from bchoice[OF this] obtain L
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   426
        where A: "\<And>j. j\<in>K \<Longrightarrow> A j = (\<Inter>l\<in>L j. E' j l)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   427
        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"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   428
        and E': "\<And>j l. j\<in>K \<Longrightarrow> l \<in> L j \<Longrightarrow> E' j l \<in> E l"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   429
        by auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   430
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   431
      { fix k l j assume "k \<in> K" "j \<in> K" "l \<in> L j" "l \<in> L k"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   432
        have "k = j"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   433
        proof (rule ccontr)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   434
          assume "k \<noteq> j"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   435
          with disjoint `K \<subseteq> J` `k \<in> K` `j \<in> K` have "I k \<inter> I j = {}"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   436
            unfolding disjoint_family_on_def by auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   437
          with L(2,3)[OF `j \<in> K`] L(2,3)[OF `k \<in> K`]
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   438
          show False using `l \<in> L k` `l \<in> L j` by auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   439
        qed }
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   440
      note L_inj = this
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   441
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   442
      def k \<equiv> "\<lambda>l. (SOME k. k \<in> K \<and> l \<in> L k)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   443
      { fix x j l assume *: "j \<in> K" "l \<in> L j"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   444
        have "k l = j" unfolding k_def
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   445
        proof (rule some_equality)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   446
          fix k assume "k \<in> K \<and> l \<in> L k"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   447
          with * L_inj show "k = j" by auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   448
        qed (insert *, simp) }
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   449
      note k_simp[simp] = this
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   450
      let "?E' l" = "E' (k l) l"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   451
      have "prob (\<Inter>j\<in>K. A j) = prob (\<Inter>l\<in>(\<Union>k\<in>K. L k). ?E' l)"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   452
        by (auto simp: A intro!: arg_cong[where f=prob])
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   453
      also have "\<dots> = (\<Prod>l\<in>(\<Union>k\<in>K. L k). prob (?E' l))"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   454
        using L K E' by (intro indep_setsD[OF indep]) (simp_all add: UN_mono)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   455
      also have "\<dots> = (\<Prod>j\<in>K. \<Prod>l\<in>L j. prob (E' j l))"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   456
        using K L L_inj by (subst setprod_UN_disjoint) auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   457
      also have "\<dots> = (\<Prod>j\<in>K. prob (A j))"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   458
        using K L E' by (auto simp add: A intro!: setprod_cong indep_setsD[OF indep, symmetric]) blast
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   459
      finally show "prob (\<Inter>j\<in>K. A j) = (\<Prod>j\<in>K. prob (A j))" .
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   460
    qed
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   461
  next
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   462
    fix j assume "j \<in> J"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   463
    show "Int_stable \<lparr> space = space M, sets = ?E j \<rparr>"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   464
    proof (rule Int_stableI)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   465
      fix a assume "a \<in> ?E j" then obtain Ka Ea
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   466
        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
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   467
      fix b assume "b \<in> ?E j" then obtain Kb Eb
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   468
        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
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   469
      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 {})"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   470
      have "a \<inter> b = INTER (Ka \<union> Kb) ?A"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   471
        by (simp add: a b set_eq_iff) auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   472
      with a b `j \<in> J` Int_stableD[OF Int_stable] show "a \<inter> b \<in> ?E j"
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   473
        by (intro CollectI exI[of _ "Ka \<union> Kb"] exI[of _ ?A]) auto
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   474
    qed
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   475
  qed
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   476
  ultimately show ?thesis
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   477
    by (simp cong: indep_sets_cong)
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   478
qed
fe7f5a26e4c6 add lemma indep_sets_collect_sigma
hoelzl
parents: 42861
diff changeset
   479
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   480
definition (in prob_space) terminal_events where
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   481
  "terminal_events A = (\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   482
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   483
lemma (in prob_space) terminal_events_sets:
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   484
  assumes A: "\<And>i. A i \<subseteq> events"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   485
  assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   486
  assumes X: "X \<in> terminal_events A"
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   487
  shows "X \<in> events"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   488
proof -
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   489
  let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   490
  interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   491
  from X have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by (auto simp: terminal_events_def)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   492
  from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   493
  then show "X \<in> events"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   494
    by induct (insert A, auto)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   495
qed
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   496
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   497
lemma (in prob_space) sigma_algebra_terminal_events:
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   498
  assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   499
  shows "sigma_algebra \<lparr> space = space M, sets = terminal_events A \<rparr>"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   500
  unfolding terminal_events_def
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   501
proof (simp add: sigma_algebra_iff2, safe)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   502
  let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   503
  interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   504
  { fix X x assume "X \<in> ?A" "x \<in> X"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   505
    then have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   506
    from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   507
    then have "X \<subseteq> space M"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   508
      by induct (insert A.sets_into_space, auto)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   509
    with `x \<in> X` show "x \<in> space M" by auto }
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   510
  { fix F :: "nat \<Rightarrow> 'a set" and n assume "range F \<subseteq> ?A"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   511
    then show "(UNION UNIV F) \<in> sigma_sets (space M) (UNION {n..} A)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   512
      by (intro sigma_sets.Union) auto }
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   513
qed (auto intro!: sigma_sets.Compl sigma_sets.Empty)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   514
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   515
lemma (in prob_space) kolmogorov_0_1_law:
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   516
  fixes A :: "nat \<Rightarrow> 'a set set"
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   517
  assumes A: "\<And>i. A i \<subseteq> events"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   518
  assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   519
  assumes indep: "indep_sets A UNIV"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   520
  and X: "X \<in> terminal_events A"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   521
  shows "prob X = 0 \<or> prob X = 1"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   522
proof -
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   523
  let ?D = "\<lparr> space = space M, sets = {D \<in> events. prob (X \<inter> D) = prob X * prob D} \<rparr>"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   524
  interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   525
  interpret T: sigma_algebra "\<lparr> space = space M, sets = terminal_events A \<rparr>"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   526
    by (rule sigma_algebra_terminal_events) fact
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   527
  have "X \<subseteq> space M" using T.space_closed X by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   528
42983
685df9c0626d use abbrevitation events == sets M
hoelzl
parents: 42982
diff changeset
   529
  have X_in: "X \<in> events"
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   530
    by (rule terminal_events_sets) fact+
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   531
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   532
  interpret D: dynkin_system ?D
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   533
  proof (rule dynkin_systemI)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   534
    fix D assume "D \<in> sets ?D" then show "D \<subseteq> space ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   535
      using sets_into_space by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   536
  next
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   537
    show "space ?D \<in> sets ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   538
      using prob_space `X \<subseteq> space M` by (simp add: Int_absorb2)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   539
  next
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   540
    fix A assume A: "A \<in> sets ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   541
    have "prob (X \<inter> (space M - A)) = prob (X - (X \<inter> A))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   542
      using `X \<subseteq> space M` by (auto intro!: arg_cong[where f=prob])
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   543
    also have "\<dots> = prob X - prob (X \<inter> A)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   544
      using X_in A by (intro finite_measure_Diff) auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   545
    also have "\<dots> = prob X * prob (space M) - prob X * prob A"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   546
      using A prob_space by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   547
    also have "\<dots> = prob X * prob (space M - A)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   548
      using X_in A sets_into_space
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   549
      by (subst finite_measure_Diff) (auto simp: field_simps)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   550
    finally show "space ?D - A \<in> sets ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   551
      using A `X \<subseteq> space M` by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   552
  next
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   553
    fix F :: "nat \<Rightarrow> 'a set" assume dis: "disjoint_family F" and "range F \<subseteq> sets ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   554
    then have F: "range F \<subseteq> events" "\<And>i. prob (X \<inter> F i) = prob X * prob (F i)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   555
      by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   556
    have "(\<lambda>i. prob (X \<inter> F i)) sums prob (\<Union>i. X \<inter> F i)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   557
    proof (rule finite_measure_UNION)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   558
      show "range (\<lambda>i. X \<inter> F i) \<subseteq> events"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   559
        using F X_in by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   560
      show "disjoint_family (\<lambda>i. X \<inter> F i)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   561
        using dis by (rule disjoint_family_on_bisimulation) auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   562
    qed
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   563
    with F have "(\<lambda>i. prob X * prob (F i)) sums prob (X \<inter> (\<Union>i. F i))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   564
      by simp
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   565
    moreover have "(\<lambda>i. prob X * prob (F i)) sums (prob X * prob (\<Union>i. F i))"
44282
f0de18b62d63 remove bounded_(bi)linear locale interpretations, to avoid duplicating so many lemmas
huffman
parents: 43920
diff changeset
   566
      by (intro sums_mult finite_measure_UNION F dis)
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   567
    ultimately have "prob (X \<inter> (\<Union>i. F i)) = prob X * prob (\<Union>i. F i)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   568
      by (auto dest!: sums_unique)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   569
    with F show "(\<Union>i. F i) \<in> sets ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   570
      by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   571
  qed
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   572
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   573
  { fix n
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   574
    have "indep_sets (\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) UNIV"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   575
    proof (rule indep_sets_collect_sigma)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   576
      have *: "(\<Union>b. case b of True \<Rightarrow> {..n} | False \<Rightarrow> {Suc n..}) = UNIV" (is "?U = _")
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   577
        by (simp split: bool.split add: set_eq_iff) (metis not_less_eq_eq)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   578
      with indep show "indep_sets A ?U" by simp
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   579
      show "disjoint_family (bool_case {..n} {Suc n..})"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   580
        unfolding disjoint_family_on_def by (auto split: bool.split)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   581
      fix m
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   582
      show "Int_stable \<lparr>space = space M, sets = A m\<rparr>"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   583
        unfolding Int_stable_def using A.Int by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   584
    qed
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   585
    also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) =
42982
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   586
      bool_case (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   587
      by (auto intro!: ext split: bool.split)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   588
    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))"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   589
      unfolding indep_set_def by simp
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   590
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   591
    have "sigma_sets (space M) (\<Union>m\<in>{..n}. A m) \<subseteq> sets ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   592
    proof (simp add: subset_eq, rule)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   593
      fix D assume D: "D \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   594
      have "X \<in> sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   595
        using X unfolding terminal_events_def by simp
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   596
      from indep_setD[OF indep D this] indep_setD_ev1[OF indep] D
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   597
      show "D \<in> events \<and> prob (X \<inter> D) = prob X * prob D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   598
        by (auto simp add: ac_simps)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   599
    qed }
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   600
  then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> sets ?D" (is "?A \<subseteq> _")
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   601
    by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   602
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   603
  have "sigma \<lparr> space = space M, sets = ?A \<rparr> =
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   604
    dynkin \<lparr> space = space M, sets = ?A \<rparr>" (is "sigma ?UA = dynkin ?UA")
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   605
  proof (rule sigma_eq_dynkin)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   606
    { fix B n assume "B \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   607
      then have "B \<subseteq> space M"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   608
        by induct (insert A sets_into_space, auto) }
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   609
    then show "sets ?UA \<subseteq> Pow (space ?UA)" by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   610
    show "Int_stable ?UA"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   611
    proof (rule Int_stableI)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   612
      fix a assume "a \<in> ?A" then guess n .. note a = this
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   613
      fix b assume "b \<in> ?A" then guess m .. note b = this
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   614
      interpret Amn: sigma_algebra "sigma \<lparr>space = space M, sets = (\<Union>i\<in>{..max m n}. A i)\<rparr>"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   615
        using A sets_into_space by (intro sigma_algebra_sigma) auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   616
      have "sigma_sets (space M) (\<Union>i\<in>{..n}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   617
        by (intro sigma_sets_subseteq UN_mono) auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   618
      with a have "a \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   619
      moreover
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   620
      have "sigma_sets (space M) (\<Union>i\<in>{..m}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   621
        by (intro sigma_sets_subseteq UN_mono) auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   622
      with b have "b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   623
      ultimately have "a \<inter> b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   624
        using Amn.Int[of a b] by (simp add: sets_sigma)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   625
      then show "a \<inter> b \<in> (\<Union>n. sigma_sets (space M) (\<Union>i\<in>{..n}. A i))" by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   626
    qed
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   627
  qed
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   628
  moreover have "sets (dynkin ?UA) \<subseteq> sets ?D"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   629
  proof (rule D.dynkin_subset)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   630
    show "sets ?UA \<subseteq> sets ?D" using `?A \<subseteq> sets ?D` by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   631
  qed simp
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   632
  ultimately have "sets (sigma ?UA) \<subseteq> sets ?D" by simp
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   633
  moreover
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   634
  have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> sigma_sets (space M) ?A"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   635
    by (intro sigma_sets_subseteq UN_mono) (auto intro: sigma_sets.Basic)
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   636
  then have "terminal_events A \<subseteq> sets (sigma ?UA)"
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   637
    unfolding sets_sigma terminal_events_def by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   638
  moreover note `X \<in> terminal_events A`
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   639
  ultimately have "X \<in> sets ?D" by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   640
  then show ?thesis by auto
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   641
qed
fa0ac7bee9ac add lemma kolmogorov_0_1_law
hoelzl
parents: 42981
diff changeset
   642
42985
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   643
lemma (in prob_space) borel_0_1_law:
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   644
  fixes F :: "nat \<Rightarrow> 'a set"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   645
  assumes F: "range F \<subseteq> events" "indep_events F UNIV"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   646
  shows "prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 0 \<or> prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 1"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   647
proof (rule kolmogorov_0_1_law[of "\<lambda>i. sigma_sets (space M) { F i }"])
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   648
  show "\<And>i. sigma_sets (space M) {F i} \<subseteq> events"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   649
    using F(1) sets_into_space
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   650
    by (subst sigma_sets_singleton) auto
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   651
  { fix i show "sigma_algebra \<lparr>space = space M, sets = sigma_sets (space M) {F i}\<rparr>"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   652
      using sigma_algebra_sigma[of "\<lparr>space = space M, sets = {F i}\<rparr>"] F sets_into_space
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   653
      by (auto simp add: sigma_def) }
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   654
  show "indep_sets (\<lambda>i. sigma_sets (space M) {F i}) UNIV"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   655
  proof (rule indep_sets_sigma_sets)
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   656
    show "indep_sets (\<lambda>i. {F i}) UNIV"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   657
      unfolding indep_sets_singleton_iff_indep_events by fact
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   658
    fix i show "Int_stable \<lparr>space = space M, sets = {F i}\<rparr>"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   659
      unfolding Int_stable_def by simp
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   660
  qed
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   661
  let "?Q n" = "\<Union>i\<in>{n..}. F i"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   662
  show "(\<Inter>n. \<Union>m\<in>{n..}. F m) \<in> terminal_events (\<lambda>i. sigma_sets (space M) {F i})"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   663
    unfolding terminal_events_def
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   664
  proof
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   665
    fix j
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   666
    interpret S: sigma_algebra "sigma \<lparr> space = space M, sets = (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})\<rparr>"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   667
      using order_trans[OF F(1) space_closed]
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   668
      by (intro sigma_algebra_sigma) (simp add: sigma_sets_singleton subset_eq)
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   669
    have "(\<Inter>n. ?Q n) = (\<Inter>n\<in>{j..}. ?Q n)"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   670
      by (intro decseq_SucI INT_decseq_offset UN_mono) auto
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   671
    also have "\<dots> \<in> sets (sigma \<lparr> space = space M, sets = (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})\<rparr>)"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   672
      using order_trans[OF F(1) space_closed]
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   673
      by (safe intro!: S.countable_INT S.countable_UN)
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   674
         (auto simp: sets_sigma sigma_sets_singleton intro!: sigma_sets.Basic bexI)
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   675
    finally show "(\<Inter>n. ?Q n) \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   676
      by (simp add: sets_sigma)
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   677
  qed
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   678
qed
1fb670792708 add lemma borel_0_1_law
hoelzl
parents: 42983
diff changeset
   679
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   680
lemma (in prob_space) indep_sets_finite:
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   681
  assumes I: "I \<noteq> {}" "finite I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   682
    and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> events" "\<And>i. i \<in> I \<Longrightarrow> space M \<in> F i"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   683
  shows "indep_sets F I \<longleftrightarrow> (\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j)))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   684
proof
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   685
  assume *: "indep_sets F I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   686
  from I show "\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   687
    by (intro indep_setsD[OF *] ballI) auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   688
next
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   689
  assume indep: "\<forall>A\<in>Pi I F. prob (\<Inter>j\<in>I. A j) = (\<Prod>j\<in>I. prob (A j))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   690
  show "indep_sets F I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   691
  proof (rule indep_setsI[OF F(1)])
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   692
    fix A J assume J: "J \<noteq> {}" "J \<subseteq> I" "finite J"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   693
    assume A: "\<forall>j\<in>J. A j \<in> F j"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   694
    let "?A j" = "if j \<in> J then A j else space M"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   695
    have "prob (\<Inter>j\<in>I. ?A j) = prob (\<Inter>j\<in>J. A j)"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   696
      using subset_trans[OF F(1) space_closed] J A
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   697
      by (auto intro!: arg_cong[where f=prob] split: split_if_asm) blast
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   698
    also
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   699
    from A F have "(\<lambda>j. if j \<in> J then A j else space M) \<in> Pi I F" (is "?A \<in> _")
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   700
      by (auto split: split_if_asm)
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   701
    with indep have "prob (\<Inter>j\<in>I. ?A j) = (\<Prod>j\<in>I. prob (?A j))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   702
      by auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   703
    also have "\<dots> = (\<Prod>j\<in>J. prob (A j))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   704
      unfolding if_distrib setprod.If_cases[OF `finite I`]
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   705
      using prob_space `J \<subseteq> I` by (simp add: Int_absorb1 setprod_1)
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   706
    finally show "prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))" ..
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   707
  qed
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   708
qed
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   709
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   710
lemma (in prob_space) indep_vars_finite:
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   711
  fixes I :: "'i set"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   712
  assumes I: "I \<noteq> {}" "finite I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   713
    and rv: "\<And>i. i \<in> I \<Longrightarrow> random_variable (sigma (M' i)) (X i)"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   714
    and Int_stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (M' i)"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   715
    and space: "\<And>i. i \<in> I \<Longrightarrow> space (M' i) \<in> sets (M' i)"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   716
  shows "indep_vars (\<lambda>i. sigma (M' i)) X I \<longleftrightarrow>
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   717
    (\<forall>A\<in>(\<Pi> i\<in>I. sets (M' i)). prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M)))"
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   718
proof -
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   719
  from rv have X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> space M \<rightarrow> space (M' i)"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   720
    unfolding measurable_def by simp
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   721
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   722
  { fix i assume "i\<in>I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   723
    have "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (sigma (M' i))}
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   724
      = sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   725
      unfolding sigma_sets_vimage_commute[OF X, OF `i \<in> I`]
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   726
      by (subst sigma_sets_sigma_sets_eq) auto }
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   727
  note this[simp]
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   728
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   729
  { fix i assume "i\<in>I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   730
    have "Int_stable \<lparr>space = space M, sets = {X i -` A \<inter> space M |A. A \<in> sets (M' i)}\<rparr>"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   731
    proof (rule Int_stableI)
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   732
      fix a assume "a \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   733
      then obtain A where "a = X i -` A \<inter> space M" "A \<in> sets (M' i)" by auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   734
      moreover
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   735
      fix b assume "b \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   736
      then obtain B where "b = X i -` B \<inter> space M" "B \<in> sets (M' i)" by auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   737
      moreover
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   738
      have "(X i -` A \<inter> space M) \<inter> (X i -` B \<inter> space M) = X i -` (A \<inter> B) \<inter> space M" by auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   739
      moreover note Int_stable[OF `i \<in> I`]
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   740
      ultimately
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   741
      show "a \<inter> b \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   742
        by (auto simp del: vimage_Int intro!: exI[of _ "A \<inter> B"] dest: Int_stableD)
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   743
    qed }
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   744
  note indep_sets_sigma_sets_iff[OF this, simp]
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   745
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   746
  { fix i assume "i \<in> I"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   747
    { fix A assume "A \<in> sets (M' i)"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   748
      then have "A \<in> sets (sigma (M' i))" by (auto simp: sets_sigma intro: sigma_sets.Basic)
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   749
      moreover
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   750
      from rv[OF `i\<in>I`] have "X i \<in> measurable M (sigma (M' i))" by auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   751
      ultimately
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   752
      have "X i -` A \<inter> space M \<in> sets M" by (auto intro: measurable_sets) }
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   753
    with X[OF `i\<in>I`] space[OF `i\<in>I`]
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   754
    have "{X i -` A \<inter> space M |A. A \<in> sets (M' i)} \<subseteq> events"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   755
      "space M \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   756
      by (auto intro!: exI[of _ "space (M' i)"]) }
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   757
  note indep_sets_finite[OF I this, simp]
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   758
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   759
  have "(\<forall>A\<in>\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> sets (M' i)}. prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))) =
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   760
    (\<forall>A\<in>\<Pi> i\<in>I. sets (M' i). prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M)))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   761
    (is "?L = ?R")
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   762
  proof safe
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   763
    fix A assume ?L and A: "A \<in> (\<Pi> i\<in>I. sets (M' i))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   764
    from `?L`[THEN bspec, of "\<lambda>i. X i -` A i \<inter> space M"] A `I \<noteq> {}`
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   765
    show "prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   766
      by (auto simp add: Pi_iff)
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   767
  next
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   768
    fix A assume ?R and A: "A \<in> (\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> sets (M' i)})"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   769
    from A have "\<forall>i\<in>I. \<exists>B. A i = X i -` B \<inter> space M \<and> B \<in> sets (M' i)" by auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   770
    from bchoice[OF this] obtain B where B: "\<forall>i\<in>I. A i = X i -` B i \<inter> space M"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   771
      "B \<in> (\<Pi> i\<in>I. sets (M' i))" by auto
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   772
    from `?R`[THEN bspec, OF B(2)] B(1) `I \<noteq> {}`
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   773
    show "prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))"
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   774
      by simp
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   775
  qed
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   776
  then show ?thesis using `I \<noteq> {}`
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   777
    by (simp add: rv indep_vars_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   778
qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   779
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   780
lemma (in prob_space) indep_vars_compose:
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   781
  assumes "indep_vars M' X I"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   782
  assumes rv:
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   783
    "\<And>i. i \<in> I \<Longrightarrow> sigma_algebra (N i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   784
    "\<And>i. i \<in> I \<Longrightarrow> Y i \<in> measurable (M' i) (N i)"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   785
  shows "indep_vars N (\<lambda>i. Y i \<circ> X i) I"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   786
  unfolding indep_vars_def
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   787
proof
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   788
  from rv `indep_vars M' X I`
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   789
  show "\<forall>i\<in>I. random_variable (N i) (Y i \<circ> X i)"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   790
    by (auto intro!: measurable_comp simp: indep_vars_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   791
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   792
  have "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   793
    using `indep_vars M' X I` by (simp add: indep_vars_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   794
  then show "indep_sets (\<lambda>i. sigma_sets (space M) {(Y i \<circ> X i) -` A \<inter> space M |A. A \<in> sets (N i)}) I"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   795
  proof (rule indep_sets_mono_sets)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   796
    fix i assume "i \<in> I"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   797
    with `indep_vars M' X I` have X: "X i \<in> space M \<rightarrow> space (M' i)"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   798
      unfolding indep_vars_def measurable_def by auto
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   799
    { fix A assume "A \<in> sets (N i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   800
      then have "\<exists>B. (Y i \<circ> X i) -` A \<inter> space M = X i -` B \<inter> space M \<and> B \<in> sets (M' i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   801
        by (intro exI[of _ "Y i -` A \<inter> space (M' i)"])
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   802
           (auto simp: vimage_compose intro!: measurable_sets rv `i \<in> I` funcset_mem[OF X]) }
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   803
    then show "sigma_sets (space M) {(Y i \<circ> X i) -` A \<inter> space M |A. A \<in> sets (N i)} \<subseteq>
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   804
      sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   805
      by (intro sigma_sets_subseteq) (auto simp: vimage_compose)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   806
  qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   807
qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   808
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   809
lemma (in prob_space) indep_varsD:
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   810
  assumes X: "indep_vars M' X I"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   811
  assumes I: "I \<noteq> {}" "finite I" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M' i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   812
  shows "prob (\<Inter>i\<in>I. X i -` A i \<inter> space M) = (\<Prod>i\<in>I. prob (X i -` A i \<inter> space M))"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   813
proof (rule indep_setsD)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   814
  show "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   815
    using X by (auto simp: indep_vars_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   816
  show "I \<subseteq> I" "I \<noteq> {}" "finite I" using I by auto
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   817
  show "\<forall>i\<in>I. X i -` A i \<inter> space M \<in> sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   818
    using I by (auto intro: sigma_sets.Basic)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   819
qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   820
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   821
lemma (in prob_space) indep_distribution_eq_measure:
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   822
  assumes I: "I \<noteq> {}" "finite I"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   823
  assumes rv: "\<And>i. random_variable (M' i) (X i)"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   824
  shows "indep_vars M' X I \<longleftrightarrow>
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   825
    (\<forall>A\<in>sets (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)).
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   826
      distribution (\<lambda>x. \<lambda>i\<in>I. X i x) A =
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   827
      finite_measure.\<mu>' (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)) A)"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   828
    (is "_ \<longleftrightarrow> (\<forall>X\<in>_. distribution ?D X = finite_measure.\<mu>' (Pi\<^isub>M I ?M) X)")
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   829
proof -
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   830
  interpret M': prob_space "?M i" for i
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   831
    using rv by (rule distribution_prob_space)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   832
  interpret P: finite_product_prob_space ?M I
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   833
    proof qed fact
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   834
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   835
  let ?D' = "(Pi\<^isub>M I ?M) \<lparr> measure := ereal \<circ> distribution ?D \<rparr>"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   836
  have "random_variable P.P ?D"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   837
    using `finite I` rv by (intro random_variable_restrict) auto
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   838
  then interpret D: prob_space ?D'
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   839
    by (rule distribution_prob_space)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   840
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   841
  show ?thesis
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   842
  proof (intro iffI ballI)
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   843
    assume "indep_vars M' X I"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   844
    fix A assume "A \<in> sets P.P"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   845
    moreover
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   846
    have "D.prob A = P.prob A"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   847
    proof (rule prob_space_unique_Int_stable)
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44282
diff changeset
   848
      show "prob_space ?D'" by unfold_locales
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44282
diff changeset
   849
      show "prob_space (Pi\<^isub>M I ?M)" by unfold_locales
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   850
      show "Int_stable P.G" using M'.Int
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   851
        by (intro Int_stable_product_algebra_generator) (simp add: Int_stable_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   852
      show "space P.G \<in> sets P.G"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   853
        using M'.top by (simp add: product_algebra_generator_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   854
      show "space ?D' = space P.G"  "sets ?D' = sets (sigma P.G)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   855
        by (simp_all add: product_algebra_def product_algebra_generator_def sets_sigma)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   856
      show "space P.P = space P.G" "sets P.P = sets (sigma P.G)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   857
        by (simp_all add: product_algebra_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   858
      show "A \<in> sets (sigma P.G)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   859
        using `A \<in> sets P.P` by (simp add: product_algebra_def)
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   860
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   861
      fix E assume E: "E \<in> sets P.G"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   862
      then have "E \<in> sets P.P"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   863
        by (simp add: sets_sigma sigma_sets.Basic product_algebra_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   864
      then have "D.prob E = distribution ?D E"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   865
        unfolding D.\<mu>'_def by simp
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   866
      also
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   867
      from E obtain F where "E = Pi\<^isub>E I F" and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets (M' i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   868
        by (auto simp: product_algebra_generator_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   869
      with `I \<noteq> {}` have "distribution ?D E = prob (\<Inter>i\<in>I. X i -` F i \<inter> space M)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   870
        using `I \<noteq> {}` by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   871
      also have "\<dots> = (\<Prod>i\<in>I. prob (X i -` F i \<inter> space M))"
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   872
        using `indep_vars M' X I` I F by (rule indep_varsD)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   873
      also have "\<dots> = P.prob E"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   874
        using F by (simp add: `E = Pi\<^isub>E I F` P.prob_times M'.\<mu>'_def distribution_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   875
      finally show "D.prob E = P.prob E" .
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   876
    qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   877
    ultimately show "distribution ?D A = P.prob A"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   878
      by (simp add: D.\<mu>'_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   879
  next
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   880
    assume eq: "\<forall>A\<in>sets P.P. distribution ?D A = P.prob A"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   881
    have [simp]: "\<And>i. sigma (M' i) = M' i"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   882
      using rv by (intro sigma_algebra.sigma_eq) simp
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   883
    have "indep_vars (\<lambda>i. sigma (M' i)) X I"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   884
    proof (subst indep_vars_finite[OF I])
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   885
      fix i assume [simp]: "i \<in> I"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   886
      show "random_variable (sigma (M' i)) (X i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   887
        using rv[of i] by simp
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   888
      show "Int_stable (M' i)" "space (M' i) \<in> sets (M' i)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   889
        using M'.Int[of _ i] M'.top by (auto simp: Int_stable_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   890
    next
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   891
      show "\<forall>A\<in>\<Pi> i\<in>I. sets (M' i). prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M))"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   892
      proof
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   893
        fix A assume A: "A \<in> (\<Pi> i\<in>I. sets (M' i))"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   894
        then have A_in_P: "(Pi\<^isub>E I A) \<in> sets P.P"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   895
          by (auto intro!: product_algebraI)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   896
        have "prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = distribution ?D (Pi\<^isub>E I A)"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   897
          using `I \<noteq> {}`by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   898
        also have "\<dots> = P.prob (Pi\<^isub>E I A)" using A_in_P eq by simp
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   899
        also have "\<dots> = (\<Prod>i\<in>I. M'.prob i (A i))"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   900
          using A by (intro P.prob_times) auto
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   901
        also have "\<dots> = (\<Prod>i\<in>I. prob (X i -` A i \<inter> space M))"
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   902
          using A by (auto intro!: setprod_cong simp: M'.\<mu>'_def Pi_iff distribution_def)
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   903
        finally show "prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M))" .
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   904
      qed
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   905
    qed
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   906
    then show "indep_vars M' X I"
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   907
      by simp
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42987
diff changeset
   908
  qed
42987
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   909
qed
73e2d802ea41 add lemma indep_rv_finite
hoelzl
parents: 42985
diff changeset
   910
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   911
lemma (in prob_space) indep_varD:
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   912
  assumes indep: "indep_var Ma A Mb B"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   913
  assumes sets: "Xa \<in> sets Ma" "Xb \<in> sets Mb"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   914
  shows "prob ((\<lambda>x. (A x, B x)) -` (Xa \<times> Xb) \<inter> space M) =
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   915
    prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   916
proof -
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   917
  have "prob ((\<lambda>x. (A x, B x)) -` (Xa \<times> Xb) \<inter> space M) =
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   918
    prob (\<Inter>i\<in>UNIV. (bool_case A B i -` bool_case Xa Xb i \<inter> space M))"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   919
    by (auto intro!: arg_cong[where f=prob] simp: UNIV_bool)
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   920
  also have "\<dots> = (\<Prod>i\<in>UNIV. prob (bool_case A B i -` bool_case Xa Xb i \<inter> space M))"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   921
    using indep unfolding indep_var_def
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   922
    by (rule indep_varsD) (auto split: bool.split intro: sets)
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   923
  also have "\<dots> = prob (A -` Xa \<inter> space M) * prob (B -` Xb \<inter> space M)"
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   924
    unfolding UNIV_bool by simp
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   925
  finally show ?thesis .
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   926
qed
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   927
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   928
lemma (in prob_space)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   929
  assumes "indep_var S X T Y"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   930
  shows indep_var_rv1: "random_variable S X"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   931
    and indep_var_rv2: "random_variable T Y"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   932
proof -
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   933
  have "\<forall>i\<in>UNIV. random_variable (bool_case S T i) (bool_case X Y i)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   934
    using assms unfolding indep_var_def indep_vars_def by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   935
  then show "random_variable S X" "random_variable T Y"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   936
    unfolding UNIV_bool by auto
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   937
qed
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   938
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   939
lemma (in prob_space) indep_var_distributionD:
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   940
  assumes indep: "indep_var S X T Y"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   941
  defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   942
  assumes "A \<in> sets P"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   943
  shows "joint_distribution X Y A = finite_measure.\<mu>' P A"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   944
proof -
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   945
  from indep have rvs: "random_variable S X" "random_variable T Y"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   946
    by (blast dest: indep_var_rv1 indep_var_rv2)+
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   947
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   948
  let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   949
  let ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   950
  interpret X: prob_space ?S by (rule distribution_prob_space) fact
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   951
  interpret Y: prob_space ?T by (rule distribution_prob_space) fact
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   952
  interpret XY: pair_prob_space ?S ?T by default
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   953
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   954
  let ?J = "XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   955
  interpret J: prob_space ?J
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   956
    by (rule joint_distribution_prob_space) (simp_all add: rvs)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   957
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 43340
diff changeset
   958
  have "finite_measure.\<mu>' (XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>) A = XY.\<mu>' A"
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   959
  proof (rule prob_space_unique_Int_stable)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   960
    show "Int_stable (pair_measure_generator ?S ?T)" (is "Int_stable ?P")
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   961
      by fact
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   962
    show "space ?P \<in> sets ?P"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   963
      unfolding space_pair_measure[simplified pair_measure_def space_sigma]
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   964
      using X.top Y.top by (auto intro!: pair_measure_generatorI)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   965
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44282
diff changeset
   966
    show "prob_space ?J" by unfold_locales
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   967
    show "space ?J = space ?P"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   968
      by (simp add: pair_measure_generator_def space_pair_measure)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   969
    show "sets ?J = sets (sigma ?P)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   970
      by (simp add: pair_measure_def)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   971
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44282
diff changeset
   972
    show "prob_space XY.P" by unfold_locales
43340
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   973
    show "space XY.P = space ?P" "sets XY.P = sets (sigma ?P)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   974
      by (simp_all add: pair_measure_generator_def pair_measure_def)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   975
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   976
    show "A \<in> sets (sigma ?P)"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   977
      using `A \<in> sets P` unfolding P_def pair_measure_def by simp
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   978
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   979
    fix X assume "X \<in> sets ?P"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   980
    then obtain A B where "A \<in> sets S" "B \<in> sets T" "X = A \<times> B"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   981
      by (auto simp: sets_pair_measure_generator)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   982
    then show "J.\<mu>' X = XY.\<mu>' X"
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   983
      unfolding J.\<mu>'_def XY.\<mu>'_def using indep
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   984
      by (simp add: XY.pair_measure_times)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   985
         (simp add: distribution_def indep_varD)
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   986
  qed
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   987
  then show ?thesis
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   988
    using `A \<in> sets P` unfolding P_def J.\<mu>'_def XY.\<mu>'_def by simp
60e181c4eae4 lemma: independence is equal to mutual information = 0
hoelzl
parents: 42989
diff changeset
   989
qed
42989
40adeda9a8d2 introduce independence of two random variables
hoelzl
parents: 42988
diff changeset
   990
42861
16375b493b64 Add formalization of probabilistic independence for families of sets
hoelzl
parents:
diff changeset
   991
end