src/HOL/Analysis/Finite_Product_Measure.thy
author nipkow
Tue, 17 Jun 2025 06:29:55 +0200
changeset 82732 71574900b6ba
parent 81097 6c81cdca5b82
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63626
diff changeset
     1
(*  Title:      HOL/Analysis/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
69722
b5163b2132c5 minor tagging updates in 13 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 69687
diff changeset
     5
section \<open>Finite Product Measure\<close>
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
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
     8
imports Binary_Product_Measure Function_Topology
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
     9
begin
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
    10
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    11
lemma Pi_choice: "(\<forall>i\<in>I. \<exists>x\<in>F i. P i x) \<longleftrightarrow> (\<exists>f\<in>Pi I F. \<forall>i\<in>I. P i (f i))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    12
  by (metis Pi_iff)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    13
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    14
lemma PiE_choice: "(\<forall>i\<in>I. \<exists>x\<in>F i. P i x) \<longleftrightarrow>(\<exists>f\<in>Pi\<^sub>E I F. \<forall>i\<in>I. P i (f i))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
    15
  unfolding Pi_choice by (metis Int_iff PiE_def restrict_PiE restrict_apply)
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
    16
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    17
lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    18
  by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    19
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
    20
subsection\<^marker>\<open>tag unimportant\<close> \<open>More about Function restricted by \<^const>\<open>extensional\<close>\<close>
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
    21
35833
7b7ae5aa396d Added product measure space
hoelzl
parents:
diff changeset
    22
definition
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    23
  "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
    24
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    25
lemma merge_apply[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    26
  "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
    27
  "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
    28
  "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
    29
  "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
    30
  "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
    31
  unfolding merge_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    32
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    33
lemma merge_commute:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    34
  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) = merge J I (y, x)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
    35
  by (force simp: merge_def)
40859
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 Pi_cancel_merge_range[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    38
  "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
    39
  "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
    40
  "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
    41
  "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
    42
  by (auto simp: Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    43
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    44
lemma Pi_cancel_merge[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    45
  "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
    46
  "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
    47
  "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
    48
  "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
    49
  by (auto simp: Pi_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    50
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    51
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
    52
  by (auto simp: extensional_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    53
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    54
lemma restrict_merge[simp]:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
    55
  "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
    56
  "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
    57
  "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
    58
  "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    59
  by (auto simp: restrict_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
    60
50042
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    61
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
    62
  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
    63
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    64
lemma PiE_cancel_merge[simp]:
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    65
  "I \<inter> J = {} \<Longrightarrow>
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
    66
    merge I J (x, y) \<in> Pi\<^sub>E (I \<union> J) B \<longleftrightarrow> x \<in> Pi I B \<and> y \<in> Pi J B"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    67
  by (auto simp: PiE_def restrict_Pi_cancel)
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    68
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    69
lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)"
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    70
  unfolding merge_def by (auto simp: fun_eq_iff)
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    71
50042
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    72
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
    73
  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
    74
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    75
lemma merge_restrict[simp]:
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    76
  "merge I J (restrict x I, y) = merge I J (x, y)"
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    77
  "merge I J (x, restrict y J) = merge I J (x, y)"
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    78
  unfolding merge_def by auto
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    79
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    80
lemma merge_x_x_eq_restrict[simp]:
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    81
  "merge I J (x, x) = restrict x (I \<union> J)"
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    82
  unfolding merge_def by auto
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    83
50042
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    84
lemma injective_vimage_restrict:
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    85
  assumes J: "J \<subseteq> I"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
    86
  and sets: "A \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" and ne: "(\<Pi>\<^sub>E i\<in>I. S i) \<noteq> {}"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
    87
  and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
50042
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    88
  shows "A = B"
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    89
proof  (intro set_eqI)
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    90
  fix x
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    91
  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
    92
  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
    93
  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
    94
  proof cases
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
    95
    assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
    96
    have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
    97
      using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
    98
      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
50042
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
    99
    then show "x \<in> A \<longleftrightarrow> x \<in> B"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   100
      using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   101
      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   102
  qed (use sets in auto)
50042
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   103
qed
6fe18351e9dd moved lemmas into projective_family; added header for theory Projective_Family
immler@in.tum.de
parents: 50041
diff changeset
   104
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   105
lemma restrict_vimage:
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   106
  "I \<inter> J = {} \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   107
    (\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^sub>E I E \<times> Pi\<^sub>E J F) = Pi (I \<union> J) (merge I J (E, F))"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   108
  by (auto simp: restrict_Pi_cancel PiE_def)
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   109
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   110
lemma merge_vimage:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   111
  "I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   112
  by (auto simp: restrict_Pi_cancel PiE_def)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   113
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   114
subsection \<open>Finite product spaces\<close>
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   115
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   116
definition\<^marker>\<open>tag important\<close> prod_emb where
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   117
  "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (\<Pi>\<^sub>E i\<in>I. space (M i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   118
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   119
lemma prod_emb_iff:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   120
  "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))"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   121
  unfolding prod_emb_def PiE_def by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   122
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   123
lemma
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   124
  shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   125
    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
   126
    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
   127
    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
   128
    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
   129
    and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   130
  by (auto simp: prod_emb_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   131
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   132
lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   133
    prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))"
62390
842917225d56 more canonical names
nipkow
parents: 61988
diff changeset
   134
  by (force simp: prod_emb_def PiE_iff if_split_mem2)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   135
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   136
lemma prod_emb_PiE_same_index[simp]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   137
    "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   138
  by (auto simp: prod_emb_def PiE_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
   139
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   140
lemma prod_emb_trans[simp]:
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   141
  "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"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   142
  by (auto simp add: Int_absorb1 prod_emb_def PiE_def)
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   143
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   144
lemma prod_emb_Pi:
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   145
  assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   146
  shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   147
  using assms sets.space_closed
62390
842917225d56 more canonical names
nipkow
parents: 61988
diff changeset
   148
  by (auto simp: prod_emb_def PiE_iff split: if_split_asm) blast+
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   149
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   150
lemma prod_emb_id:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   151
  "B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   152
  by (auto simp: prod_emb_def subset_eq extensional_restrict)
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   153
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   154
lemma prod_emb_mono:
50041
afe886a04198 removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents: 50038
diff changeset
   155
  "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
   156
  by (auto simp: prod_emb_def)
afe886a04198 removed redundant/unnecessary assumptions from projective_family
immler@in.tum.de
parents: 50038
diff changeset
   157
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   158
definition\<^marker>\<open>tag important\<close> PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   159
  "PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   160
    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   161
    (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   162
    (\<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
   163
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   164
definition\<^marker>\<open>tag important\<close> prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   165
  "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) `
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   166
    {(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
   167
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   168
abbreviation
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   169
  "Pi\<^sub>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
   170
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   171
syntax
81097
6c81cdca5b82 more inner syntax markup: HOL-Analysis;
wenzelm
parents: 80914
diff changeset
   172
  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"
6c81cdca5b82 more inner syntax markup: HOL-Analysis;
wenzelm
parents: 80914
diff changeset
   173
    (\<open>(\<open>indent=3 notation=\<open>binder \<Pi>\<^sub>M\<close>\<close>\<Pi>\<^sub>M _\<in>_./ _)\<close>  10)
80768
c7723cc15de8 more markup for syntax consts;
wenzelm
parents: 78516
diff changeset
   174
syntax_consts
c7723cc15de8 more markup for syntax consts;
wenzelm
parents: 78516
diff changeset
   175
  "_PiM" == PiM
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   176
translations
61988
34b51f436e92 clarified print modes;
wenzelm
parents: 61808
diff changeset
   177
  "\<Pi>\<^sub>M x\<in>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
   178
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   179
lemma extend_measure_cong:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   180
  assumes "\<Omega> = \<Omega>'" "I = I'" "G = G'" "\<And>i. i \<in> I' \<Longrightarrow> \<mu> i = \<mu>' i"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   181
  shows "extend_measure \<Omega> I G \<mu> = extend_measure \<Omega>' I' G' \<mu>'"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   182
  unfolding extend_measure_def by (auto simp add: assms)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   183
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   184
lemma Pi_cong_sets:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   185
    "\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   186
  by auto
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   187
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   188
lemma PiM_cong:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   189
  assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   190
  shows "PiM I M = PiM J N"
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 59425
diff changeset
   191
  unfolding PiM_def
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   192
proof (rule extend_measure_cong, goal_cases)
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 59425
diff changeset
   193
  case 1
7e741e22d7fc tuned proofs;
wenzelm
parents: 59425
diff changeset
   194
  show ?case using assms
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   195
    by (subst assms(1), intro PiE_cong[of J "\<lambda>i. space (M i)" "\<lambda>i. space (N i)"]) simp_all
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   196
next
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 59425
diff changeset
   197
  case 2
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   198
  have "\<And>K. K \<subseteq> J \<Longrightarrow> (\<Pi> j\<in>K. sets (M j)) = (\<Pi> j\<in>K. sets (N j))"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   199
    using assms by (intro Pi_cong_sets) auto
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   200
  thus ?case by (auto simp: assms)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   201
next
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 59425
diff changeset
   202
  case 3
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   203
  show ?case using assms
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   204
    by (intro ext) (auto simp: prod_emb_def dest: PiE_mem)
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   205
next
60580
7e741e22d7fc tuned proofs;
wenzelm
parents: 59425
diff changeset
   206
  case (4 x)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   207
  thus ?case using assms
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
   208
    by (auto intro!: prod.cong split: if_split_asm)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   209
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   210
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   211
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   212
lemma prod_algebra_sets_into_space:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   213
  "prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   214
  by (auto simp: prod_emb_def prod_algebra_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   215
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   216
lemma prod_algebra_eq_finite:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   217
  assumes I: "finite I"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   218
  shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   219
proof (intro iffI set_eqI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   220
  fix A assume "A \<in> ?L"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   221
  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)"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   222
    and A: "A = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   223
    by (auto simp: prod_algebra_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   224
  let ?A = "\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   225
  have A: "A = ?A"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   226
    unfolding A using J by (intro prod_emb_PiE sets.sets_into_space) auto
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   227
  show "A \<in> ?R" unfolding A using J sets.top
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   228
    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
   229
next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   230
  fix A assume "A \<in> ?R"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   231
  then obtain X where A: "A = (\<Pi>\<^sub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   232
  then have A: "A = prod_emb I M I (\<Pi>\<^sub>E i\<in>I. X i)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   233
    by (simp add: prod_emb_PiE_same_index[OF sets.sets_into_space] Pi_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   234
  from X I show "A \<in> ?L" unfolding A
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   235
    by (auto simp: prod_algebra_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   236
qed
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   237
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   238
lemma prod_algebraI:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   239
  "finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i))
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   240
    \<Longrightarrow> prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j) \<in> prod_algebra I M"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   241
  by (auto simp: prod_algebra_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
   242
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   243
lemma prod_algebraI_finite:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   244
  "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   245
  using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   246
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   247
lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   248
proof (safe intro!: Int_stableI)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   249
  fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   250
  then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   251
    by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int)
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   252
qed
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   253
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   254
lemma prod_algebraE:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   255
  assumes A: "A \<in> prod_algebra I M"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   256
  obtains J E where "A = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   257
    "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   258
  using A by (auto simp: prod_algebra_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42950
diff changeset
   259
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   260
lemma prod_algebraE_all:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   261
  assumes A: "A \<in> prod_algebra I M"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   262
  obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   263
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   264
  from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   265
    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
   266
    by (auto simp: prod_algebra_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   267
  from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   268
    using sets.sets_into_space by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   269
  then have "A = (\<Pi>\<^sub>E i\<in>I. if i\<in>J then E i else space (M i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   270
    using A J by (auto simp: prod_emb_PiE)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   271
  moreover have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   272
    using sets.top E by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   273
  ultimately show ?thesis using that by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   274
qed
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   275
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   276
lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   277
  unfolding Int_stable_def
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   278
proof safe
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   279
  fix A assume "A \<in> prod_algebra I M"
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   280
  from prod_algebraE[OF this] obtain J E where A:
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   281
    "A = prod_emb I M J (Pi\<^sub>E J E)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   282
    "finite J"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   283
    "J \<noteq> {} \<or> I = {}"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   284
    "J \<subseteq> I"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   285
    "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   286
    by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   287
  fix B assume "B \<in> prod_algebra I M"
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   288
  from prod_algebraE[OF this] obtain K F where B:
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   289
    "B = prod_emb I M K (Pi\<^sub>E K F)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   290
    "finite K"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   291
    "K \<noteq> {} \<or> I = {}"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   292
    "K \<subseteq> I"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   293
    "\<And>i. i \<in> K \<Longrightarrow> F i \<in> sets (M i)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   294
    by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   295
  have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter>
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   296
      (if i \<in> K then F i else space (M i)))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   297
    unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4)
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   298
      B(5)[THEN sets.sets_into_space]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   299
    apply (subst (1 2 3) prod_emb_PiE)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   300
    apply (simp_all add: subset_eq PiE_Int)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   301
    apply blast
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   302
    apply (intro PiE_cong)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   303
    apply auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   304
    done
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   305
  also have "\<dots> \<in> prod_algebra I M"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   306
    using A B by (auto intro!: prod_algebraI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   307
  finally show "A \<inter> B \<in> prod_algebra I M" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   308
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   309
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   310
proposition prod_algebra_mono:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   311
  assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   312
  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   313
  shows "prod_algebra I E \<subseteq> prod_algebra I F"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   314
proof
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   315
  fix A assume "A \<in> prod_algebra I E"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   316
  then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   317
    and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   318
    and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   319
    by (auto simp: prod_algebra_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   320
  moreover
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   321
  from space have "(\<Pi>\<^sub>E i\<in>I. space (E i)) = (\<Pi>\<^sub>E i\<in>I. space (F i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   322
    by (rule PiE_cong)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   323
  with A have "A = prod_emb I F J (\<Pi>\<^sub>E i\<in>J. G i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   324
    by (simp add: prod_emb_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   325
  moreover
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   326
  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
   327
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   328
  ultimately show "A \<in> prod_algebra I F"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   329
    apply (simp add: prod_algebra_def image_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   330
    apply (intro exI[of _ J] exI[of _ G] conjI)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   331
    apply auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   332
    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
   333
qed
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   334
proposition prod_algebra_cong:
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   335
  assumes "I = J" and "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   336
  shows "prod_algebra I M = prod_algebra J N"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   337
  by (metis assms prod_algebra_mono sets_eq_imp_space_eq subsetI subset_antisym)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   338
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   339
lemma space_in_prod_algebra: "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   340
proof cases
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   341
  assume "I = {}" then show ?thesis
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   342
    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
   343
next
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   344
  assume "I \<noteq> {}"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   345
  then obtain i where "i \<in> I" by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   346
  then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   347
    by (auto simp: prod_emb_def)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   348
  then show ?thesis
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   349
    by (simp add: \<open>i \<in> I\<close> prod_algebraI) 
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   350
qed
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
   351
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   352
lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   353
  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
   354
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   355
lemma prod_emb_subset_PiM[simp]: "prod_emb I M K X \<subseteq> space (PiM I M)"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   356
  by (auto simp: prod_emb_def space_PiM)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   357
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   358
lemma space_PiM_empty_iff[simp]: "space (PiM I M) = {} \<longleftrightarrow>  (\<exists>i\<in>I. space (M i) = {})"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   359
  by (auto simp: space_PiM PiE_eq_empty_iff)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   360
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   361
lemma undefined_in_PiM_empty[simp]: "(\<lambda>x. undefined) \<in> space (PiM {} M)"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   362
  by (auto simp: space_PiM)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   363
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   364
lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   365
  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
   366
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   367
proposition sets_PiM_single: "sets (PiM I M) =
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   368
    sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   369
    (is "_ = sigma_sets ?\<Omega> ?R")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   370
  unfolding sets_PiM
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   371
proof (rule sigma_sets_eqI)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   372
  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
   373
  fix A assume "A \<in> prod_algebra I M"
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   374
  from prod_algebraE[OF this] obtain J X where X:
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   375
    "A = prod_emb I M J (Pi\<^sub>E J X)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   376
    "finite J"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   377
    "J \<noteq> {} \<or> I = {}"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   378
    "J \<subseteq> I"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   379
    "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   380
    by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   381
  show "A \<in> sigma_sets ?\<Omega> ?R"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   382
  proof cases
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   383
    assume "I = {}"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   384
    with X show ?thesis
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   385
      by (metis (no_types, lifting) PiE_cong R.top empty_iff prod_emb_PiE subset_eq)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   386
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   387
    assume "I \<noteq> {}"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   388
    with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   389
      by (auto simp: prod_emb_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   390
    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   391
      using X \<open>I \<noteq> {}\<close> by (intro R.finite_INT sigma_sets.Basic) auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   392
    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   393
  qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   394
next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   395
  fix A assume "A \<in> ?R"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   396
  then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   397
    by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   398
  then have "A = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. B)"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   399
     by (auto simp: prod_emb_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   400
  also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   401
    using A by (intro sigma_sets.Basic prod_algebraI) auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   402
  finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   403
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   404
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   405
lemma sets_PiM_eq_proj:
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   406
  assumes "I \<noteq> {}"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   407
  shows "sets (PiM I M) = sets (SUP i\<in>I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   408
        (is "?lhs = ?rhs")
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   409
proof -
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   410
  have "?lhs = 
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   411
        sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   412
    by (simp add: sets_PiM_single)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   413
  also have "\<dots> = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i))
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   414
                (\<Union>x\<in>I. sets (vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>xa. xa x) (M x)))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   415
    apply (subst arg_cong [of _ _ Sup, OF image_cong, OF refl])
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   416
     apply (rule sets_vimage_algebra2)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   417
    by (auto intro!: arg_cong2[where f=sigma_sets])
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   418
  also have "... = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i))
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   419
     (\<Union> (sets ` (\<lambda>i. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i)) ` I))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   420
    by simp
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   421
  also have "... = ?rhs"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   422
    by (subst sets_Sup_eq[where X="\<Pi>\<^sub>E i\<in>I. space (M i)"]) (use assms in auto)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   423
  finally show ?thesis .
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   424
qed
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   425
69739
nipkow
parents: 69722
diff changeset
   426
lemma
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   427
  shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   428
    and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\<lambda>k. undefined} }"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   429
  by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   430
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   431
proposition sets_PiM_sigma:
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   432
  assumes \<Omega>_cover: "\<And>i. i \<in> I \<Longrightarrow> \<exists>S\<subseteq>E i. countable S \<and> \<Omega> i = \<Union>S"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   433
  assumes E: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (\<Omega> i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   434
  assumes J: "\<And>j. j \<in> J \<Longrightarrow> finite j" "\<Union>J = I"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   435
  defines "P \<equiv> {{f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i} | A j. j \<in> J \<and> A \<in> Pi j E}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   436
  shows "sets (\<Pi>\<^sub>M i\<in>I. sigma (\<Omega> i) (E i)) = sets (sigma (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   437
proof cases
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   438
  assume "I = {}"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   439
  with \<open>\<Union>J = I\<close> have "P = {{\<lambda>_. undefined}} \<or> P = {}"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   440
    by (auto simp: P_def)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   441
  with \<open>I = {}\<close> show ?thesis
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   442
    by (auto simp add: sets_PiM_empty sigma_sets_empty_eq)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   443
next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   444
  let ?F = "\<lambda>i. {(\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega> |A. A \<in> E i}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   445
  assume "I \<noteq> {}"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   446
  then have "sets (Pi\<^sub>M I (\<lambda>i. sigma (\<Omega> i) (E i))) =
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69064
diff changeset
   447
      sets (SUP i\<in>I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<lambda>x. x i) (sigma (\<Omega> i) (E i)))"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   448
    by (subst sets_PiM_eq_proj) (auto simp: space_measure_of_conv)
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69064
diff changeset
   449
  also have "\<dots> = sets (SUP i\<in>I. sigma (Pi\<^sub>E I \<Omega>) (?F i))"
63333
158ab2239496 Probability: show that measures form a complete lattice
hoelzl
parents: 63040
diff changeset
   450
    using E by (intro sets_SUP_cong arg_cong[where f=sets] vimage_algebra_sigma) auto
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   451
  also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   452
    using \<open>I \<noteq> {}\<close> by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   453
  also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) P)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   454
  proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   455
    show "(\<Union>i\<in>I. ?F i) \<subseteq> Pow (Pi\<^sub>E I \<Omega>)" "P \<subseteq> Pow (Pi\<^sub>E I \<Omega>)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   456
      by (auto simp: P_def)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   457
  next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   458
    interpret P: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   459
      by (auto intro!: sigma_algebra_sigma_sets simp: P_def)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   460
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   461
    fix Z assume "Z \<in> (\<Union>i\<in>I. ?F i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   462
    then obtain i A where i: "i \<in> I" "A \<in> E i" and Z_def: "Z = (\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega>"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   463
      by auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   464
    from \<open>i \<in> I\<close> J obtain j where j: "i \<in> j" "j \<in> J" "j \<subseteq> I" "finite j"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   465
      by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   466
    obtain S where S: "\<And>i. i \<in> j \<Longrightarrow> S i \<subseteq> E i" "\<And>i. i \<in> j \<Longrightarrow> countable (S i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   467
      "\<And>i. i \<in> j \<Longrightarrow> \<Omega> i = \<Union>(S i)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   468
      by (metis subset_eq \<Omega>_cover \<open>j \<subseteq> I\<close>)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   469
    define A' where "A' n = n(i := A)" for n
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   470
    then have A'_i: "\<And>n. A' n i = A"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   471
      by simp
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   472
    { fix n assume "n \<in> Pi\<^sub>E (j - {i}) S"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   473
      then have "A' n \<in> Pi j E"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   474
        unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def \<open>A \<in> E i\<close> )
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   475
      with \<open>j \<in> J\<close> have "{f \<in> Pi\<^sub>E I \<Omega>. \<forall>i\<in>j. f i \<in> A' n i} \<in> P"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   476
        by (auto simp: P_def) }
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   477
    note A'_in_P = this
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   478
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   479
    { fix x assume "x i \<in> A" "x \<in> Pi\<^sub>E I \<Omega>"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   480
      with S(3) \<open>j \<subseteq> I\<close> have "\<forall>i\<in>j. \<exists>s\<in>S i. x i \<in> s"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   481
        by (auto simp: PiE_def Pi_def)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   482
      then obtain s where s: "\<And>i. i \<in> j \<Longrightarrow> s i \<in> S i" "\<And>i. i \<in> j \<Longrightarrow> x i \<in> s i"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   483
        by metis
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   484
      with \<open>x i \<in> A\<close> have "\<exists>n\<in>Pi\<^sub>E (j-{i}) S. \<forall>i\<in>j. x i \<in> A' n i"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   485
        by (intro bexI[of _ "restrict (s(i := A)) (j-{i})"]) (auto simp: A'_def split: if_splits) }
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   486
    then have "Z = (\<Union>n\<in>Pi\<^sub>E (j-{i}) S. {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A' n i})"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   487
      unfolding Z_def
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   488
      by (auto simp add: set_eq_iff ball_conj_distrib \<open>i\<in>j\<close> A'_i dest: bspec[OF _ \<open>i\<in>j\<close>]
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   489
               cong: conj_cong)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   490
    also have "\<dots> \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   491
      using \<open>finite j\<close> S(2)
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   492
      by (intro P.countable_UN' countable_PiE) (simp_all add: image_subset_iff A'_in_P)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   493
    finally show "Z \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P" .
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   494
  next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   495
    interpret F: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<Union>i\<in>I. ?F i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   496
      by (auto intro!: sigma_algebra_sigma_sets)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   497
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   498
    fix b assume "b \<in> P"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   499
    then obtain A j where b: "b = {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i}" "j \<in> J" "A \<in> Pi j E"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   500
      by (auto simp: P_def)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   501
    show "b \<in> sigma_sets (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   502
    proof cases
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   503
      assume "j = {}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   504
      with b have "b = (\<Pi>\<^sub>E i\<in>I. \<Omega> i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   505
        by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   506
      then show ?thesis
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   507
        by blast
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   508
    next
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   509
      assume "j \<noteq> {}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   510
      with J b(2,3) have eq: "b = (\<Inter>i\<in>j. ((\<lambda>x. x i) -` A i \<inter> Pi\<^sub>E I \<Omega>))"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   511
        unfolding b(1)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   512
        by (auto simp: PiE_def Pi_def)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   513
      show ?thesis
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   514
        unfolding eq using \<open>A \<in> Pi j E\<close> \<open>j \<in> J\<close> J(2)
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   515
        by (intro F.finite_INT J \<open>j \<in> J\<close> \<open>j \<noteq> {}\<close> sigma_sets.Basic) blast
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   516
    qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   517
  qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   518
  finally show "?thesis" .
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   519
qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   520
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   521
lemma sets_PiM_in_sets:
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   522
  assumes space: "space N = (\<Pi>\<^sub>E i\<in>I. space (M i))"
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   523
  assumes sets: "\<And>i A. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {x\<in>space N. x i \<in> A} \<in> sets N"
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   524
  shows "sets (\<Pi>\<^sub>M i \<in> I. M i) \<subseteq> sets N"
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   525
  unfolding sets_PiM_single space[symmetric]
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   526
  by (intro sets.sigma_sets_subset subsetI) (auto intro: sets)
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   527
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   528
lemma sets_PiM_cong[measurable_cong]:
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 58876
diff changeset
   529
  assumes "I = J" "\<And>i. i \<in> J \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (PiM I M) = sets (PiM J N)"
58606
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   530
  using assms sets_eq_imp_space_eq[OF assms(2)] by (simp add: sets_PiM_single cong: PiE_cong conj_cong)
9c66f7c541fb add Giry monad
hoelzl
parents: 57512
diff changeset
   531
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   532
lemma sets_PiM_I:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   533
  assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   534
  shows "prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   535
proof cases
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   536
  assume "J = {}"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   537
  then have "prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j) = (\<Pi>\<^sub>E j\<in>I. space (M j))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   538
    by (auto simp: prod_emb_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   539
  then show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   540
    by (auto simp add: sets_PiM intro!: sigma_sets_top)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   541
next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   542
  assume "J \<noteq> {}" with assms show ?thesis
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   543
    by (force simp add: sets_PiM prod_algebra_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   544
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   545
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   546
proposition measurable_PiM:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   547
  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   548
  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>
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   549
    f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   550
  shows "f \<in> measurable N (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   551
  using sets_PiM prod_algebra_sets_into_space space
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   552
proof (rule measurable_sigma_sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   553
  fix A assume "A \<in> prod_algebra I M"
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   554
  from prod_algebraE[OF this] obtain J X where
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   555
    "A = prod_emb I M J (Pi\<^sub>E J X)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   556
    "finite J"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   557
    "J \<noteq> {} \<or> I = {}"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   558
    "J \<subseteq> I"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   559
    "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   560
    by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   561
  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
   562
qed
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   563
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   564
lemma measurable_PiM_Collect:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   565
  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   566
  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>
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   567
    {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   568
  shows "f \<in> measurable N (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   569
  using sets_PiM prod_algebra_sets_into_space space
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   570
proof (rule measurable_sigma_sets)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   571
  fix A assume "A \<in> prod_algebra I M"
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   572
  from prod_algebraE[OF this] obtain J X
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   573
    where X:
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   574
      "A = prod_emb I M J (Pi\<^sub>E J X)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   575
      "finite J"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   576
      "J \<noteq> {} \<or> I = {}"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   577
      "J \<subseteq> I"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   578
      "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   579
    by auto
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   580
  then have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}"
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   581
    using space by (auto simp: prod_emb_def del: PiE_I)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   582
  also have "\<dots> \<in> sets N" using X(3,2,4,5) by (rule sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   583
  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
   584
qed
41095
c335d880ff82 cleanup bijectivity btw. product spaces and pairs
hoelzl
parents: 41026
diff changeset
   585
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   586
lemma measurable_PiM_single:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   587
  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   588
  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"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   589
  shows "f \<in> measurable N (PiM I M)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   590
  using sets_PiM_single
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   591
proof (rule measurable_sigma_sets)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   592
  fix A assume "A \<in> {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   593
  then obtain B i where "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   594
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   595
  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
   596
  also have "\<dots> \<in> sets N" using B by (rule sets)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   597
  finally show "f -` A \<inter> space N \<in> sets N" .
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   598
qed (auto simp: space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   599
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   600
lemma measurable_PiM_single':
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   601
  assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   602
    and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   603
  shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^sub>M I M)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   604
proof (rule measurable_PiM_single)
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   605
  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
   606
  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
   607
    by auto
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   608
  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
   609
    using A f by (auto intro!: measurable_sets)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   610
qed fact
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   611
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   612
lemma sets_PiM_I_finite[measurable]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   613
  assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   614
  shows "(\<Pi>\<^sub>E j\<in>I. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   615
  using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] \<open>finite I\<close> sets by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   616
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   617
lemma measurable_component_singleton[measurable (raw)]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   618
  assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (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
   619
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
   620
  fix A assume "A \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   621
  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) = prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. A)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   622
    using sets.sets_into_space \<open>i \<in> I\<close>
62390
842917225d56 more canonical names
nipkow
parents: 61988
diff changeset
   623
    by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: if_split_asm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   624
  then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   625
    using \<open>A \<in> sets (M i)\<close> \<open>i \<in> I\<close> by (auto intro!: sets_PiM_I)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   626
qed (use \<open>i \<in> I\<close> in \<open>auto simp: space_PiM\<close>)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   627
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   628
lemma measurable_component_singleton'[measurable_dest]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   629
  assumes f: "f \<in> measurable N (Pi\<^sub>M I M)"
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59088
diff changeset
   630
  assumes g: "g \<in> measurable L N"
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   631
  assumes i: "i \<in> I"
59353
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59088
diff changeset
   632
  shows "(\<lambda>x. (f (g x)) i) \<in> measurable L (M i)"
f0707dc3d9aa measurability prover: removed app splitting, replaced by more powerful destruction rules
hoelzl
parents: 59088
diff changeset
   633
  using measurable_compose[OF measurable_compose[OF g f] measurable_component_singleton, OF i] .
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   634
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   635
lemma measurable_PiM_component_rev:
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   636
  "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
   637
  by simp
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   638
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   639
lemma measurable_case_nat[measurable (raw)]:
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   640
  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
   641
    "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N"
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   642
  shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N"
50021
d96a3f468203 add support for function application to measurability prover
hoelzl
parents: 50003
diff changeset
   643
  by (cases i) simp_all
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   644
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   645
lemma measurable_case_nat'[measurable (raw)]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   646
  assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   647
  shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   648
  using fg[THEN measurable_space]
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   649
  by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split)
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50087
diff changeset
   650
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   651
lemma measurable_add_dim[measurable]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   652
  "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   653
    (is "?f \<in> measurable ?P ?I")
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   654
proof (rule measurable_PiM_single)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   655
  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
   656
  have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} =
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   657
    (if j = i then space (Pi\<^sub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   658
    using sets.sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   659
  also have "\<dots> \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   660
    using A j
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   661
    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
   662
  finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   663
qed (auto simp: space_pair_measure space_PiM PiE_def)
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   664
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   665
proposition measurable_fun_upd:
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   666
  assumes I: "I = J \<union> {i}"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   667
  assumes f[measurable]: "f \<in> measurable N (PiM J M)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   668
  assumes h[measurable]: "h \<in> measurable N (M i)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   669
  shows "(\<lambda>x. (f x) (i := h x)) \<in> measurable N (PiM I M)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   670
proof (intro measurable_PiM_single')
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   671
  fix j assume "j \<in> I" then show "(\<lambda>\<omega>. ((f \<omega>)(i := h \<omega>)) j) \<in> measurable N (M j)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   672
    unfolding I by (cases "j = i") auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   673
next
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   674
  show "(\<lambda>x. (f x)(i := h x)) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   675
    using I f[THEN measurable_space] h[THEN measurable_space]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   676
    by (auto simp: space_PiM PiE_iff extensional_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   677
qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   678
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   679
lemma measurable_component_update:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   680
  "x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   681
  by simp
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   682
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   683
lemma measurable_merge[measurable]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   684
  "merge I J \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   685
    (is "?f \<in> measurable ?P ?U")
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   686
proof (rule measurable_PiM_single)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   687
  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
   688
  then have "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} =
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   689
    (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
   690
    by (auto simp: merge_def)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   691
  also have "\<dots> \<in> sets ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   692
    using A
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   693
    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   694
  finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" .
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   695
qed (auto simp: space_pair_measure space_PiM PiE_iff merge_def extensional_def)
42988
d8f3fc934ff6 add lemma indep_distribution_eq_measure
hoelzl
parents: 42950
diff changeset
   696
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   697
lemma measurable_restrict[measurable (raw)]:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   698
  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   699
  shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^sub>M I M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   700
proof (rule measurable_PiM_single)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   701
  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   702
  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
   703
    by auto
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   704
  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
   705
    using A X by (auto intro!: measurable_sets)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   706
next
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   707
  show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   708
    using X by (auto simp add: PiE_def dest: measurable_space)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   709
qed 
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   710
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   711
lemma measurable_abs_UNIV:
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   712
  "(\<And>n. (\<lambda>\<omega>. f n \<omega>) \<in> measurable M (N n)) \<Longrightarrow> (\<lambda>\<omega> n. f n \<omega>) \<in> measurable M (PiM UNIV N)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   713
  by (intro measurable_PiM_single) (auto dest: measurable_space)
e7fd64f82876 add various lemmas
hoelzl
parents: 56996
diff changeset
   714
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   715
lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   716
  by (intro measurable_restrict measurable_component_singleton) auto
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   717
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   718
lemma measurable_restrict_subset':
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   719
  assumes "J \<subseteq> L" "\<And>x. x \<in> J \<Longrightarrow> sets (M x) = sets (N x)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   720
  shows "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   721
  by (metis (no_types) assms measurable_cong_sets measurable_restrict_subset sets_PiM_cong)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
   722
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   723
lemma measurable_prod_emb[intro, simp]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   724
  "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)"
50038
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   725
  unfolding prod_emb_def space_PiM[symmetric]
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   726
  by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
8e32c9254535 moved lemmas further up
immler@in.tum.de
parents: 50021
diff changeset
   727
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   728
lemma merge_in_prod_emb:
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   729
  assumes "y \<in> space (PiM I M)" "x \<in> X" and X: "X \<in> sets (Pi\<^sub>M J M)" and "J \<subseteq> I"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   730
  shows "merge J I (x, y) \<in> prod_emb I M J X"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   731
  using assms sets.sets_into_space[OF X]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   732
  by (simp add: merge_def prod_emb_def subset_eq space_PiM PiE_def extensional_restrict Pi_iff
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   733
           cong: if_cong restrict_cong)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   734
     (simp add: extensional_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   735
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   736
lemma prod_emb_eq_emptyD:
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   737
  assumes J: "J \<subseteq> I" and ne: "space (PiM I M) \<noteq> {}" and X: "X \<in> sets (Pi\<^sub>M J M)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   738
    and *: "prod_emb I M J X = {}"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   739
  shows "X = {}"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   740
proof safe
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   741
  fix x assume "x \<in> X"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   742
  obtain \<omega> where "\<omega> \<in> space (PiM I M)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   743
    using ne by blast
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   744
  from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   745
qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   746
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   747
lemma sets_in_Pi_aux:
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   748
  "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
   749
  {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
   750
  by (simp add: subset_eq Pi_iff)
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   751
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   752
lemma sets_in_Pi[measurable (raw)]:
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   753
  "finite I \<Longrightarrow> f \<in> measurable N (PiM I M) \<Longrightarrow>
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   754
  (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents: 50245
diff changeset
   755
  Measurable.pred N (\<lambda>x. f x \<in> Pi I F)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   756
  unfolding pred_def
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   757
  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
   758
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   759
lemma sets_in_extensional_aux:
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   760
  "{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   761
  by (smt (verit) PiE_iff mem_Collect_eq sets.top space_PiM subsetI subset_antisym)
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   762
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   763
lemma sets_in_extensional[measurable (raw)]:
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents: 50245
diff changeset
   764
  "f \<in> measurable N (PiM I M) \<Longrightarrow> Measurable.pred N (\<lambda>x. f x \<in> extensional I)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   765
  unfolding pred_def
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   766
  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
   767
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   768
lemma sets_PiM_I_countable:
61363
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   769
  assumes I: "countable I" and E: "\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i)" shows "Pi\<^sub>E I E \<in> sets (Pi\<^sub>M I M)"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   770
proof cases
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   771
  assume "I \<noteq> {}"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   772
  then have "Pi\<^sub>E I E = (\<Inter>i\<in>I. prod_emb I M {i} (Pi\<^sub>E {i} E))"
61363
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   773
    using E[THEN sets.sets_into_space] by (auto simp: PiE_iff prod_emb_def fun_eq_iff)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   774
  also have "\<dots> \<in> sets (PiM I M)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   775
    using I \<open>I \<noteq> {}\<close> by (simp add: E sets.countable_INT' sets_PiM_I subset_eq)
61363
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   776
  finally show ?thesis .
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   777
qed (simp add: sets_PiM_empty)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   778
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   779
lemma sets_PiM_D_countable:
61363
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   780
  assumes A: "A \<in> PiM I M"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   781
  shows "\<exists>J\<subseteq>I. \<exists>X\<in>PiM J M. countable J \<and> A = prod_emb I M J X"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   782
  using A[unfolded sets_PiM_single]
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   783
proof induction
61363
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   784
  case (Basic A)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   785
  then obtain i X where *: "i \<in> I" "X \<in> sets (M i)" and "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> X}"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   786
    by auto
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   787
  then have A: "A = prod_emb I M {i} (\<Pi>\<^sub>E _\<in>{i}. X)"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   788
    by (auto simp: prod_emb_def)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   789
  then show ?case
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   790
    by (intro exI[of _ "{i}"] conjI bexI[of _ "\<Pi>\<^sub>E _\<in>{i}. X"])
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   791
       (auto intro: countable_finite * sets_PiM_I_finite)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   792
next
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   793
  case Empty then show ?case
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   794
    by (intro exI[of _ "{}"] conjI bexI[of _ "{}"]) auto
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   795
next
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   796
  case (Compl A)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   797
  then obtain J X where "J \<subseteq> I" "X \<in> sets (Pi\<^sub>M J M)" "countable J" "A = prod_emb I M J X"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   798
    by auto
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   799
  then show ?case
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   800
    by (intro exI[of _ J] bexI[of _ "space (PiM J M) - X"] conjI)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   801
       (auto simp add: space_PiM prod_emb_PiE intro!: sets_PiM_I_countable)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   802
next
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   803
  case (Union K)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   804
  obtain J X where J: "\<And>i. J i \<subseteq> I" "\<And>i. countable (J i)" and X: "\<And>i. X i \<in> sets (Pi\<^sub>M (J i) M)"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   805
    and K: "\<And>i. K i = prod_emb I M (J i) (X i)"
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   806
    by (metis Union.IH)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   807
  show ?case
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   808
  proof (intro exI bexI conjI)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   809
    show "(\<Union>i. J i) \<subseteq> I" "countable (\<Union>i. J i)" 
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
   810
      using J by auto
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69260
diff changeset
   811
    with J show "\<Union>(K ` UNIV) = prod_emb I M (\<Union>i. J i) (\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i))"
61363
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   812
      by (simp add: K[abs_def] SUP_upper)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   813
  qed(auto intro: X)
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   814
qed
76ac925927aa measurable sets on product spaces are embeddings of countable products
hoelzl
parents: 61362
diff changeset
   815
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   816
proposition measure_eqI_PiM_finite:
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   817
  assumes [simp]: "finite I" "sets P = PiM I M" "sets Q = PiM I M"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   818
  assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = Q (Pi\<^sub>E I A)"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   819
  assumes A: "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = space (PiM I M)" "\<And>i::nat. P (A i) \<noteq> \<infinity>"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   820
  shows "P = Q"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   821
proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   822
  show "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. P (A i) \<noteq> \<infinity>"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   823
    unfolding space_PiM[symmetric] by fact+
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   824
  fix X assume "X \<in> prod_algebra I M"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   825
  then obtain J E where X: "X = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   826
    and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   827
    by (force elim!: prod_algebraE)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   828
  then show "emeasure P X = emeasure Q X"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   829
    unfolding X by (subst (1 2) prod_emb_Pi) (auto simp: eq)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   830
qed (simp_all add: sets_PiM)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   831
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   832
proposition measure_eqI_PiM_infinite:
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   833
  assumes [simp]: "sets P = PiM I M" "sets Q = PiM I M"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   834
  assumes eq: "\<And>A J. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow>
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   835
    P (prod_emb I M J (Pi\<^sub>E J A)) = Q (prod_emb I M J (Pi\<^sub>E J A))"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   836
  assumes A: "finite_measure P"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   837
  shows "P = Q"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   838
proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   839
  interpret finite_measure P by fact
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   840
  define i where "i = (SOME i. i \<in> I)"
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   841
  have i: "I \<noteq> {} \<Longrightarrow> i \<in> I"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   842
    unfolding i_def by (rule someI_ex) auto
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   843
  define A where "A n =
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   844
    (if I = {} then prod_emb I M {} (\<Pi>\<^sub>E i\<in>{}. {}) else prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)))"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   845
    for n :: nat
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   846
  then show "range A \<subseteq> prod_algebra I M"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   847
    using prod_algebraI[of "{}" I "\<lambda>i. space (M i)" M] by (auto intro!: prod_algebraI i)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   848
  have "\<And>i. A i = space (PiM I M)"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   849
    by (auto simp: prod_emb_def space_PiM PiE_iff A_def i ex_in_conv[symmetric] exI)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   850
  then show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. emeasure P (A i) \<noteq> \<infinity>"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   851
    by (auto simp: space_PiM)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   852
next
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   853
  fix X assume X: "X \<in> prod_algebra I M"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
   854
  then obtain J E where X: "X = prod_emb I M J (\<Pi>\<^sub>E j\<in>J. E j)"
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   855
    and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   856
    by (force elim!: prod_algebraE)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   857
  then show "emeasure P X = emeasure Q X"
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   858
    by (auto intro!: eq)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   859
qed (auto simp: sets_PiM)
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   860
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   861
locale\<^marker>\<open>tag unimportant\<close> product_sigma_finite =
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   862
  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
   863
  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
   864
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   865
sublocale\<^marker>\<open>tag unimportant\<close> 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
   866
  by (rule sigma_finite_measures)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   867
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
   868
locale\<^marker>\<open>tag unimportant\<close> finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   869
  fixes I :: "'i set"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   870
  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
   871
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   872
proposition (in finite_product_sigma_finite) sigma_finite_pairs:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   873
  "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   874
    (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   875
    (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k) \<and>
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   876
    (\<Union>k. \<Pi>\<^sub>E i\<in>I. F i k) = space (PiM I M)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   877
proof -
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   878
  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
   879
    using M.sigma_finite_incseq by metis
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   880
  then obtain F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   881
    where "\<forall>x. range (F x) \<subseteq> sets (M x) \<and> incseq (F x) \<and> \<Union> (range (F x)) = space (M x) \<and> (\<forall>k. emeasure (M x) (F x k) \<noteq> \<infinity>)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   882
    by metis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   883
  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
   884
    by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   885
  let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   886
  note space_PiM[simp]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   887
  show ?thesis
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
   888
  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
   889
    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
   890
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   891
    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
   892
  next
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   893
    fix x assume "x \<in> (\<Union>i. ?F i)" with F(1) show "x \<in> space (PiM I M)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   894
      by (auto simp: PiE_def dest!: sets.sets_into_space)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   895
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   896
    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
   897
    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
   898
    show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def PiE_def)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   899
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   900
    fix i show "?F i \<subseteq> ?F (Suc i)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   901
      using \<open>\<And>i. incseq (F i)\<close>[THEN incseq_SucD] by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   902
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   903
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   904
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   905
lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   906
proof -
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   907
  let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ennreal)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   908
  have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   909
  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   910
    show "positive (PiM {} M) ?\<mu>"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   911
      by (auto simp: positive_def)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   912
    show "countably_additive (PiM {} M) ?\<mu>"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   913
      by (rule sets.countably_additiveI_finite)
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   914
         (auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: )
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   915
  qed (auto simp: prod_emb_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   916
  also have "(prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   917
    by (auto simp: prod_emb_def)
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   918
  finally show ?thesis
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   919
    by simp
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   920
qed
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   921
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   922
lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   923
  by (rule measure_eqI) (auto simp add: sets_PiM_empty)
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
   924
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   925
lemma (in product_sigma_finite) emeasure_PiM:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   926
  "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   927
proof (induct I arbitrary: A rule: finite_induct)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
   928
  case (insert i I)
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
   929
  interpret finite_product_sigma_finite M I by standard fact
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
   930
  have "finite (insert i I)" using \<open>finite I\<close> by auto
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
   931
  interpret I': finite_product_sigma_finite M "insert i I" by standard fact
41661
baf1964bc468 use pre-image measure, instead of image
hoelzl
parents: 41659
diff changeset
   932
  let ?h = "(\<lambda>(f, y). f(i := y))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   933
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   934
  let ?P = "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   935
  let ?\<mu> = "emeasure ?P"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   936
  let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   937
  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
   938
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   939
  have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) =
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   940
    (\<Prod>i\<in>insert i I. emeasure (M i) (A i))"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   941
  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   942
    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
   943
    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
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   944
    let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   945
    let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^sub>E j\<in>J-{i}. E j)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   946
    have "?\<mu> ?p =
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   947
      emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i))"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   948
      by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   949
    also have "?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   950
      using J E[rule_format, THEN sets.sets_into_space]
62390
842917225d56 more canonical names
nipkow
parents: 61988
diff changeset
   951
      by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: if_split_asm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   952
    also have "emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   953
      emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   954
      using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   955
    also have "?p' = (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   956
      using J E[rule_format, THEN sets.sets_into_space]
62390
842917225d56 more canonical names
nipkow
parents: 61988
diff changeset
   957
      by (auto simp: prod_emb_iff PiE_def Pi_iff split: if_split_asm) blast+
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   958
    also have "emeasure (Pi\<^sub>M I M) (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   959
      (\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
   960
      using E by (subst insert) (auto intro!: prod.cong)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   961
    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
   962
       emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)"
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68833
diff changeset
   963
      using insert by (auto simp: mult.commute intro!: arg_cong2[where f="(*)"] prod.cong)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   964
    also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
   965
      using insert(1,2) J E by (intro prod.mono_neutral_right) auto
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   966
    finally show "?\<mu> ?p = \<dots>" .
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   967
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   968
    show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \<in> Pow (\<Pi>\<^sub>E i\<in>insert i I. space (M i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   969
      using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   970
  next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
   971
    show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   972
      using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   973
  next
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   974
    show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and>
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   975
      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
   976
      using insert by auto
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
   977
  qed (auto intro!: prod.cong)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
   978
  with insert show ?case
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   979
    by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space)
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
   980
qed simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   981
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   982
lemma (in product_sigma_finite) PiM_eqI:
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   983
  assumes I[simp]: "finite I" and P: "sets P = PiM I M"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   984
  assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   985
  shows "P = PiM I M"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   986
proof -
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   987
  interpret finite_product_sigma_finite M I
74362
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   988
    by standard fact
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   989
  from sigma_finite_pairs obtain C where C:
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   990
    "\<forall>i\<in>I. range (C i) \<subseteq> sets (M i)" "\<forall>k. \<forall>i\<in>I. emeasure (M i) (C i k) \<noteq> \<infinity>"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   991
    "incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. C i k)" "(\<Union>k. \<Pi>\<^sub>E i\<in>I. C i k) = space (Pi\<^sub>M I M)"
0135a0c77b64 tuned proofs --- avoid 'guess';
wenzelm
parents: 70365
diff changeset
   992
    by auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
   993
  show ?thesis
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   994
  proof (rule measure_eqI_PiM_finite[OF I refl P, symmetric])
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   995
    show "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>M I M) (Pi\<^sub>E I A) = P (Pi\<^sub>E I A)" for A
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   996
      by (simp add: eq emeasure_PiM)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   997
    define A where "A n = (\<Pi>\<^sub>E i\<in>I. C i n)" for n
61362
48d1b147f094 generalize eqI theorems for product measures
hoelzl
parents: 61359
diff changeset
   998
    with C show "range A \<subseteq> prod_algebra I M" "\<And>i. emeasure (Pi\<^sub>M I M) (A i) \<noteq> \<infinity>" "(\<Union>i. A i) = space (PiM I M)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
   999
      by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_prod_eq_top)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1000
  qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1001
qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1002
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1003
lemma (in product_sigma_finite) sigma_finite:
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1004
  assumes "finite I"
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1005
  shows "sigma_finite_measure (PiM I M)"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1006
proof
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1007
  interpret finite_product_sigma_finite M I by standard fact
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1008
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1009
  obtain F where F: "\<And>j. countable (F j)" "\<And>j f. f \<in> F j \<Longrightarrow> f \<in> sets (M j)"
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1010
    "\<And>j f. f \<in> F j \<Longrightarrow> emeasure (M j) f \<noteq> \<infinity>" and
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 69739
diff changeset
  1011
    in_space: "\<And>j. space (M j) = \<Union>(F j)"
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1012
    using sigma_finite_countable by (metis subset_eq)
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
  1013
  moreover have "(\<Union>(Pi\<^sub>E I ` Pi\<^sub>E I F)) = space (Pi\<^sub>M I M)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1014
    using in_space by (auto simp: space_PiM PiE_iff intro!: PiE_choice[THEN iffD1])
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1015
  ultimately show "\<exists>A. countable A \<and> A \<subseteq> sets (Pi\<^sub>M I M) \<and> \<Union>A = space (Pi\<^sub>M I M) \<and> (\<forall>a\<in>A. emeasure (Pi\<^sub>M I M) a \<noteq> \<infinity>)"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
  1016
    by (intro exI[of _ "Pi\<^sub>E I ` Pi\<^sub>E I F"])
57447
87429bdecad5 import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents: 57418
diff changeset
  1017
       (auto intro!: countable_PiE sets_PiM_I_finite
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
  1018
             simp: PiE_iff emeasure_PiM finite_index ennreal_prod_eq_top)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1019
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1020
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1021
sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1022
  using sigma_finite[OF finite_index] .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1023
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1024
lemma (in finite_product_sigma_finite) measure_times:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1025
  "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1026
  using emeasure_PiM[OF finite_index] by auto
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1027
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1028
lemma (in product_sigma_finite) nn_integral_empty:
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1029
  "0 \<le> f (\<lambda>k. undefined) \<Longrightarrow> integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1030
  by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1031
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69918
diff changeset
  1032
lemma\<^marker>\<open>tag important\<close> (in product_sigma_finite) distr_merge:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1033
  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1034
  shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1035
   (is "?D = ?P")
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1036
proof (rule PiM_eqI)
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1037
  interpret I: finite_product_sigma_finite M I by standard fact
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1038
  interpret J: finite_product_sigma_finite M J by standard fact
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1039
  fix A assume A: "\<And>i. i \<in> I \<union> J \<Longrightarrow> A i \<in> sets (M i)"
64910
6108dddad9f0 more symbols via abbrevs;
wenzelm
parents: 64272
diff changeset
  1040
  have *: "(merge I J -` Pi\<^sub>E (I \<union> J) A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)) = Pi\<^sub>E I A \<times> Pi\<^sub>E J A"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1041
    using A[THEN sets.sets_into_space] by (auto simp: space_PiM space_pair_measure)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1042
  from A fin show "emeasure (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) (Pi\<^sub>E (I \<union> J) A) =
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1043
      (\<Prod>i\<in>I \<union> J. emeasure (M i) (A i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1044
    by (subst emeasure_distr)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
  1045
       (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times prod.union_disjoint)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1046
qed (use fin in simp_all)
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1047
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1048
proposition (in product_sigma_finite) product_nn_integral_fold:
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1049
  assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1050
  and f[measurable]: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1051
shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^sub>M J M)) \<partial>(Pi\<^sub>M I M))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1052
        (is "?lhs = ?rhs")
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1053
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1054
  interpret I: finite_product_sigma_finite M I by standard fact
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1055
  interpret J: finite_product_sigma_finite M J by standard fact
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1056
  interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1057
  have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1058
    using measurable_comp[OF measurable_merge f] by (simp add: comp_def)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1059
  have "?lhs = integral\<^sup>N (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) f"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1060
    by (simp add: I.finite_index J.finite_index assms(1) distr_merge)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1061
  also have "... = \<integral>\<^sup>+ x. f (merge I J x) \<partial>(Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1062
    by (simp add: nn_integral_distr)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1063
  also have "... = ?rhs"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1064
    using P.Fubini P.nn_integral_snd by force
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1065
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1066
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1067
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1068
lemma (in product_sigma_finite) distr_singleton:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1069
  "distr (Pi\<^sub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1070
proof (intro measure_eqI[symmetric])
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1071
  interpret I: finite_product_sigma_finite M "{i}" by standard simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1072
  fix A assume A: "A \<in> sets (M i)"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
  1073
  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M {i} M) = (\<Pi>\<^sub>E i\<in>{i}. A)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
  1074
    using sets.sets_into_space by (auto simp: space_PiM)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
  1075
  then show "emeasure (M i) A = emeasure ?D A"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1076
    using A I.measure_times[of "\<lambda>_. A"]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1077
    by (simp add: emeasure_distr measurable_component_singleton)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1078
qed simp
41831
91a2b435dd7a use measure_preserving in ..._vimage lemmas
hoelzl
parents: 41706
diff changeset
  1079
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1080
lemma (in product_sigma_finite) product_nn_integral_singleton:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1081
  assumes f: "f \<in> borel_measurable (M i)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1082
  shows "integral\<^sup>N (Pi\<^sub>M {i} M) (\<lambda>x. f (x i)) = integral\<^sup>N (M i) f"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1083
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1084
  interpret I: finite_product_sigma_finite M "{i}" by standard simp
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1085
  from f show ?thesis
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1086
    by (metis distr_singleton insert_iff measurable_component_singleton nn_integral_distr)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1087
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39098
diff changeset
  1088
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1089
proposition (in product_sigma_finite) product_nn_integral_insert:
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1090
  assumes I[simp]: "finite I" "i \<notin> I"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1091
    and f: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1092
  shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^sub>M I M))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1093
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1094
  interpret I: finite_product_sigma_finite M I by standard auto
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1095
  interpret i: finite_product_sigma_finite M "{i}" by standard 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
  1096
  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
  1097
    using f by auto
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1098
  show ?thesis
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1099
    unfolding product_nn_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f]
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1100
  proof (rule nn_integral_cong, subst product_nn_integral_singleton[symmetric])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1101
    fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1102
    let ?f = "\<lambda>y. f (x(i := y))"
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1103
    show "?f \<in> borel_measurable (M i)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
  1104
      using measurable_comp[OF measurable_component_update f, OF x \<open>i \<notin> I\<close>]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1105
      unfolding comp_def .
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1106
    show "(\<integral>\<^sup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^sub>M {i} M) = (\<integral>\<^sup>+ y. f (x(i := y i)) \<partial>Pi\<^sub>M {i} M)"
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49779
diff changeset
  1107
      using x
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1108
      by (auto intro!: nn_integral_cong arg_cong[where f=f]
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
  1109
               simp add: space_PiM extensional_def PiE_def)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1110
  qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1111
qed
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1112
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1113
lemma (in product_sigma_finite) product_nn_integral_insert_rev:
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1114
  assumes I[simp]: "finite I" "i \<notin> I"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1115
    and [measurable]: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1116
  shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x(i := y)) \<partial>(Pi\<^sub>M I M)) \<partial>(M i))"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1117
  apply (subst product_nn_integral_insert[OF assms])
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1118
  apply (rule pair_sigma_finite.Fubini')
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1119
  apply (simp add: local.sigma_finite pair_sigma_finite.intro sigma_finite_measures)
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1120
  apply measurable
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1121
  done
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1122
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1123
lemma (in product_sigma_finite) product_nn_integral_prod:
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1124
  assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 56994
diff changeset
  1125
  shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>N (M i) (f i))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1126
using assms proof (induction I)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1127
  case (insert i I)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
  1128
  note insert.prems[measurable]
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61565
diff changeset
  1129
  note \<open>finite I\<close>[intro, simp]
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1130
  interpret I: finite_product_sigma_finite M I by standard auto
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1131
  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))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
  1132
    using insert by (auto intro!: prod.cong)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1133
  have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
  1134
    using sets.sets_into_space insert
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64008
diff changeset
  1135
    by (intro borel_measurable_prod_ennreal
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
  1136
              measurable_comp[OF measurable_component_singleton, unfolded comp_def])
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1137
       auto
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41831
diff changeset
  1138
  then show ?case
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1139
    using product_nn_integral_insert[OF insert(1,2)]
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1140
    by (simp add: insert(2-) * nn_integral_multc nn_integral_cmult)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1141
qed (simp add: space_PiM)
41096
843c40bbc379 integral over setprod
hoelzl
parents: 41095
diff changeset
  1142
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1143
proposition (in product_sigma_finite) product_nn_integral_pair:
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 61378
diff changeset
  1144
  assumes [measurable]: "case_prod f \<in> borel_measurable (M x \<Otimes>\<^sub>M M y)"
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1145
  assumes xy: "x \<noteq> y"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1146
  shows "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {x, y} M) = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1147
proof -
59425
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1148
  interpret psm: pair_sigma_finite "M x" "M y"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1149
    unfolding pair_sigma_finite_def using sigma_finite_measures by simp_all
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1150
  have "{x, y} = {y, x}" by auto
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1151
  also have "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {y, x} M) = (\<integral>\<^sup>+y. \<integral>\<^sup>+\<sigma>. f (\<sigma> x) y \<partial>PiM {x} M \<partial>M y)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1152
    using xy by (subst product_nn_integral_insert_rev) simp_all
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1153
  also have "... = (\<integral>\<^sup>+y. \<integral>\<^sup>+x. f x y \<partial>M x \<partial>M y)"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1154
    by (intro nn_integral_cong, subst product_nn_integral_singleton) simp_all
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1155
  also have "... = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))"
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1156
    by (subst psm.nn_integral_snd[symmetric]) simp_all
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1157
  finally show ?thesis .
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1158
qed
c5e79df8cc21 import general thms from Density_Compiler
hoelzl
parents: 59353
diff changeset
  1159
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1160
lemma (in product_sigma_finite) distr_component:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1161
  "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P")
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1162
proof (intro PiM_eqI)
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63333
diff changeset
  1163
  fix A assume A: "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)"
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63333
diff changeset
  1164
  then have "(\<lambda>x. \<lambda>i\<in>{i}. x) -` Pi\<^sub>E {i} A \<inter> space (M i) = A i"
65036
ab7e11730ad8 Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
paulson <lp15@cam.ac.uk>
parents: 64910
diff changeset
  1165
    by (fastforce dest: sets.sets_into_space)
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63333
diff changeset
  1166
  with A show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x)) (Pi\<^sub>E {i} A) = (\<Prod>i\<in>{i}. emeasure (M i) (A i))"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1167
    by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 61169
diff changeset
  1168
qed simp_all
41026
bea75746dc9d folding on arbitrary Lebesgue integrable functions
hoelzl
parents: 41023
diff changeset
  1169
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1170
lemma (in product_sigma_finite)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1171
  assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1172
  shows emeasure_fold_integral:
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1173
    "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" 
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1174
         (is "?lhs = ?rhs")
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1175
    and emeasure_fold_measurable:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1176
    "(\<lambda>x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M))) \<in> borel_measurable (Pi\<^sub>M I M)" (is ?B)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1177
proof -
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1178
  interpret I: finite_product_sigma_finite M I by standard fact
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 61166
diff changeset
  1179
  interpret J: finite_product_sigma_finite M J by standard fact
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1180
  interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" ..
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1181
  have merge: "merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) \<in> sets (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1182
    by (intro measurable_sets[OF _ A] measurable_merge assms)
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1183
  have "?lhs = emeasure (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) A"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1184
    by (simp add: I.finite_index J.finite_index assms(1) distr_merge)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1185
  also have "... = emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M))"
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1186
    by (meson A emeasure_distr measurable_merge)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1187
  also have "... = ?rhs"
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1188
    apply (subst J.emeasure_pair_measure_alt[OF merge])
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1189
    by (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1190
  finally show "?lhs = ?rhs" .
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1191
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1192
  show ?B
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1193
    using IJ.measurable_emeasure_Pair1[OF merge]
56154
f0a927235162 more complete set of lemmas wrt. image and composition
haftmann
parents: 55415
diff changeset
  1194
    by (simp add: vimage_comp comp_def space_pair_measure cong: measurable_cong)
49776
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1195
qed
199d1d5bb17e tuned product measurability
hoelzl
parents: 47761
diff changeset
  1196
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1197
lemma sets_Collect_single:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 50387
diff changeset
  1198
  "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 49999
diff changeset
  1199
  by simp
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1200
lemma pair_measure_eq_distr_PiM:
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1201
  fixes M1 :: "'a measure" and M2 :: "'a measure"
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1202
  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1203
  shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1204
    (is "?P = ?D")
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1205
proof (rule pair_measure_eqI[OF assms])
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1206
  interpret B: product_sigma_finite "case_bool M1 M2"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1207
    unfolding product_sigma_finite_def using assms by (auto split: bool.split)
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1208
  let ?B = "Pi\<^sub>M UNIV (case_bool M1 M2)"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1209
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1210
  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
  1211
    by auto
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1212
  fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1213
  have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (case_bool M1 M2 i) (case_bool A B i))"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1214
    by (simp add: UNIV_bool ac_simps)
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1215
  also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (case_bool A B))"
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1216
    using A B by (subst B.emeasure_PiM) (auto split: bool.split)
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1217
  also have "Pi\<^sub>E UNIV (case_bool A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
  1218
    using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space]
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50104
diff changeset
  1219
    by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split)
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1220
  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
  1221
    using A B
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1222
      measurable_component_singleton[of True UNIV "case_bool M1 M2"]
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
  1223
      measurable_component_singleton[of False UNIV "case_bool M1 M2"]
50104
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1224
    by (subst emeasure_distr) (auto simp: measurable_pair_iff)
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1225
qed simp
de19856feb54 move theorems to be more generally useable
hoelzl
parents: 50099
diff changeset
  1226
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1227
lemma infprod_in_sets[intro]:
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1228
  fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1229
  shows "Pi UNIV E \<in> sets (\<Pi>\<^sub>M i\<in>UNIV::nat set. M i)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1230
proof -
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1231
  have "Pi UNIV E = (\<Inter>i. prod_emb UNIV M {..i} (\<Pi>\<^sub>E j\<in>{..i}. E j))"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1232
    using E E[THEN sets.sets_into_space]
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1233
    by (auto simp: prod_emb_def Pi_iff extensional_def)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1234
  with E show ?thesis by auto
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1235
qed
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
  1236
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1237
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1238
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1239
subsection \<open>Measurability\<close>
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1240
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1241
text \<open>There are two natural sigma-algebras on a product space: the borel sigma algebra,
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1242
generated by open sets in the product, and the product sigma algebra, countably generated by
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1243
products of measurable sets along finitely many coordinates. The second one is defined and studied
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1244
in \<^file>\<open>Finite_Product_Measure.thy\<close>.
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1245
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1246
These sigma-algebra share a lot of natural properties (measurability of coordinates, for instance),
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1247
but there is a fundamental difference: open sets are generated by arbitrary unions, not only
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1248
countable ones, so typically many open sets will not be measurable with respect to the product sigma
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1249
algebra (while all sets in the product sigma algebra are borel). The two sigma algebras coincide
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1250
only when everything is countable (i.e., the product is countable, and the borel sigma algebra in
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1251
the factor is countably generated).
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1252
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1253
In this paragraph, we develop basic measurability properties for the borel sigma algebra, and
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1254
compare it with the product sigma algebra as explained above.
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1255
\<close>
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1256
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1257
lemma measurable_product_coordinates [measurable (raw)]:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1258
  "(\<lambda>x. x i) \<in> measurable borel borel"
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
  1259
by (rule borel_measurable_continuous_onI[OF continuous_on_product_coordinates])
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1260
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1261
lemma measurable_product_then_coordinatewise:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1262
  fixes f::"'a \<Rightarrow> 'b \<Rightarrow> ('c::topological_space)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1263
  assumes [measurable]: "f \<in> borel_measurable M"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1264
  shows "(\<lambda>x. f x i) \<in> borel_measurable M"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1265
proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1266
  have "(\<lambda>x. f x i) = (\<lambda>y. y i) o f"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1267
    unfolding comp_def by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1268
  then show ?thesis by simp
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1269
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1270
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1271
text \<open>To compare the Borel sigma algebra with the product sigma algebra, we give a presentation
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1272
of the product sigma algebra that is more similar to the one we used above for the product
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1273
topology.\<close>
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1274
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1275
lemma sets_PiM_finite:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1276
  "sets (Pi\<^sub>M I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i))
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1277
        {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1278
proof
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1279
  have "{(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}} \<subseteq> sets (Pi\<^sub>M I M)"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1280
  proof clarify
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1281
    fix X assume H: "\<forall>i. X i \<in> sets (M i)" "finite {i. X i \<noteq> space (M i)}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1282
    then have *: "X i \<in> sets (M i)" for i by simp
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1283
    define J where "J = {i \<in> I. X i \<noteq> space (M i)}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1284
    have "finite J" "J \<subseteq> I" unfolding J_def using H by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1285
    define Y where "Y = (\<Pi>\<^sub>E j\<in>J. X j)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1286
    have "prod_emb I M J Y \<in> sets (Pi\<^sub>M I M)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1287
      unfolding Y_def apply (rule sets_PiM_I) using \<open>finite J\<close> \<open>J \<subseteq> I\<close> * by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1288
    moreover have "prod_emb I M J Y = (\<Pi>\<^sub>E i\<in>I. X i)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1289
      unfolding prod_emb_def Y_def J_def using H sets.sets_into_space[OF *]
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1290
      by (auto simp add: PiE_iff, blast)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1291
    ultimately show "Pi\<^sub>E I X \<in> sets (Pi\<^sub>M I M)" by simp
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1292
  qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1293
  then show "sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1294
              \<subseteq> sets (Pi\<^sub>M I M)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1295
    by (metis (mono_tags, lifting) sets.sigma_sets_subset' sets.top space_PiM)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1296
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1297
  have *: "\<exists>X. {f. (\<forall>i\<in>I. f i \<in> space (M i)) \<and> f \<in> extensional I \<and> f i \<in> A} = Pi\<^sub>E I X \<and>
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1298
                (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1299
    if "i \<in> I" "A \<in> sets (M i)" for i A
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1300
  proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1301
    define X where "X = (\<lambda>j. if j = i then A else space (M j))"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1302
    have "{f. (\<forall>i\<in>I. f i \<in> space (M i)) \<and> f \<in> extensional I \<and> f i \<in> A} = Pi\<^sub>E I X"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1303
      unfolding X_def using sets.sets_into_space[OF \<open>A \<in> sets (M i)\<close>] \<open>i \<in> I\<close>
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1304
      by (auto simp add: PiE_iff extensional_def, metis subsetCE, metis)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1305
    moreover have "X j \<in> sets (M j)" for j
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1306
      unfolding X_def using \<open>A \<in> sets (M i)\<close> by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1307
    moreover have "finite {j. X j \<noteq> space (M j)}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1308
      unfolding X_def by simp
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1309
    ultimately show ?thesis by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1310
  qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1311
  show "sets (Pi\<^sub>M I M) \<subseteq> sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {(\<Pi>\<^sub>E i\<in>I. X i) |X. (\<forall>i. X i \<in> sets (M i)) \<and> finite {i. X i \<noteq> space (M i)}}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1312
    unfolding sets_PiM_single
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1313
    by (intro sigma_sets_mono') (auto simp add: PiE_iff *)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1314
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1315
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1316
lemma sets_PiM_subset_borel:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1317
  "sets (Pi\<^sub>M UNIV (\<lambda>_. borel)) \<subseteq> sets borel"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1318
proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1319
  have *: "Pi\<^sub>E UNIV X \<in> sets borel" if [measurable]: "\<And>i. X i \<in> sets borel" "finite {i. X i \<noteq> UNIV}" for X::"'a \<Rightarrow> 'b set"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1320
  proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1321
    define I where "I = {i. X i \<noteq> UNIV}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1322
    have "finite I" unfolding I_def using that by simp
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1323
    have "Pi\<^sub>E UNIV X = (\<Inter>i\<in>I. (\<lambda>x. x i)-`(X i) \<inter> space borel) \<inter> space borel"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1324
      unfolding I_def by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1325
    also have "... \<in> sets borel"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1326
      using that \<open>finite I\<close> by measurable
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1327
    finally show ?thesis by simp
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1328
  qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1329
  then have "{(\<Pi>\<^sub>E i\<in>UNIV. X i) |X::('a \<Rightarrow> 'b set). (\<forall>i. X i \<in> sets borel) \<and> finite {i. X i \<noteq> space borel}} \<subseteq> sets borel"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1330
    by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1331
  then show ?thesis unfolding sets_PiM_finite space_borel
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1332
    by (simp add: * sets.sigma_sets_subset')
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1333
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1334
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1335
proposition sets_PiM_equal_borel:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1336
  "sets (Pi\<^sub>M UNIV (\<lambda>i::('a::countable). borel::('b::second_countable_topology measure))) = sets borel"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1337
proof
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1338
  obtain K::"('a \<Rightarrow> 'b) set set" where K: "topological_basis K" "countable K"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1339
            "\<And>k. k \<in> K \<Longrightarrow> \<exists>X. (k = Pi\<^sub>E UNIV X) \<and> (\<forall>i. open (X i)) \<and> finite {i. X i \<noteq> UNIV}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1340
    using product_topology_countable_basis by fast
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1341
  have *: "k \<in> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))" if "k \<in> K" for k
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1342
  proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1343
    obtain X where H: "k = Pi\<^sub>E UNIV X" "\<And>i. open (X i)" "finite {i. X i \<noteq> UNIV}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1344
      using K(3)[OF \<open>k \<in> K\<close>] by blast
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1345
    show ?thesis unfolding H(1) sets_PiM_finite space_borel using borel_open[OF H(2)] H(3) by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1346
  qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1347
  have **: "U \<in> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))" if "open U" for U::"('a \<Rightarrow> 'b) set"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1348
  proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1349
    obtain B where "B \<subseteq> K" "U = (\<Union>B)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1350
      using \<open>open U\<close> \<open>topological_basis K\<close> by (metis topological_basis_def)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1351
    have "countable B" using \<open>B \<subseteq> K\<close> \<open>countable K\<close> countable_subset by blast
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1352
    moreover have "k \<in> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))" if "k \<in> B" for k
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1353
      using \<open>B \<subseteq> K\<close> * that by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1354
    ultimately show ?thesis unfolding \<open>U = (\<Union>B)\<close> by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1355
  qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1356
  have "sigma_sets UNIV (Collect open) \<subseteq> sets (Pi\<^sub>M UNIV (\<lambda>i::'a. (borel::('b measure))))"
78516
56a408fa2440 substantial tidy-up, shortening many proofs
paulson <lp15@cam.ac.uk>
parents: 74362
diff changeset
  1357
    by (metis "**" mem_Collect_eq open_UNIV sets.sigma_sets_subset' subsetI)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1358
  then show "sets (borel::('a \<Rightarrow> 'b) measure) \<subseteq> sets (Pi\<^sub>M UNIV (\<lambda>_. borel))"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1359
    unfolding borel_def by auto
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1360
qed (simp add: sets_PiM_subset_borel)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1361
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1362
lemma measurable_coordinatewise_then_product:
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1363
  fixes f::"'a \<Rightarrow> ('b::countable) \<Rightarrow> ('c::second_countable_topology)"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1364
  assumes [measurable]: "\<And>i. (\<lambda>x. f x i) \<in> borel_measurable M"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1365
  shows "f \<in> borel_measurable M"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1366
proof -
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1367
  have "f \<in> measurable M (Pi\<^sub>M UNIV (\<lambda>_. borel))"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1368
    by (rule measurable_PiM_single', auto simp add: assms)
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1369
  then show ?thesis using sets_PiM_equal_borel measurable_cong_sets by blast
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1370
qed
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1371
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69745
diff changeset
  1372
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
  1373
end