src/HOL/Probability/Finite_Product_Measure.thy
author hoelzl
Fri, 16 Nov 2012 18:45:57 +0100
changeset 50104 de19856feb54
parent 50099 a58bb401af80
child 50123 69b35a75caf3
permissions -rw-r--r--
move theorems to be more generally useable
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     1
(*  Title:      HOL/Probability/Finite_Product_Measure.thy
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     2
    Author:     Johannes Hölzl, TU München
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     3
*)
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     4
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     5
header {*Finite product measures*}
42067
66c8281349ec standardized headers
hoelzl
parents: 41981
diff changeset
     6
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     7
theory Finite_Product_Measure
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 42067
diff changeset
     8
imports Binary_Product_Measure
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
     9
begin
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
    10
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    11
lemma split_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    12
  by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    13
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    14
abbreviation
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    15
  "Pi\<^isub>E A B \<equiv> Pi A B \<inter> extensional A"
39094
67da17aced5a measure unique
hellerar
parents: 39080
diff changeset
    16
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    17
syntax
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    18
  "_PiE"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PIE _:_./ _)" 10)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    19
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    20
syntax (xsymbols)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    21
  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    22
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    23
syntax (HTML output)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    24
  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    25
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    26
translations
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    27
  "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
    28
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    29
abbreviation
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    30
  funcset_extensional :: "['a set, 'b set] => ('a => 'b) set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    31
    (infixr "->\<^isub>E" 60) where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    32
  "A ->\<^isub>E B \<equiv> Pi\<^isub>E A (%_. B)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    33
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    34
notation (xsymbols)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    35
  funcset_extensional  (infixr "\<rightarrow>\<^isub>E" 60)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    36
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    37
lemma extensional_insert[intro, simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    38
  assumes "a \<in> extensional (insert i I)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    39
  shows "a(i := b) \<in> extensional (insert i I)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    40
  using assms unfolding extensional_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    41
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    42
lemma extensional_Int[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    43
  "extensional I \<inter> extensional I' = extensional (I \<inter> I')"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    44
  unfolding extensional_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents: 36649
diff changeset
    45
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    46
lemma extensional_UNIV[simp]: "extensional UNIV = UNIV"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    47
  by (auto simp: extensional_def)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    48
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    49
lemma restrict_extensional_sub[intro]: "A \<subseteq> B \<Longrightarrow> restrict f A \<in> extensional B"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    50
  unfolding restrict_def extensional_def by auto
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    51
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    52
lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \<inter> B)"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    53
  unfolding restrict_def by (simp add: fun_eq_iff)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    54
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
    55
definition
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    56
  "merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    57
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    58
lemma merge_apply[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    59
  "I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    60
  "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    61
  "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    62
  "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    63
  "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I J (x, y) i = undefined"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    64
  unfolding merge_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    65
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    66
lemma merge_commute:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    67
  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) = merge J I (y, x)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
    68
  by (force simp: merge_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    69
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    70
lemma Pi_cancel_merge_range[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    71
  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    72
  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    73
  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    74
  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    75
  by (auto simp: Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    76
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    77
lemma Pi_cancel_merge[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    78
  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    79
  "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    80
  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    81
  "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    82
  by (auto simp: Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    83
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    84
lemma extensional_merge[simp]: "merge I J (x, y) \<in> extensional (I \<union> J)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    85
  by (auto simp: extensional_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    86
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    87
lemma restrict_Pi_cancel: "restrict x I \<in> Pi I A \<longleftrightarrow> x \<in> Pi I A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    88
  by (auto simp: restrict_def Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    89
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    90
lemma restrict_merge[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    91
  "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    92
  "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    93
  "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    94
  "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    95
  by (auto simp: restrict_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    96
50042
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    97
lemma split_merge: "P (merge I J (x,y) i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    98
  unfolding merge_def by auto
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    99
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   100
lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I J (x, y) \<in> extensional K"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   101
  unfolding merge_def extensional_def by auto
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   102
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   103
lemma injective_vimage_restrict:
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   104
  assumes J: "J \<subseteq> I"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   105
  and sets: "A \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" and ne: "(\<Pi>\<^isub>E i\<in>I. S i) \<noteq> {}"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   106
  and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   107
  shows "A = B"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   108
proof  (intro set_eqI)
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   109
  fix x
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   110
  from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   111
  have "J \<inter> (I - J) = {}" by auto
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   112
  show "x \<in> A \<longleftrightarrow> x \<in> B"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   113
  proof cases
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   114
    assume x: "x \<in> (\<Pi>\<^isub>E i\<in>J. S i)"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   115
    have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   116
      using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub split: split_merge)
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   117
    then show "x \<in> A \<longleftrightarrow> x \<in> B"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   118
      using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub eq split: split_merge)
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   119
  next
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   120
    assume "x \<notin> (\<Pi>\<^isub>E i\<in>J. S i)" with sets show "x \<in> A \<longleftrightarrow> x \<in> B" by auto
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   121
  qed
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   122
qed
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   123
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   124
lemma extensional_insert_undefined[intro, simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   125
  assumes "a \<in> extensional (insert i I)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   126
  shows "a(i := undefined) \<in> extensional I"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   127
  using assms unfolding extensional_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   128
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   129
lemma extensional_insert_cancel[intro, simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   130
  assumes "a \<in> extensional I"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   131
  shows "a \<in> extensional (insert i I)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   132
  using assms unfolding extensional_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   133
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   134
lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   135
  unfolding merge_def by (auto simp: fun_eq_iff)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   136
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   137
lemma Pi_Int: "Pi I E \<inter> Pi I F = (\<Pi> i\<in>I. E i \<inter> F i)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   138
  by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   139
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   140
lemma PiE_Int: "(Pi\<^isub>E I A) \<inter> (Pi\<^isub>E I B) = Pi\<^isub>E I (\<lambda>x. A x \<inter> B x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   141
  by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   142
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   143
lemma Pi_cancel_fupd_range[simp]: "i \<notin> I \<Longrightarrow> x \<in> Pi I (B(i := b)) \<longleftrightarrow> x \<in> Pi I B"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   144
  by (auto simp: Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   145
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   146
lemma Pi_split_insert_domain[simp]: "x \<in> Pi (insert i I) X \<longleftrightarrow> x \<in> Pi I X \<and> x i \<in> X i"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   147
  by (auto simp: Pi_def)
39088
ca17017c10e6 Measurable on product space is equiv. to measurable components
hoelzl
parents: 39082
diff changeset
   148
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   149
lemma Pi_split_domain[simp]: "x \<in> Pi (I \<union> J) X \<longleftrightarrow> x \<in> Pi I X \<and> x \<in> Pi J X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   150
  by (auto simp: Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   151
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   152
lemma Pi_cancel_fupd[simp]: "i \<notin> I \<Longrightarrow> x(i := a) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   153
  by (auto simp: Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   154
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   155
lemma restrict_vimage:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   156
  assumes "I \<inter> J = {}"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   157
  shows "(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I J (E, F))"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   158
  using assms by (auto simp: restrict_Pi_cancel)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   159
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   160
lemma merge_vimage:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   161
  assumes "I \<inter> J = {}"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   162
  shows "merge I J -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E"
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   163
  using assms by (auto simp: restrict_Pi_cancel)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   164
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   165
lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   166
  by (auto simp: restrict_def)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   167
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   168
lemma merge_restrict[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   169
  "merge I J (restrict x I, y) = merge I J (x, y)"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   170
  "merge I J (x, restrict y J) = merge I J (x, y)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   171
  unfolding merge_def by auto
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   172
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   173
lemma merge_x_x_eq_restrict[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   174
  "merge I J (x, x) = restrict x (I \<union> J)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   175
  unfolding merge_def by auto
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   176
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   177
lemma Pi_fupd_iff: "i \<in> I \<Longrightarrow> f \<in> Pi I (B(i := A)) \<longleftrightarrow> f \<in> Pi (I - {i}) B \<and> f i \<in> A"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   178
  apply auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   179
  apply (drule_tac x=x in Pi_mem)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   180
  apply (simp_all split: split_if_asm)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   181
  apply (drule_tac x=i in Pi_mem)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   182
  apply (auto dest!: Pi_mem)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   183
  done
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   184
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   185
lemma Pi_UN:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   186
  fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   187
  assumes "finite I" and mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   188
  shows "(\<Union>n. Pi I (A n)) = (\<Pi> i\<in>I. \<Union>n. A n i)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   189
proof (intro set_eqI iffI)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   190
  fix f assume "f \<in> (\<Pi> i\<in>I. \<Union>n. A n i)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   191
  then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   192
  from bchoice[OF this] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   193
  obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   194
    using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   195
  have "f \<in> Pi I (A k)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   196
  proof (intro Pi_I)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   197
    fix i assume "i \<in> I"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   198
    from mono[OF this, of "n i" k] k[OF this] n[OF this]
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   199
    show "f i \<in> A k i" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   200
  qed
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   201
  then show "f \<in> (\<Union>n. Pi I (A n))" by auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   202
qed auto
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   203
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   204
lemma PiE_cong:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   205
  assumes "\<And>i. i\<in>I \<Longrightarrow> A i = B i"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   206
  shows "Pi\<^isub>E I A = Pi\<^isub>E I B"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   207
  using assms by (auto intro!: Pi_cong)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   208
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   209
lemma restrict_upd[simp]:
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   210
  "i \<notin> I \<Longrightarrow> (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)"
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   211
  by (auto simp: fun_eq_iff)
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   212
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   213
lemma Pi_eq_subset:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   214
  assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   215
  assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \<in> I"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   216
  shows "F i \<subseteq> F' i"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   217
proof
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   218
  fix x assume "x \<in> F i"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   219
  with ne have "\<forall>j. \<exists>y. ((j \<in> I \<longrightarrow> y \<in> F j \<and> (i = j \<longrightarrow> x = y)) \<and> (j \<notin> I \<longrightarrow> y = undefined))" by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   220
  from choice[OF this] guess f .. note f = this
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   221
  then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   222
  then have "f \<in> Pi\<^isub>E I F'" using assms by simp
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   223
  then show "x \<in> F' i" using f `i \<in> I` by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   224
qed
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   225
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   226
lemma Pi_eq_iff_not_empty:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   227
  assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   228
  shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   229
proof (intro iffI ballI)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   230
  fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \<in> I"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   231
  show "F i = F' i"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   232
    using Pi_eq_subset[of I F F', OF ne eq i]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   233
    using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i]
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   234
    by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   235
qed auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   236
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   237
lemma Pi_eq_empty_iff:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   238
  "Pi\<^isub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   239
proof
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   240
  assume "Pi\<^isub>E I F = {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   241
  show "\<exists>i\<in>I. F i = {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   242
  proof (rule ccontr)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   243
    assume "\<not> ?thesis"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   244
    then have "\<forall>i. \<exists>y. (i \<in> I \<longrightarrow> y \<in> F i) \<and> (i \<notin> I \<longrightarrow> y = undefined)" by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   245
    from choice[OF this] guess f ..
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   246
    then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   247
    with `Pi\<^isub>E I F = {}` show False by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   248
  qed
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   249
qed auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   250
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   251
lemma Pi_eq_iff:
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   252
  "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   253
proof (intro iffI disjCI)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   254
  assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   255
  assume "\<not> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   256
  then have "(\<forall>i\<in>I. F i \<noteq> {}) \<and> (\<forall>i\<in>I. F' i \<noteq> {})"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   257
    using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   258
  with Pi_eq_iff_not_empty[of I F F'] show "\<forall>i\<in>I. F i = F' i" by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   259
next
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   260
  assume "(\<forall>i\<in>I. F i = F' i) \<or> (\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {})"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   261
  then show "Pi\<^isub>E I F = Pi\<^isub>E I F'"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   262
    using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   263
qed
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   264
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   265
lemma funset_eq_UN_fun_upd_I:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   266
  assumes *: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f(a := d) \<in> F A"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   267
  and **: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f a \<in> G (f(a:=d))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   268
  and ***: "\<And>f x. \<lbrakk> f \<in> F A ; x \<in> G f \<rbrakk> \<Longrightarrow> f(a:=x) \<in> F (insert a A)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   269
  shows "F (insert a A) = (\<Union>f\<in>F A. fun_upd f a ` (G f))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   270
proof safe
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   271
  fix f assume f: "f \<in> F (insert a A)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   272
  show "f \<in> (\<Union>f\<in>F A. fun_upd f a ` G f)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   273
  proof (rule UN_I[of "f(a := d)"])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   274
    show "f(a := d) \<in> F A" using *[OF f] .
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   275
    show "f \<in> fun_upd (f(a:=d)) a ` G (f(a:=d))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   276
    proof (rule image_eqI[of _ _ "f a"])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   277
      show "f a \<in> G (f(a := d))" using **[OF f] .
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   278
    qed simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   279
  qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   280
next
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   281
  fix f x assume "f \<in> F A" "x \<in> G f"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   282
  from ***[OF this] show "f(a := x) \<in> F (insert a A)" .
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   283
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   284
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   285
lemma extensional_funcset_insert_eq[simp]:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   286
  assumes "a \<notin> A"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   287
  shows "extensional (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensional A \<inter> (A \<rightarrow> B). (\<lambda>b. f(a := b)) ` B)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   288
  apply (rule funset_eq_UN_fun_upd_I)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   289
  using assms
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   290
  by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   291
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   292
lemma finite_extensional_funcset[simp, intro]:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   293
  assumes "finite A" "finite B"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   294
  shows "finite (extensional A \<inter> (A \<rightarrow> B))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   295
  using assms by induct auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   296
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   297
lemma finite_PiE[simp, intro]:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   298
  assumes fin: "finite A" "\<And>i. i \<in> A \<Longrightarrow> finite (B i)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   299
  shows "finite (Pi\<^isub>E A B)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   300
proof -
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   301
  have *: "(Pi\<^isub>E A B) \<subseteq> extensional A \<inter> (A \<rightarrow> (\<Union>i\<in>A. B i))" by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   302
  show ?thesis
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   303
    using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   304
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   305
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   306
section "Finite product spaces"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   307
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   308
section "Products"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   309
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   310
definition prod_emb where
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   311
  "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   312
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   313
lemma prod_emb_iff: 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   314
  "f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   315
  unfolding prod_emb_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   316
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   317
lemma
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   318
  shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   319
    and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   320
    and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   321
    and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   322
    and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   323
    and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   324
  by (auto simp: prod_emb_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   325
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   326
lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   327
    prod_emb I M J (\<Pi>\<^isub>E i\<in>J. E i) = (\<Pi>\<^isub>E i\<in>I. if i \<in> J then E i else space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   328
  by (force simp: prod_emb_def Pi_iff split_if_mem2)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   329
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   330
lemma prod_emb_PiE_same_index[simp]: "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^isub>E I E) = Pi\<^isub>E I E"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   331
  by (auto simp: prod_emb_def Pi_iff)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   332
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   333
lemma prod_emb_trans[simp]:
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   334
  "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   335
  by (auto simp add: Int_absorb1 prod_emb_def)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   336
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   337
lemma prod_emb_Pi:
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   338
  assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   339
  shows "prod_emb K M J (Pi\<^isub>E J X) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then X i else space (M i))"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   340
  using assms space_closed
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   341
  by (auto simp: prod_emb_def Pi_iff split: split_if_asm) blast+
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   342
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   343
lemma prod_emb_id:
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   344
  "B \<subseteq> (\<Pi>\<^isub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   345
  by (auto simp: prod_emb_def Pi_iff subset_eq extensional_restrict)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   346
50041
afe886a04198 removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents: 50038
diff changeset
   347
lemma prod_emb_mono:
afe886a04198 removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents: 50038
diff changeset
   348
  "F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G"
afe886a04198 removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents: 50038
diff changeset
   349
  by (auto simp: prod_emb_def)
afe886a04198 removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents: 50038
diff changeset
   350
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   351
definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   352
  "PiM I M = extend_measure (\<Pi>\<^isub>E i\<in>I. space (M i))
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   353
    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   354
    (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j))
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   355
    (\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   356
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   357
definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   358
  "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j)) `
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   359
    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   360
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   361
abbreviation
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   362
  "Pi\<^isub>M I M \<equiv> PiM I M"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   363
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   364
syntax
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   365
  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3PIM _:_./ _)" 10)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   366
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   367
syntax (xsymbols)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   368
  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"  10)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   369
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   370
syntax (HTML output)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   371
  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"  10)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   372
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   373
translations
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   374
  "PIM x:I. M" == "CONST PiM I (%x. M)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   375
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   376
lemma prod_algebra_sets_into_space:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   377
  "prod_algebra I M \<subseteq> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   378
  using assms by (auto simp: prod_emb_def prod_algebra_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   379
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   380
lemma prod_algebra_eq_finite:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   381
  assumes I: "finite I"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   382
  shows "prod_algebra I M = {(\<Pi>\<^isub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   383
proof (intro iffI set_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   384
  fix A assume "A \<in> ?L"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   385
  then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   386
    and A: "A = prod_emb I M J (PIE j:J. E j)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   387
    by (auto simp: prod_algebra_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   388
  let ?A = "\<Pi>\<^isub>E i\<in>I. if i \<in> J then E i else space (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   389
  have A: "A = ?A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   390
    unfolding A using J by (intro prod_emb_PiE sets_into_space) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   391
  show "A \<in> ?R" unfolding A using J top
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   392
    by (intro CollectI exI[of _ "\<lambda>i. if i \<in> J then E i else space (M i)"]) simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   393
next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   394
  fix A assume "A \<in> ?R"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   395
  then obtain X where "A = (\<Pi>\<^isub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   396
  then have A: "A = prod_emb I M I (\<Pi>\<^isub>E i\<in>I. X i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   397
    using sets_into_space by (force simp: prod_emb_def Pi_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   398
  from X I show "A \<in> ?L" unfolding A
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   399
    by (auto simp: prod_algebra_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   400
qed
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   401
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   402
lemma prod_algebraI:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   403
  "finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i))
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   404
    \<Longrightarrow> prod_emb I M J (PIE j:J. E j) \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   405
  by (auto simp: prod_algebra_def Pi_iff)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   406
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   407
lemma prod_algebraI_finite:
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   408
  "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^isub>E I E) \<in> prod_algebra I M"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   409
  using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets_into_space] by simp
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   410
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   411
lemma Int_stable_PiE: "Int_stable {Pi\<^isub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   412
proof (safe intro!: Int_stableI)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   413
  fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   414
  then show "\<exists>G. Pi\<^isub>E J E \<inter> Pi\<^isub>E J F = Pi\<^isub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   415
    by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"])
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   416
qed
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   417
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   418
lemma prod_algebraE:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   419
  assumes A: "A \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   420
  obtains J E where "A = prod_emb I M J (PIE j:J. E j)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   421
    "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)" 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   422
  using A by (auto simp: prod_algebra_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42950
diff changeset
   423
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   424
lemma prod_algebraE_all:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   425
  assumes A: "A \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   426
  obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   427
proof -
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   428
  from A obtain E J where A: "A = prod_emb I M J (Pi\<^isub>E J E)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   429
    and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   430
    by (auto simp: prod_algebra_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   431
  from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   432
    using sets_into_space by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   433
  then have "A = (\<Pi>\<^isub>E i\<in>I. if i\<in>J then E i else space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   434
    using A J by (auto simp: prod_emb_PiE)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   435
  moreover then have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   436
    using top E by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   437
  ultimately show ?thesis using that by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   438
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   439
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   440
lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   441
proof (unfold Int_stable_def, safe)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   442
  fix A assume "A \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   443
  from prod_algebraE[OF this] guess J E . note A = this
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   444
  fix B assume "B \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   445
  from prod_algebraE[OF this] guess K F . note B = this
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   446
  have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^isub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter> 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   447
      (if i \<in> K then F i else space (M i)))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   448
    unfolding A B using A(2,3,4) A(5)[THEN sets_into_space] B(2,3,4) B(5)[THEN sets_into_space]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   449
    apply (subst (1 2 3) prod_emb_PiE)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   450
    apply (simp_all add: subset_eq PiE_Int)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   451
    apply blast
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   452
    apply (intro PiE_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   453
    apply auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   454
    done
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   455
  also have "\<dots> \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   456
    using A B by (auto intro!: prod_algebraI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   457
  finally show "A \<inter> B \<in> prod_algebra I M" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   458
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   459
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   460
lemma prod_algebra_mono:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   461
  assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   462
  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   463
  shows "prod_algebra I E \<subseteq> prod_algebra I F"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   464
proof
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   465
  fix A assume "A \<in> prod_algebra I E"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   466
  then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   467
    and A: "A = prod_emb I E J (\<Pi>\<^isub>E i\<in>J. G i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   468
    and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   469
    by (auto simp: prod_algebra_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   470
  moreover
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   471
  from space have "(\<Pi>\<^isub>E i\<in>I. space (E i)) = (\<Pi>\<^isub>E i\<in>I. space (F i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   472
    by (rule PiE_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   473
  with A have "A = prod_emb I F J (\<Pi>\<^isub>E i\<in>J. G i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   474
    by (simp add: prod_emb_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   475
  moreover
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   476
  from sets G J have "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (F i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   477
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   478
  ultimately show "A \<in> prod_algebra I F"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   479
    apply (simp add: prod_algebra_def image_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   480
    apply (intro exI[of _ J] exI[of _ G] conjI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   481
    apply auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   482
    done
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   483
qed
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   484
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   485
lemma prod_algebra_cong:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   486
  assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   487
  shows "prod_algebra I M = prod_algebra J N"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   488
proof -
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   489
  have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   490
    using sets_eq_imp_space_eq[OF sets] by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   491
  with sets show ?thesis unfolding `I = J`
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   492
    by (intro antisym prod_algebra_mono) auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   493
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   494
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   495
lemma space_in_prod_algebra:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   496
  "(\<Pi>\<^isub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   497
proof cases
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   498
  assume "I = {}" then show ?thesis
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   499
    by (auto simp add: prod_algebra_def image_iff prod_emb_def)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   500
next
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   501
  assume "I \<noteq> {}"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   502
  then obtain i where "i \<in> I" by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   503
  then have "(\<Pi>\<^isub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   504
    by (auto simp: prod_emb_def Pi_iff)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   505
  also have "\<dots> \<in> prod_algebra I M"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   506
    using `i \<in> I` by (intro prod_algebraI) auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   507
  finally show ?thesis .
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   508
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   509
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   510
lemma space_PiM: "space (\<Pi>\<^isub>M i\<in>I. M i) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   511
  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   512
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   513
lemma sets_PiM: "sets (\<Pi>\<^isub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) (prod_algebra I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   514
  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   515
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   516
lemma sets_PiM_single: "sets (PiM I M) =
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   517
    sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   518
    (is "_ = sigma_sets ?\<Omega> ?R")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   519
  unfolding sets_PiM
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   520
proof (rule sigma_sets_eqI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   521
  interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   522
  fix A assume "A \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   523
  from prod_algebraE[OF this] guess J X . note X = this
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   524
  show "A \<in> sigma_sets ?\<Omega> ?R"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   525
  proof cases
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   526
    assume "I = {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   527
    with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   528
    with `I = {}` show ?thesis by (auto intro!: sigma_sets_top)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   529
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   530
    assume "I \<noteq> {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   531
    with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^isub>E i\<in>I. space (M i)). f j \<in> X j})"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   532
      using sets_into_space[OF X(5)]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   533
      by (auto simp: prod_emb_PiE[OF _ sets_into_space] Pi_iff split: split_if_asm) blast
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   534
    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   535
      using X `I \<noteq> {}` by (intro R.finite_INT sigma_sets.Basic) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   536
    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   537
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   538
next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   539
  fix A assume "A \<in> ?R"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   540
  then obtain i B where A: "A = {f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)" 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   541
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   542
  then have "A = prod_emb I M {i} (\<Pi>\<^isub>E i\<in>{i}. B)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   543
    using sets_into_space[OF A(3)]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   544
    apply (subst prod_emb_PiE)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   545
    apply (auto simp: Pi_iff split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   546
    apply blast
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   547
    done
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   548
  also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   549
    using A by (intro sigma_sets.Basic prod_algebraI) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   550
  finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   551
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   552
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   553
lemma sets_PiM_I:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   554
  assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   555
  shows "prod_emb I M J (PIE j:J. E j) \<in> sets (PIM i:I. M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   556
proof cases
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   557
  assume "J = {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   558
  then have "prod_emb I M J (PIE j:J. E j) = (PIE j:I. space (M j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   559
    by (auto simp: prod_emb_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   560
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   561
    by (auto simp add: sets_PiM intro!: sigma_sets_top)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   562
next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   563
  assume "J \<noteq> {}" with assms show ?thesis
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   564
    by (force simp add: sets_PiM prod_algebra_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   565
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   566
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   567
lemma measurable_PiM:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   568
  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   569
  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   570
    f -` prod_emb I M J (Pi\<^isub>E J X) \<inter> space N \<in> sets N" 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   571
  shows "f \<in> measurable N (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   572
  using sets_PiM prod_algebra_sets_into_space space
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   573
proof (rule measurable_sigma_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   574
  fix A assume "A \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   575
  from prod_algebraE[OF this] guess J X .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   576
  with sets[of J X] show "f -` A \<inter> space N \<in> sets N" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   577
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   578
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   579
lemma measurable_PiM_Collect:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   580
  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   581
  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   582
    {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N" 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   583
  shows "f \<in> measurable N (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   584
  using sets_PiM prod_algebra_sets_into_space space
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   585
proof (rule measurable_sigma_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   586
  fix A assume "A \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   587
  from prod_algebraE[OF this] guess J X . note X = this
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   588
  have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   589
    using sets_into_space[OF X(5)] X(2-) space unfolding X(1)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   590
    by (subst prod_emb_PiE) (auto simp: Pi_iff split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   591
  also have "\<dots> \<in> sets N" using X(3,2,4,5) by (rule sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   592
  finally show "f -` A \<inter> space N \<in> sets N" .
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   593
qed
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   594
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   595
lemma measurable_PiM_single:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   596
  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   597
  assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N" 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   598
  shows "f \<in> measurable N (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   599
  using sets_PiM_single
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   600
proof (rule measurable_sigma_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   601
  fix A assume "A \<in> {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   602
  then obtain B i where "A = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   603
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   604
  with space have "f -` A \<inter> space N = {\<omega> \<in> space N. f \<omega> i \<in> B}" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   605
  also have "\<dots> \<in> sets N" using B by (rule sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   606
  finally show "f -` A \<inter> space N \<in> sets N" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   607
qed (auto simp: space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   608
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   609
lemma measurable_PiM_single':
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   610
  assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   611
    and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   612
  shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^isub>M I M)"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   613
proof (rule measurable_PiM_single)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   614
  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   615
  then have "{\<omega> \<in> space N. f i \<omega> \<in> A} = f i -` A \<inter> space N"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   616
    by auto
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   617
  then show "{\<omega> \<in> space N. f i \<omega> \<in> A} \<in> sets N"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   618
    using A f by (auto intro!: measurable_sets)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   619
qed fact
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   620
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   621
lemma sets_PiM_I_finite[measurable]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   622
  assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   623
  shows "(PIE j:I. E j) \<in> sets (PIM i:I. M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   624
  using sets_PiM_I[of I I E M] sets_into_space[OF sets] `finite I` sets by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   625
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   626
lemma measurable_component_singleton:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   627
  assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   628
proof (unfold measurable_def, intro CollectI conjI ballI)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   629
  fix A assume "A \<in> sets (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   630
  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = prod_emb I M {i} (\<Pi>\<^isub>E j\<in>{i}. A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   631
    using sets_into_space `i \<in> I`
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   632
    by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: split_if_asm)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   633
  then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   634
    using `A \<in> sets (M i)` `i \<in> I` by (auto intro!: sets_PiM_I)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   635
qed (insert `i \<in> I`, auto simp: space_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   636
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   637
lemma measurable_component_singleton'[measurable_app]:
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   638
  assumes f: "f \<in> measurable N (Pi\<^isub>M I M)"
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   639
  assumes i: "i \<in> I"
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   640
  shows "(\<lambda>x. (f x) i) \<in> measurable N (M i)"
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   641
  using measurable_compose[OF f measurable_component_singleton, OF i] .
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   642
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   643
lemma measurable_PiM_component_rev[measurable (raw)]:
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   644
  "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<lambda>x. f (x i)) \<in> measurable (PiM I M) N"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   645
  by simp
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   646
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   647
lemma measurable_nat_case[measurable (raw)]:
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   648
  assumes [measurable (raw)]: "i = 0 \<Longrightarrow> f \<in> measurable M N"
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   649
    "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N"
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   650
  shows "(\<lambda>x. nat_case (f x) (g x) i) \<in> measurable M N"
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   651
  by (cases i) simp_all
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   652
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   653
lemma measurable_nat_case'[measurable (raw)]:
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   654
  assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   655
  shows "(\<lambda>x. nat_case (f x) (g x)) \<in> measurable N (\<Pi>\<^isub>M i\<in>UNIV. M)"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   656
  using fg[THEN measurable_space]
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   657
  by (auto intro!: measurable_PiM_single' simp add: space_PiM Pi_iff split: nat.split)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   658
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   659
lemma measurable_add_dim[measurable]:
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   660
  "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   661
    (is "?f \<in> measurable ?P ?I")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   662
proof (rule measurable_PiM_single)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   663
  fix j A assume j: "j \<in> insert i I" and A: "A \<in> sets (M j)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   664
  have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} =
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   665
    (if j = i then space (Pi\<^isub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   666
    using sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   667
  also have "\<dots> \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   668
    using A j
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   669
    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   670
  finally show "{\<omega> \<in> space ?P. prod_case (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   671
qed (auto simp: space_pair_measure space_PiM)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   672
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   673
lemma measurable_component_update:
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   674
  "x \<in> space (Pi\<^isub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)"
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   675
  by simp
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   676
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   677
lemma measurable_merge[measurable]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   678
  "merge I J \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   679
    (is "?f \<in> measurable ?P ?U")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   680
proof (rule measurable_PiM_single)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   681
  fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   682
  then have "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} =
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   683
    (if i \<in> I then ((\<lambda>x. x i) \<circ> fst) -` A \<inter> space ?P else ((\<lambda>x. x i) \<circ> snd) -` A \<inter> space ?P)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   684
    by (auto simp: merge_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   685
  also have "\<dots> \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   686
    using A
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   687
    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   688
  finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" .
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   689
qed (auto simp: space_pair_measure space_PiM Pi_iff merge_def extensional_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42950
diff changeset
   690
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   691
lemma measurable_restrict[measurable (raw)]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   692
  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   693
  shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^isub>M I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   694
proof (rule measurable_PiM_single)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   695
  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   696
  then have "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} = X i -` A \<inter> space N"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   697
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   698
  then show "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} \<in> sets N"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   699
    using A X by (auto intro!: measurable_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   700
qed (insert X, auto dest: measurable_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   701
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   702
lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^isub>M L M) (Pi\<^isub>M J M)"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   703
  by (intro measurable_restrict measurable_component_singleton) auto
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   704
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   705
lemma measurable_prod_emb[intro, simp]:
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   706
  "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^isub>M L M)"
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   707
  unfolding prod_emb_def space_PiM[symmetric]
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   708
  by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   709
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   710
lemma sets_in_Pi_aux:
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   711
  "finite I \<Longrightarrow> (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   712
  {x\<in>space (PiM I M). x \<in> Pi I F} \<in> sets (PiM I M)"
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   713
  by (simp add: subset_eq Pi_iff)
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   714
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   715
lemma sets_in_Pi[measurable (raw)]:
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   716
  "finite I \<Longrightarrow> f \<in> measurable N (PiM I M) \<Longrightarrow>
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   717
  (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   718
  Sigma_Algebra.pred N (\<lambda>x. f x \<in> Pi I F)"
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   719
  unfolding pred_def
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   720
  by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_Pi_aux]) auto
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   721
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   722
lemma sets_in_extensional_aux:
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   723
  "{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)"
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   724
proof -
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   725
  have "{x\<in>space (PiM I M). x \<in> extensional I} = space (PiM I M)"
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   726
    by (auto simp add: extensional_def space_PiM)
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   727
  then show ?thesis by simp
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   728
qed
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   729
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   730
lemma sets_in_extensional[measurable (raw)]:
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   731
  "f \<in> measurable N (PiM I M) \<Longrightarrow> Sigma_Algebra.pred N (\<lambda>x. f x \<in> extensional I)"
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   732
  unfolding pred_def
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   733
  by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   734
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   735
locale product_sigma_finite =
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   736
  fixes M :: "'i \<Rightarrow> 'a measure"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   737
  assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   738
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   739
sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   740
  by (rule sigma_finite_measures)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   741
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   742
locale finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   743
  fixes I :: "'i set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   744
  assumes finite_index: "finite I"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   745
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   746
lemma (in finite_product_sigma_finite) sigma_finite_pairs:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   747
  "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   748
    (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   749
    (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and>
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   750
    (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space (PiM I M)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   751
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   752
  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   753
    using M.sigma_finite_incseq by metis
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   754
  from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   755
  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. emeasure (M i) (F i k) \<noteq> \<infinity>"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   756
    by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   757
  let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   758
  note space_PiM[simp]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   759
  show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   760
  proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   761
    fix i show "range (F i) \<subseteq> sets (M i)" by fact
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   762
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   763
    fix i k show "emeasure (M i) (F i k) \<noteq> \<infinity>" by fact
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   764
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   765
    fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   766
      using `\<And>i. range (F i) \<subseteq> sets (M i)` sets_into_space
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   767
      by auto blast
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   768
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   769
    fix f assume "f \<in> space (PiM I M)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   770
    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   771
    show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   772
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   773
    fix i show "?F i \<subseteq> ?F (Suc i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   774
      using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   775
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   776
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   777
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   778
lemma
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   779
  shows space_PiM_empty: "space (Pi\<^isub>M {} M) = {\<lambda>k. undefined}"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   780
    and sets_PiM_empty: "sets (Pi\<^isub>M {} M) = { {}, {\<lambda>k. undefined} }"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   781
  by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   782
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   783
lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   784
proof -
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   785
  let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ereal)"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   786
  have "emeasure (Pi\<^isub>M {} M) (prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = 1"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   787
  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   788
    show "positive (PiM {} M) ?\<mu>"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   789
      by (auto simp: positive_def)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   790
    show "countably_additive (PiM {} M) ?\<mu>"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   791
      by (rule countably_additiveI_finite)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   792
         (auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: )
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   793
  qed (auto simp: prod_emb_def)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   794
  also have "(prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   795
    by (auto simp: prod_emb_def)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   796
  finally show ?thesis
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   797
    by simp
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   798
qed
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   799
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   800
lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   801
  by (rule measure_eqI) (auto simp add: sets_PiM_empty one_ereal_def)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   802
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   803
lemma (in product_sigma_finite) emeasure_PiM:
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   804
  "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   805
proof (induct I arbitrary: A rule: finite_induct)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   806
  case (insert i I)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   807
  interpret finite_product_sigma_finite M I by default fact
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   808
  have "finite (insert i I)" using `finite I` by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   809
  interpret I': finite_product_sigma_finite M "insert i I" by default fact
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   810
  let ?h = "(\<lambda>(f, y). f(i := y))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   811
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   812
  let ?P = "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M) ?h"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   813
  let ?\<mu> = "emeasure ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   814
  let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   815
  let ?f = "\<lambda>J E j. if j \<in> J then emeasure (M j) (E j) else emeasure (M j) (space (M j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   816
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   817
  have "emeasure (Pi\<^isub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^isub>E (insert i I) A)) =
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   818
    (\<Prod>i\<in>insert i I. emeasure (M i) (A i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   819
  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   820
    fix J E assume "(J \<noteq> {} \<or> insert i I = {}) \<and> finite J \<and> J \<subseteq> insert i I \<and> E \<in> (\<Pi> j\<in>J. sets (M j))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   821
    then have J: "J \<noteq> {}" "finite J" "J \<subseteq> insert i I" and E: "\<forall>j\<in>J. E j \<in> sets (M j)" by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   822
    let ?p = "prod_emb (insert i I) M J (Pi\<^isub>E J E)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   823
    let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^isub>E j\<in>J-{i}. E j)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   824
    have "?\<mu> ?p =
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   825
      emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   826
      by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   827
    also have "?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   828
      using J E[rule_format, THEN sets_into_space]
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   829
      by (force simp: space_pair_measure space_PiM Pi_iff prod_emb_iff split: split_if_asm)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   830
    also have "emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   831
      emeasure (Pi\<^isub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   832
      using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   833
    also have "?p' = (\<Pi>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   834
      using J E[rule_format, THEN sets_into_space]
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   835
      by (auto simp: prod_emb_iff Pi_iff split: split_if_asm) blast+
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   836
    also have "emeasure (Pi\<^isub>M I M) (\<Pi>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   837
      (\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   838
      using E by (subst insert) (auto intro!: setprod_cong)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   839
    also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) *
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   840
       emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   841
      using insert by (auto simp: mult_commute intro!: arg_cong2[where f="op *"] setprod_cong)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   842
    also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   843
      using insert(1,2) J E by (intro setprod_mono_one_right) auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   844
    finally show "?\<mu> ?p = \<dots>" .
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   845
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   846
    show "prod_emb (insert i I) M J (Pi\<^isub>E J E) \<in> Pow (\<Pi>\<^isub>E i\<in>insert i I. space (M i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   847
      using J E[rule_format, THEN sets_into_space] by (auto simp: prod_emb_iff)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   848
  next
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   849
    show "positive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   850
      using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   851
  next
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   852
    show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and>
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   853
      insert i I \<subseteq> insert i I \<and> A \<in> (\<Pi> j\<in>insert i I. sets (M j))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   854
      using insert by auto
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   855
  qed (auto intro!: setprod_cong)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   856
  with insert show ?case
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   857
    by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets_into_space)
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   858
qed simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   859
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   860
lemma (in product_sigma_finite) sigma_finite: 
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   861
  assumes "finite I"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   862
  shows "sigma_finite_measure (PiM I M)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   863
proof -
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   864
  interpret finite_product_sigma_finite M I by default fact
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   865
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   866
  from sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   867
  then have F: "\<And>j. j \<in> I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   868
    "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> I. F j k)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   869
    "(\<Union>k. \<Pi>\<^isub>E j \<in> I. F j k) = space (Pi\<^isub>M I M)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   870
    "\<And>k. \<And>j. j \<in> I \<Longrightarrow> emeasure (M j) (F j k) \<noteq> \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   871
    by blast+
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   872
  let ?F = "\<lambda>k. \<Pi>\<^isub>E j \<in> I. F j k"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   873
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   874
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   875
  proof (unfold_locales, intro exI[of _ ?F] conjI allI)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   876
    show "range ?F \<subseteq> sets (Pi\<^isub>M I M)" using F(1) `finite I` by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   877
  next
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   878
    from F(3) show "(\<Union>i. ?F i) = space (Pi\<^isub>M I M)" by simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   879
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   880
    fix j
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   881
    from F `finite I` setprod_PInf[of I, OF emeasure_nonneg, of M]
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   882
    show "emeasure (\<Pi>\<^isub>M i\<in>I. M i) (?F j) \<noteq> \<infinity>"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   883
      by (subst emeasure_PiM) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   884
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   885
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   886
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   887
sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^isub>M I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   888
  using sigma_finite[OF finite_index] .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   889
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   890
lemma (in finite_product_sigma_finite) measure_times:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   891
  "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^isub>M I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   892
  using emeasure_PiM[OF finite_index] by auto
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
   893
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   894
lemma (in product_sigma_finite) positive_integral_empty:
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   895
  assumes pos: "0 \<le> f (\<lambda>k. undefined)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   896
  shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   897
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   898
  interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   899
  have "\<And>A. emeasure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   900
    using assms by (subst measure_times) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   901
  then show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   902
    unfolding positive_integral_def simple_function_def simple_integral_def[abs_def]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   903
  proof (simp add: space_PiM_empty sets_PiM_empty, intro antisym)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   904
    show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   905
      by (intro SUP_upper) (auto simp: le_fun_def split: split_max)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   906
    show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 44890
diff changeset
   907
      by (intro SUP_least) (auto simp: le_fun_def simp: max_def split: split_if_asm)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   908
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   909
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   910
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   911
lemma (in product_sigma_finite) distr_merge:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   912
  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   913
  shows "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M) (merge I J) = Pi\<^isub>M (I \<union> J) M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   914
   (is "?D = ?P")
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   915
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   916
  interpret I: finite_product_sigma_finite M I by default fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   917
  interpret J: finite_product_sigma_finite M J by default fact
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   918
  have "finite (I \<union> J)" using fin by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   919
  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   920
  interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   921
  let ?g = "merge I J"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   922
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   923
  from IJ.sigma_finite_pairs obtain F where
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   924
    F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   925
       "incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   926
       "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   927
       "\<And>k. \<forall>i\<in>I\<union>J. emeasure (M i) (F i k) \<noteq> \<infinity>"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   928
    by auto
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   929
  let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   930
  
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   931
  show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   932
  proof (rule measure_eqI_generator_eq[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   933
    show "Int_stable (prod_algebra (I \<union> J) M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   934
      by (rule Int_stable_prod_algebra)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   935
    show "prod_algebra (I \<union> J) M \<subseteq> Pow (\<Pi>\<^isub>E i \<in> I \<union> J. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   936
      by (rule prod_algebra_sets_into_space)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   937
    show "sets ?P = sigma_sets (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   938
      by (rule sets_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   939
    then show "sets ?D = sigma_sets (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   940
      by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   941
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   942
    show "range ?F \<subseteq> prod_algebra (I \<union> J) M" using F
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   943
      using fin by (auto simp: prod_algebra_eq_finite)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   944
    show "(\<Union>i. \<Pi>\<^isub>E ia\<in>I \<union> J. F ia i) = (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   945
      using F(3) by (simp add: space_PiM)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   946
  next
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   947
    fix k
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   948
    from F `finite I` setprod_PInf[of "I \<union> J", OF emeasure_nonneg, of M]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   949
    show "emeasure ?P (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   950
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   951
    fix A assume A: "A \<in> prod_algebra (I \<union> J) M"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   952
    with fin obtain F where A_eq: "A = (Pi\<^isub>E (I \<union> J) F)" and F: "\<forall>i\<in>J. F i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   953
      by (auto simp add: prod_algebra_eq_finite)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   954
    let ?B = "Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   955
    let ?X = "?g -` A \<inter> space ?B"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   956
    have "Pi\<^isub>E I F \<subseteq> space (Pi\<^isub>M I M)" "Pi\<^isub>E J F \<subseteq> space (Pi\<^isub>M J M)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   957
      using F[rule_format, THEN sets_into_space] by (force simp: space_PiM)+
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   958
    then have X: "?X = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   959
      unfolding A_eq by (subst merge_vimage) (auto simp: space_pair_measure space_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   960
    have "emeasure ?D A = emeasure ?B ?X"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   961
      using A by (intro emeasure_distr measurable_merge) (auto simp: sets_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   962
    also have "emeasure ?B ?X = (\<Prod>i\<in>I. emeasure (M i) (F i)) * (\<Prod>i\<in>J. emeasure (M i) (F i))"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   963
      using `finite J` `finite I` F unfolding X
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   964
      by (simp add: J.emeasure_pair_measure_Times I.measure_times J.measure_times Pi_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   965
    also have "\<dots> = (\<Prod>i\<in>I \<union> J. emeasure (M i) (F i))"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   966
      using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   967
    also have "\<dots> = emeasure ?P (Pi\<^isub>E (I \<union> J) F)"
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   968
      using `finite J` `finite I` F unfolding A
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   969
      by (intro IJ.measure_times[symmetric]) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   970
    finally show "emeasure ?P A = emeasure ?D A" using A_eq by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   971
  qed
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   972
qed
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   973
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   974
lemma (in product_sigma_finite) product_positive_integral_fold:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   975
  assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   976
  and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   977
  shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f =
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   978
    (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
   979
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   980
  interpret I: finite_product_sigma_finite M I by default fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
   981
  interpret J: finite_product_sigma_finite M J by default fact
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41706
diff changeset
   982
  interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   983
  have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   984
    using measurable_comp[OF measurable_merge f] by (simp add: comp_def)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   985
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   986
    apply (subst distr_merge[OF IJ, symmetric])
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   987
    apply (subst positive_integral_distr[OF measurable_merge f])
49999
dfb63b9b8908 for the product measure it is enough if only one measure is sigma-finite
hoelzl
parents: 49784
diff changeset
   988
    apply (subst J.positive_integral_fst_measurable(2)[symmetric, OF P_borel])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   989
    apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   990
    done
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   991
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   992
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   993
lemma (in product_sigma_finite) distr_singleton:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   994
  "distr (Pi\<^isub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   995
proof (intro measure_eqI[symmetric])
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41706
diff changeset
   996
  interpret I: finite_product_sigma_finite M "{i}" by default simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   997
  fix A assume A: "A \<in> sets (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   998
  moreover then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M {i} M) = (\<Pi>\<^isub>E i\<in>{i}. A)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   999
    using sets_into_space by (auto simp: space_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1000
  ultimately show "emeasure (M i) A = emeasure ?D A"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1001
    using A I.measure_times[of "\<lambda>_. A"]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1002
    by (simp add: emeasure_distr measurable_component_singleton)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1003
qed simp
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41706
diff changeset
  1004
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1005
lemma (in product_sigma_finite) product_positive_integral_singleton:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1006
  assumes f: "f \<in> borel_measurable (M i)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1007
  shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1008
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1009
  interpret I: finite_product_sigma_finite M "{i}" by default simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1010
  from f show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1011
    apply (subst distr_singleton[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1012
    apply (subst positive_integral_distr[OF measurable_component_singleton])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1013
    apply simp_all
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1014
    done
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1015
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1016
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1017
lemma (in product_sigma_finite) product_positive_integral_insert:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1018
  assumes I[simp]: "finite I" "i \<notin> I"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1019
    and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1020
  shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1021
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1022
  interpret I: finite_product_sigma_finite M I by default auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1023
  interpret i: finite_product_sigma_finite M "{i}" by default auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1024
  have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1025
    using f by auto
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1026
  show ?thesis
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1027
    unfolding product_positive_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f]
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1028
  proof (rule positive_integral_cong, subst product_positive_integral_singleton[symmetric])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1029
    fix x assume x: "x \<in> space (Pi\<^isub>M I M)"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1030
    let ?f = "\<lambda>y. f (x(i := y))"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1031
    show "?f \<in> borel_measurable (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1032
      using measurable_comp[OF measurable_component_update f, OF x `i \<notin> I`]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1033
      unfolding comp_def .
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1034
    show "(\<integral>\<^isup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^isub>M {i} M) = (\<integral>\<^isup>+ y. f (x(i := y i)) \<partial>Pi\<^isub>M {i} M)"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1035
      using x
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1036
      by (auto intro!: positive_integral_cong arg_cong[where f=f]
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1037
               simp add: space_PiM extensional_def)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1038
  qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1039
qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1040
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1041
lemma (in product_sigma_finite) product_positive_integral_setprod:
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42988
diff changeset
  1042
  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1043
  assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1044
  and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1045
  shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1046
using assms proof induct
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1047
  case (insert i I)
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1048
  note `finite I`[intro, simp]
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1049
  interpret I: finite_product_sigma_finite M I by default auto
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1050
  have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1051
    using insert by (auto intro!: setprod_cong)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1052
  have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1053
    using sets_into_space insert
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1054
    by (intro borel_measurable_ereal_setprod
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1055
              measurable_comp[OF measurable_component_singleton, unfolded comp_def])
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1056
       auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1057
  then show ?case
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1058
    apply (simp add: product_positive_integral_insert[OF insert(1,2) prod])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1059
    apply (simp add: insert(2-) * pos borel setprod_ereal_pos positive_integral_multc)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1060
    apply (subst positive_integral_cmult)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1061
    apply (auto simp add: pos borel insert(2-) setprod_ereal_pos positive_integral_positive)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1062
    done
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1063
qed (simp add: space_PiM)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1064
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1065
lemma (in product_sigma_finite) product_integral_singleton:
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1066
  assumes f: "f \<in> borel_measurable (M i)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1067
  shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1068
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1069
  interpret I: finite_product_sigma_finite M "{i}" by default simp
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42988
diff changeset
  1070
  have *: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M i)"
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42988
diff changeset
  1071
    "(\<lambda>x. ereal (- f x)) \<in> borel_measurable (M i)"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1072
    using assms by auto
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1073
  show ?thesis
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1074
    unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] ..
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1075
qed
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1076
lemma (in product_sigma_finite) distr_component:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1077
  "distr (M i) (Pi\<^isub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^isub>M {i} M" (is "?D = ?P")
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1078
proof (intro measure_eqI[symmetric])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1079
  interpret I: finite_product_sigma_finite M "{i}" by default simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1080
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1081
  have eq: "\<And>x. x \<in> extensional {i} \<Longrightarrow> (\<lambda>j\<in>{i}. x i) = x"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1082
    by (auto simp: extensional_def restrict_def)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1083
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1084
  fix A assume A: "A \<in> sets ?P"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1085
  then have "emeasure ?P A = (\<integral>\<^isup>+x. indicator A x \<partial>?P)" 
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1086
    by simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1087
  also have "\<dots> = (\<integral>\<^isup>+x. indicator ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i)) (x i) \<partial>PiM {i} M)" 
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1088
    by (intro positive_integral_cong) (auto simp: space_PiM indicator_def simp: eq)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1089
  also have "\<dots> = emeasure ?D A"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1090
    using A by (simp add: product_positive_integral_singleton emeasure_distr)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1091
  finally show "emeasure (Pi\<^isub>M {i} M) A = emeasure ?D A" .
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1092
qed simp
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1093
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1094
lemma (in product_sigma_finite) product_integral_fold:
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1095
  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1096
  and f: "integrable (Pi\<^isub>M (I \<union> J) M) f"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1097
  shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)"
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1098
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1099
  interpret I: finite_product_sigma_finite M I by default fact
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1100
  interpret J: finite_product_sigma_finite M J by default fact
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1101
  have "finite (I \<union> J)" using fin by auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1102
  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1103
  interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1104
  let ?M = "merge I J"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1105
  let ?f = "\<lambda>x. f (?M x)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1106
  from f have f_borel: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1107
    by auto
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1108
  have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1109
    using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1110
  have f_int: "integrable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ?f"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1111
    by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f)
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1112
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1113
    apply (subst distr_merge[symmetric, OF IJ fin])
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1114
    apply (subst integral_distr[OF measurable_merge f_borel])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1115
    apply (subst P.integrable_fst_measurable(2)[symmetric, OF f_int])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1116
    apply simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1117
    done
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1118
qed
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1119
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1120
lemma (in product_sigma_finite)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1121
  assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1122
  shows emeasure_fold_integral:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1123
    "emeasure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. emeasure (Pi\<^isub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1124
    and emeasure_fold_measurable:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1125
    "(\<lambda>x. emeasure (Pi\<^isub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1126
proof -
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1127
  interpret I: finite_product_sigma_finite M I by default fact
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1128
  interpret J: finite_product_sigma_finite M J by default fact
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1129
  interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" ..
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1130
  have merge: "merge I J -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1131
    by (intro measurable_sets[OF _ A] measurable_merge assms)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1132
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1133
  show ?I
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1134
    apply (subst distr_merge[symmetric, OF IJ])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1135
    apply (subst emeasure_distr[OF measurable_merge A])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1136
    apply (subst J.emeasure_pair_measure_alt[OF merge])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1137
    apply (auto intro!: positive_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1138
    done
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1139
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1140
  show ?B
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1141
    using IJ.measurable_emeasure_Pair1[OF merge]
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1142
    by (simp add: vimage_compose[symmetric] comp_def space_pair_measure cong: measurable_cong)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1143
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1144
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1145
lemma (in product_sigma_finite) product_integral_insert:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1146
  assumes I: "finite I" "i \<notin> I"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1147
    and f: "integrable (Pi\<^isub>M (insert i I) M) f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1148
  shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1149
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1150
  have "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = integral\<^isup>L (Pi\<^isub>M (I \<union> {i}) M) f"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1151
    by simp
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1152
  also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^isub>M {i} M) \<partial>Pi\<^isub>M I M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1153
    using f I by (intro product_integral_fold) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1154
  also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1155
  proof (rule integral_cong, subst product_integral_singleton[symmetric])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1156
    fix x assume x: "x \<in> space (Pi\<^isub>M I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1157
    have f_borel: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1158
      using f by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1159
    show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1160
      using measurable_comp[OF measurable_component_update f_borel, OF x `i \<notin> I`]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1161
      unfolding comp_def .
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1162
    from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^isub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^isub>M {i} M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1163
      by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1164
  qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1165
  finally show ?thesis .
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1166
qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1167
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1168
lemma (in product_sigma_finite) product_integrable_setprod:
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1169
  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1170
  assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1171
  shows "integrable (Pi\<^isub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1172
proof -
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1173
  interpret finite_product_sigma_finite M I by default fact
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1174
  have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1175
    using integrable unfolding integrable_def by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1176
  have borel: "?f \<in> borel_measurable (Pi\<^isub>M I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1177
    using measurable_comp[OF measurable_component_singleton[of _ I M] f] by (auto simp: comp_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1178
  moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1179
  proof (unfold integrable_def, intro conjI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1180
    show "(\<lambda>x. abs (?f x)) \<in> borel_measurable (Pi\<^isub>M I M)"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1181
      using borel by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1182
    have "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>Pi\<^isub>M I M) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>Pi\<^isub>M I M)"
43920
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42988
diff changeset
  1183
      by (simp add: setprod_ereal abs_setprod)
cedb5cb948fd Rename extreal => ereal
hoelzl
parents: 42988
diff changeset
  1184
    also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. ereal (abs (f i x)) \<partial>M i))"
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1185
      using f by (subst product_positive_integral_setprod) auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1186
    also have "\<dots> < \<infinity>"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1187
      using integrable[THEN integrable_abs]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1188
      by (simp add: setprod_PInf integrable_def positive_integral_positive)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1189
    finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>(Pi\<^isub>M I M)) \<noteq> \<infinity>" by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1190
    have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^isub>M I M)) = (\<integral>\<^isup>+x. 0 \<partial>(Pi\<^isub>M I M))"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1191
      by (intro positive_integral_cong_pos) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1192
    then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^isub>M I M)) \<noteq> \<infinity>" by simp
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1193
  qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1194
  ultimately show ?thesis
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1195
    by (rule integrable_abs_iff[THEN iffD1])
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1196
qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1197
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1198
lemma (in product_sigma_finite) product_integral_setprod:
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1199
  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1200
  assumes "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1201
  shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1202
using assms proof induct
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1203
  case empty
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1204
  interpret finite_measure "Pi\<^isub>M {} M"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1205
    by rule (simp add: space_PiM)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1206
  show ?case by (simp add: space_PiM measure_def)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1207
next
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1208
  case (insert i I)
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1209
  then have iI: "finite (insert i I)" by auto
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1210
  then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1211
    integrable (Pi\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1212
    by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41661
diff changeset
  1213
  interpret I: finite_product_sigma_finite M I by default fact
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1214
  have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1215
    using `i \<notin> I` by (auto intro!: setprod_cong)
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1216
  show ?case
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1217
    unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1218
    by (simp add: * insert integral_multc integral_cmult[OF prod] subset_insertI)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1219
qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1220
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1221
lemma sets_Collect_single:
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1222
  "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^isub>M I M). x i \<in> A } \<in> sets (Pi\<^isub>M I M)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
  1223
  by simp
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1224
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1225
lemma sigma_prod_algebra_sigma_eq_infinite:
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1226
  fixes E :: "'i \<Rightarrow> 'a set set"
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1227
  assumes S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1228
    and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1229
  assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1230
    and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1231
  defines "P == {{f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> E i}"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1232
  shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1233
proof
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1234
  let ?P = "sigma (space (Pi\<^isub>M I M)) P"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1235
  have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>M I M))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1236
    using E_closed by (auto simp: space_PiM P_def Pi_iff subset_eq)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1237
  then have space_P: "space ?P = (\<Pi>\<^isub>E i\<in>I. space (M i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1238
    by (simp add: space_PiM)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1239
  have "sets (PiM I M) =
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1240
      sigma_sets (space ?P) {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1241
    using sets_PiM_single[of I M] by (simp add: space_P)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1242
  also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1243
  proof (safe intro!: sigma_sets_subset)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1244
    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1245
    then have "(\<lambda>x. x i) \<in> measurable ?P (sigma (space (M i)) (E i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1246
      apply (subst measurable_iff_measure_of)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1247
      apply (simp_all add: P_closed)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1248
      using E_closed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1249
      apply (force simp: subset_eq space_PiM)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1250
      apply (force simp: subset_eq space_PiM)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1251
      apply (auto simp: P_def intro!: sigma_sets.Basic exI[of _ i])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1252
      apply (rule_tac x=Aa in exI)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1253
      apply (auto simp: space_PiM)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1254
      done
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1255
    from measurable_sets[OF this, of A] A `i \<in> I` E_closed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1256
    have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1257
      by (simp add: E_generates)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1258
    also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A}"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1259
      using P_closed by (auto simp: space_PiM)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1260
    finally show "\<dots> \<in> sets ?P" .
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1261
  qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1262
  finally show "sets (PiM I M) \<subseteq> sigma_sets (space (PiM I M)) P"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1263
    by (simp add: P_closed)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1264
  show "sigma_sets (space (PiM I M)) P \<subseteq> sets (PiM I M)"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1265
    unfolding P_def space_PiM[symmetric]
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1266
    by (intro sigma_sets_subset) (auto simp: E_generates sets_Collect_single)
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1267
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1268
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1269
lemma sigma_prod_algebra_sigma_eq:
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1270
  fixes E :: "'i \<Rightarrow> 'a set set" and S :: "'i \<Rightarrow> nat \<Rightarrow> 'a set"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1271
  assumes "finite I"
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1272
  assumes S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1273
    and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1274
  assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1275
    and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1276
  defines "P == { Pi\<^isub>E I F | F. \<forall>i\<in>I. F i \<in> E i }"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1277
  shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1278
proof
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1279
  let ?P = "sigma (space (Pi\<^isub>M I M)) P"
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1280
  from `finite I`[THEN ex_bij_betw_finite_nat] guess T ..
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1281
  then have T: "\<And>i. i \<in> I \<Longrightarrow> T i < card I" "\<And>i. i\<in>I \<Longrightarrow> the_inv_into I T (T i) = i"
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1282
    by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1283
  have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>M I M))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1284
    using E_closed by (auto simp: space_PiM P_def Pi_iff subset_eq)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1285
  then have space_P: "space ?P = (\<Pi>\<^isub>E i\<in>I. space (M i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1286
    by (simp add: space_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1287
  have "sets (PiM I M) =
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1288
      sigma_sets (space ?P) {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1289
    using sets_PiM_single[of I M] by (simp add: space_P)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1290
  also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1291
  proof (safe intro!: sigma_sets_subset)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1292
    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1293
    have "(\<lambda>x. x i) \<in> measurable ?P (sigma (space (M i)) (E i))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1294
    proof (subst measurable_iff_measure_of)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1295
      show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1296
      from space_P `i \<in> I` show "(\<lambda>x. x i) \<in> space ?P \<rightarrow> space (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1297
        by (auto simp: Pi_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1298
      show "\<forall>A\<in>E i. (\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1299
      proof
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1300
        fix A assume A: "A \<in> E i"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1301
        then have "(\<lambda>x. x i) -` A \<inter> space ?P = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1302
          using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1303
        also have "\<dots> = (\<Pi>\<^isub>E j\<in>I. \<Union>n. if i = j then A else S j n)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1304
          by (intro PiE_cong) (simp add: S_union)
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1305
        also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>\<^isub>E j\<in>I. if i = j then A else S j (xs ! T j))"
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1306
          using T
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1307
          apply (auto simp: Pi_iff bchoice_iff)
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1308
          apply (rule_tac x="map (\<lambda>n. f (the_inv_into I T n)) [0..<card I]" in exI)
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1309
          apply (auto simp: bij_betw_def)
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1310
          done
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1311
        also have "\<dots> \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1312
        proof (safe intro!: countable_UN)
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1313
          fix xs show "(\<Pi>\<^isub>E j\<in>I. if i = j then A else S j (xs ! T j)) \<in> sets ?P"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1314
            using A S_in_E
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1315
            by (simp add: P_closed)
49779
1484b4b82855 remove incseq assumption from sigma_prod_algebra_sigma_eq
hoelzl
parents: 49776
diff changeset
  1316
               (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j (xs ! T j)"])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1317
        qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1318
        finally show "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1319
          using P_closed by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1320
      qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1321
    qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1322
    from measurable_sets[OF this, of A] A `i \<in> I` E_closed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1323
    have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1324
      by (simp add: E_generates)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1325
    also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1326
      using P_closed by (auto simp: space_PiM)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1327
    finally show "\<dots> \<in> sets ?P" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1328
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1329
  finally show "sets (PiM I M) \<subseteq> sigma_sets (space (PiM I M)) P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1330
    by (simp add: P_closed)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1331
  show "sigma_sets (space (PiM I M)) P \<subseteq> sets (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1332
    using `finite I`
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
  1333
    by (auto intro!: sigma_sets_subset sets_PiM_I_finite simp: E_generates P_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1334
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1335
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1336
lemma pair_measure_eq_distr_PiM:
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1337
  fixes M1 :: "'a measure" and M2 :: "'a measure"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1338
  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1339
  shows "(M1 \<Otimes>\<^isub>M M2) = distr (Pi\<^isub>M UNIV (bool_case M1 M2)) (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. (x True, x False))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1340
    (is "?P = ?D")
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1341
proof (rule pair_measure_eqI[OF assms])
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1342
  interpret B: product_sigma_finite "bool_case M1 M2"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1343
    unfolding product_sigma_finite_def using assms by (auto split: bool.split)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1344
  let ?B = "Pi\<^isub>M UNIV (bool_case M1 M2)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1345
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1346
  have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1347
    by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1348
  fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1349
  have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (bool_case M1 M2 i) (bool_case A B i))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1350
    by (simp add: UNIV_bool ac_simps)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1351
  also have "\<dots> = emeasure ?B (Pi\<^isub>E UNIV (bool_case A B))"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1352
    using A B by (subst B.emeasure_PiM) (auto split: bool.split)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1353
  also have "Pi\<^isub>E UNIV (bool_case A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1354
    using A[THEN sets_into_space] B[THEN sets_into_space]
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1355
    by (auto simp: Pi_iff all_bool_eq space_PiM split: bool.split)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1356
  finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1357
    using A B
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1358
      measurable_component_singleton[of True UNIV "bool_case M1 M2"]
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1359
      measurable_component_singleton[of False UNIV "bool_case M1 M2"]
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1360
    by (subst emeasure_distr) (auto simp: measurable_pair_iff)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1361
qed simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1362
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1363
end