src/HOL/Probability/Infinite_Product_Measure.thy
author blanchet
Thu, 18 Sep 2014 16:47:40 +0200
changeset 58376 c9d3074f83b3
parent 58184 db1381d811ab
child 58876 1888e3cb8048
permissions -rw-r--r--
moved datatype realizer to 'old_datatype' and colleagues
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     1
(*  Title:      HOL/Probability/Infinite_Product_Measure.thy
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     2
    Author:     Johannes Hölzl, TU München
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     3
*)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     4
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     5
header {*Infinite Product Measure*}
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     6
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     7
theory Infinite_Product_Measure
50039
bfd5198cbe40 added projective_family; generalized generator in product_prob_space to projective_family
immler@in.tum.de
parents: 50038
diff changeset
     8
  imports Probability_Measure Caratheodory Projective_Family
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
     9
begin
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    10
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    11
lemma (in product_prob_space) emeasure_PiM_emb_not_empty:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    12
  assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    13
  shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    14
proof cases
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    15
  assume "finite I" with X show ?thesis by simp
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    16
next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    17
  let ?\<Omega> = "\<Pi>\<^sub>E i\<in>I. space (M i)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    18
  let ?G = generator
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    19
  assume "\<not> finite I"
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44928
diff changeset
    20
  then have I_not_empty: "I \<noteq> {}" by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    21
  interpret G!: algebra ?\<Omega> generator by (rule algebra_generator) fact
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    22
  note mu_G_mono =
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    23
    G.additive_increasing[OF positive_mu_G[OF I_not_empty] additive_mu_G[OF I_not_empty],
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    24
      THEN increasingD]
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    25
  write mu_G  ("\<mu>G")
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    26
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    27
  { fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> ?G"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    28
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    29
    from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    30
      by (metis rev_finite_subset subsetI)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    31
    moreover from Z guess K' X' by (rule generatorE)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    32
    moreover def K \<equiv> "insert k K'"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    33
    moreover def X \<equiv> "emb K K' X'"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    34
    ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^sub>M K M)" "Z = emb I K X"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    35
      "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^sub>M K M) X"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    36
      by (auto simp: subset_insertI)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    37
    let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^sub>M (K - J) M)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    38
    { fix y assume y: "y \<in> space (Pi\<^sub>M J M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    39
      note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    40
      moreover
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    41
      have **: "?M y \<in> sets (Pi\<^sub>M (K - J) M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    42
        using J K y by (intro merge_sets) auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    43
      ultimately
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    44
      have ***: "((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<in> ?G"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    45
        using J K by (intro generatorI) auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    46
      have "\<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^sub>M I M)) = emeasure (Pi\<^sub>M (K - J) M) (?M y)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    47
        unfolding * using K J by (subst mu_G_eq[OF _ _ _ **]) auto
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    48
      note * ** *** this }
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    49
    note merge_in_G = this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    50
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    51
    have "finite (K - J)" using K by auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    52
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    53
    interpret J: finite_product_prob_space M J by default fact+
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    54
    interpret KmJ: finite_product_prob_space M "K - J" by default fact+
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    55
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    56
    have "\<mu>G Z = emeasure (Pi\<^sub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    57
      using K J by simp
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    58
    also have "\<dots> = (\<integral>\<^sup>+ x. emeasure (Pi\<^sub>M (K - J) M) (?M x) \<partial>Pi\<^sub>M J M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    59
      using K J by (subst emeasure_fold_integral) auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    60
    also have "\<dots> = (\<integral>\<^sup>+ y. \<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<partial>Pi\<^sub>M J M)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    61
      (is "_ = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)")
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 55642
diff changeset
    62
    proof (intro nn_integral_cong)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    63
      fix x assume x: "x \<in> space (Pi\<^sub>M J M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    64
      with K merge_in_G(2)[OF this]
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    65
      show "emeasure (Pi\<^sub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    66
        unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst mu_G_eq) auto
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    67
    qed
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    68
    finally have fold: "\<mu>G Z = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)" .
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    69
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    70
    { fix x assume x: "x \<in> space (Pi\<^sub>M J M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    71
      then have "\<mu>G (?MZ x) \<le> 1"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    72
        unfolding merge_in_G(4)[OF x] `Z = emb I K X`
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    73
        by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) }
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    74
    note le_1 = this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    75
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    76
    let ?q = "\<lambda>y. \<mu>G ((\<lambda>x. merge J (I - J) (y,x)) -` Z \<inter> space (Pi\<^sub>M I M))"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
    77
    have "?q \<in> borel_measurable (Pi\<^sub>M J M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    78
      unfolding `Z = emb I K X` using J K merge_in_G(3)
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    79
      by (simp add: merge_in_G  mu_G_eq emeasure_fold_measurable cong: measurable_cong)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    80
    note this fold le_1 merge_in_G(3) }
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    81
  note fold = this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    82
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    83
  have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    84
  proof (rule G.caratheodory_empty_continuous[OF positive_mu_G additive_mu_G])
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    85
    fix A assume "A \<in> ?G"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    86
    with generatorE guess J X . note JX = this
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
    87
    interpret JK: finite_product_prob_space M J by default fact+ 
46898
1570b30ee040 tuned proofs -- eliminated pointless chaining of facts after 'interpret';
wenzelm
parents: 46731
diff changeset
    88
    from JX show "\<mu>G A \<noteq> \<infinity>" by simp
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    89
  next
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
    90
    fix A assume A: "range A \<subseteq> ?G" "decseq A" "(\<Inter>i. A i) = {}"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    91
    then have "decseq (\<lambda>i. \<mu>G (A i))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    92
      by (auto intro!: mu_G_mono simp: decseq_def)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    93
    moreover
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    94
    have "(INF i. \<mu>G (A i)) = 0"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    95
    proof (rule ccontr)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    96
      assume "(INF i. \<mu>G (A i)) \<noteq> 0" (is "?a \<noteq> 0")
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    97
      moreover have "0 \<le> ?a"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
    98
        using A positive_mu_G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
    99
      ultimately have "0 < ?a" by auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   100
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   101
      have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^sub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (limP J M (\<lambda>J. (Pi\<^sub>M J M))) X"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   102
        using A by (intro allI generator_Ex) auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   103
      then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^sub>M (J' n) M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   104
        and A': "\<And>n. A n = emb I (J' n) (X' n)"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   105
        unfolding choice_iff by blast
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   106
      moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   107
      moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   108
      ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^sub>M (J n) M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   109
        by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   110
      with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   111
        unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   112
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   113
      have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   114
        unfolding J_def by force
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   115
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   116
      interpret J: finite_product_prob_space M "J i" for i by default fact+
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   117
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   118
      have a_le_1: "?a \<le> 1"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
   119
        using mu_G_spec[of "J 0" "A 0" "X 0"] J A_eq
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 43920
diff changeset
   120
        by (auto intro!: INF_lower2[of 0] J.measure_le_1)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   121
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   122
      let ?M = "\<lambda>K Z y. (\<lambda>x. merge K (I - K) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   123
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   124
      { fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   125
        then have Z_sets: "\<And>n. Z n \<in> ?G" by auto
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   126
        fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   127
        interpret J': finite_product_prob_space M J' by default fact+
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   128
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   129
        let ?q = "\<lambda>n y. \<mu>G (?M J' (Z n) y)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   130
        let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^sub>M J' M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   131
        { fix n
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   132
          have "?q n \<in> borel_measurable (Pi\<^sub>M J' M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   133
            using Z J' by (intro fold(1)) auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   134
          then have "?Q n \<in> sets (Pi\<^sub>M J' M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   135
            by (rule measurable_sets) auto }
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   136
        note Q_sets = this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   137
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   138
        have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^sub>M J' M) (?Q n))"
44928
7ef6505bde7f renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents: 43920
diff changeset
   139
        proof (intro INF_greatest)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   140
          fix n
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   141
          have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   142
          also have "\<dots> \<le> (\<integral>\<^sup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^sub>M J' M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   143
            unfolding fold(2)[OF J' `Z n \<in> ?G`]
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 55642
diff changeset
   144
          proof (intro nn_integral_mono)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   145
            fix x assume x: "x \<in> space (Pi\<^sub>M J' M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   146
            then have "?q n x \<le> 1 + 0"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   147
              using J' Z fold(3) Z_sets by auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   148
            also have "\<dots> \<le> 1 + ?a / 2^(k+1)"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   149
              using `0 < ?a` by (intro add_mono) auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   150
            finally have "?q n x \<le> 1 + ?a / 2^(k+1)" .
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   151
            with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   152
              by (auto split: split_indicator simp del: power_Suc)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   153
          qed
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   154
          also have "\<dots> = emeasure (Pi\<^sub>M J' M) (?Q n) + ?a / 2^(k+1)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   155
            using `0 \<le> ?a` Q_sets J'.emeasure_space_1
56996
891e992e510f renamed positive_integral to nn_integral
hoelzl
parents: 55642
diff changeset
   156
            by (subst nn_integral_add) auto
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   157
          finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^sub>M J' M) (?Q n)" using `?a \<le> 1`
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   158
            by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^sub>M J' M) (?Q n)"])
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   159
               (auto simp: field_simps)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   160
        qed
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   161
        also have "\<dots> = emeasure (Pi\<^sub>M J' M) (\<Inter>n. ?Q n)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   162
        proof (intro INF_emeasure_decseq)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   163
          show "range ?Q \<subseteq> sets (Pi\<^sub>M J' M)" using Q_sets by auto
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   164
          show "decseq ?Q"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   165
            unfolding decseq_def
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   166
          proof (safe intro!: vimageI[OF refl])
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   167
            fix m n :: nat assume "m \<le> n"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   168
            fix x assume x: "x \<in> space (Pi\<^sub>M J' M)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   169
            assume "?a / 2^(k+1) \<le> ?q n x"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   170
            also have "?q n x \<le> ?q m x"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
   171
            proof (rule mu_G_mono)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   172
              from fold(4)[OF J', OF Z_sets x]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   173
              show "?M J' (Z n) x \<in> ?G" "?M J' (Z m) x \<in> ?G" by auto
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   174
              show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   175
                using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   176
            qed
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   177
            finally show "?a / 2^(k+1) \<le> ?q m x" .
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   178
          qed
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   179
        qed simp
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   180
        finally have "(\<Inter>n. ?Q n) \<noteq> {}"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   181
          using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   182
        then have "\<exists>w\<in>space (Pi\<^sub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   183
      note Ex_w = this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   184
46731
5302e932d1e5 avoid undeclared variables in let bindings;
wenzelm
parents: 45777
diff changeset
   185
      let ?q = "\<lambda>k n y. \<mu>G (?M (J k) (A n) y)"
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   186
58184
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   187
      let ?P = "\<lambda>w k. w \<in> space (Pi\<^sub>M (J k) M) \<and> (\<forall>n. ?a / 2 ^ (Suc k) \<le> ?q k n w)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   188
      have "\<exists>w. \<forall>k. ?P (w k) k \<and> restrict (w (Suc k)) (J k) = w k"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   189
      proof (rule dependent_nat_choice)
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   190
        have "\<forall>n. ?a / 2 ^ 0 \<le> \<mu>G (A n)" by (auto intro: INF_lower)
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   191
        from Ex_w[OF A(1,2) this J(1-3), of 0] show "\<exists>w. ?P w 0" by auto
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   192
      next
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   193
        fix w k assume Suc: "?P w k"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   194
        show "\<exists>w'. ?P w' (Suc k) \<and> restrict w' (J k) = w"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   195
        proof cases
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   196
          assume [simp]: "J k = J (Suc k)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   197
          have "?a / 2 ^ (Suc (Suc k)) \<le> ?a / 2 ^ (k + 1)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   198
            using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: field_simps)
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   199
          with Suc show ?thesis
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   200
            by (auto intro!: exI[of _ w] simp: extensional_restrict space_PiM intro: order_trans)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   201
        next
58184
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   202
          assume "J k \<noteq> J (Suc k)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   203
          with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   204
          have "range (\<lambda>n. ?M (J k) (A n) w) \<subseteq> ?G" "decseq (\<lambda>n. ?M (J k) (A n) w)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   205
            "\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) w)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   206
            using `decseq A` fold(4)[OF J(1-3) A_eq(2), of w k] Suc
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   207
            by (auto simp: decseq_def)
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   208
          from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"]
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   209
          obtain w' where w': "w' \<in> space (Pi\<^sub>M ?D M)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   210
            "\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) w) w')" by auto
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   211
          let ?w = "merge (J k) ?D (w, w')"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   212
          have [simp]: "\<And>x. merge (J k) (I - J k) (w, merge ?D (I - ?D) (w', x)) =
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   213
            merge (J (Suc k)) (I - (J (Suc k))) (?w, x)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   214
            using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"]
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   215
            by (auto intro!: ext split: split_merge)
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   216
          have *: "\<And>n. ?M ?D (?M (J k) (A n) w) w' = ?M (J (Suc k)) (A n) ?w"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   217
            using w'(1) J(3)[of "Suc k"]
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   218
            by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   219
          show ?thesis
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   220
            using Suc w' J_mono[of k "Suc k"] unfolding *
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   221
            by (intro exI[of _ ?w])
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   222
               (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM PiE_iff)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   223
        qed
58184
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   224
      qed
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   225
      then obtain w where w:
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   226
        "\<And>k. w k \<in> space (Pi\<^sub>M (J k) M)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   227
        "\<And>k n. ?a / 2 ^ (Suc k) \<le> ?q k n (w k)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   228
        "\<And>k. restrict (w (Suc k)) (J k) = w k"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   229
        by metis
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   230
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   231
      { fix k
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   232
        from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   233
        then have "?M (J k) (A k) (w k) \<noteq> {}"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
   234
          using positive_mu_G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \<le> 1`
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   235
          by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   236
        then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto
49780
92a58f80b20c merge should operate on pairs
hoelzl
parents: 49776
diff changeset
   237
        then have "merge (J k) (I - J k) (w k, x) \<in> A k" by auto
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   238
        then have "\<exists>x\<in>A k. restrict x (J k) = w k"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   239
          using `w k \<in> space (Pi\<^sub>M (J k) M)`
58184
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   240
          by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM) }
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   241
      note w_ext = this
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   242
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   243
      { fix k l i assume "k \<le> l" "i \<in> J k"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   244
        { fix l have "w k i = w (k + l) i"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   245
          proof (induct l)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   246
            case (Suc l)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   247
            from `i \<in> J k` J_mono[of k "k + l"] have "i \<in> J (k + l)" by auto
58184
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   248
            with w(3)[of "k + l"]
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   249
            have "w (k + l) i = w (k + Suc l) i"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   250
              by (auto simp: restrict_def fun_eq_iff split: split_if_asm)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   251
            with Suc show ?case by simp
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   252
          qed simp }
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   253
        from this[of "l - k"] `k \<le> l` have "w l i = w k i" by simp }
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   254
      note w_mono = this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   255
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   256
      def w' \<equiv> "\<lambda>i. if i \<in> (\<Union>k. J k) then w (LEAST k. i \<in> J k) i else if i \<in> I then (SOME x. x \<in> space (M i)) else undefined"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   257
      { fix i k assume k: "i \<in> J k"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   258
        have "w k i = w (LEAST k. i \<in> J k) i"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   259
          by (intro w_mono Least_le k LeastI[of _ k])
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   260
        then have "w' i = w k i"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   261
          unfolding w'_def using k by auto }
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   262
      note w'_eq = this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   263
      have w'_simps1: "\<And>i. i \<notin> I \<Longrightarrow> w' i = undefined"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   264
        using J by (auto simp: w'_def)
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   265
      have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)"
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   266
        using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]])
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   267
      { fix i assume "i \<in> I" then have "w' i \<in> space (M i)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   268
          using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ }
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   269
      note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   270
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   271
      have w': "w' \<in> space (Pi\<^sub>M I M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   272
        using w(1) by (auto simp add: Pi_iff extensional_def space_PiM)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   273
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   274
      { fix n
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50099
diff changeset
   275
        have "restrict w' (J n) = w n" using w(1)[of n]
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50099
diff changeset
   276
          by (auto simp add: fun_eq_iff space_PiM)
58184
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   277
        with w_ext[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)"
db1381d811ab cleanup Wfrec; introduce dependent_wf/wellorder_choice
hoelzl
parents: 57512
diff changeset
   278
          by auto
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   279
        then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) }
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   280
      then have "w' \<in> (\<Inter>i. A i)" by auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   281
      with `(\<Inter>i. A i) = {}` show False by auto
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   282
    qed
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   283
    ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"
51351
dd1dd470690b generalized lemmas in Extended_Real_Limits
hoelzl
parents: 50252
diff changeset
   284
      using LIMSEQ_INF[of "\<lambda>i. \<mu>G (A i)"] by simp
45777
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44928
diff changeset
   285
  qed fact+
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44928
diff changeset
   286
  then guess \<mu> .. note \<mu> = this
c36637603821 remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents: 44928
diff changeset
   287
  show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   288
  proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \<mu> J X])
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   289
    from assms show "(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
   290
      by (simp add: Pi_iff)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   291
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   292
    fix J X assume J: "(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: 51351
diff changeset
   293
    then show "emb I J (Pi\<^sub>E J X) \<in> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   294
      by (auto simp: Pi_iff prod_emb_def dest: sets.sets_into_space)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   295
    have "emb I J (Pi\<^sub>E J X) \<in> generator"
50003
8c213922ed49 use measurability prover
hoelzl
parents: 50000
diff changeset
   296
      using J `I \<noteq> {}` by (intro generatorI') (auto simp: Pi_iff)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   297
    then have "\<mu> (emb I J (Pi\<^sub>E J X)) = \<mu>G (emb I J (Pi\<^sub>E J X))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   298
      using \<mu> by simp
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   299
    also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
50252
4aa34bd43228 eliminated slightly odd identifiers;
wenzelm
parents: 50244
diff changeset
   300
      using J  `I \<noteq> {}` by (subst mu_G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   301
    also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}.
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   302
      if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56996
diff changeset
   303
      using J `I \<noteq> {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   304
    finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = \<dots>" .
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   305
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   306
    let ?F = "\<lambda>j. 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
   307
    have "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = (\<Prod>j\<in>J. ?F j)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56996
diff changeset
   308
      using X `I \<noteq> {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1)
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   309
    then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) =
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   310
      emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   311
      using X by (auto simp add: emeasure_PiM) 
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   312
  next
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   313
    show "positive (sets (Pi\<^sub>M I M)) \<mu>" "countably_additive (sets (Pi\<^sub>M I M)) \<mu>"
49804
ace9b5a83e60 infprod generator works also with empty index set
hoelzl
parents: 49784
diff changeset
   314
      using \<mu> unfolding sets_PiM_generator by (auto simp: measure_space_def)
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   315
  qed
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   316
qed
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   317
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   318
sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^sub>M I M"
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   319
proof
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   320
  show "emeasure (Pi\<^sub>M I M) (space (Pi\<^sub>M I M)) = 1"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   321
  proof cases
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   322
    assume "I = {}" then show ?thesis by (simp add: space_PiM_empty)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   323
  next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   324
    assume "I \<noteq> {}"
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   325
    then obtain i where i: "i \<in> I" by auto
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   326
    then have "emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)) = (space (Pi\<^sub>M I M))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   327
      by (auto simp: prod_emb_def space_PiM)
53374
a14d2a854c02 tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents: 53015
diff changeset
   328
    with i show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   329
      using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   330
      by (simp add: emeasure_PiM emeasure_space_1)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   331
  qed
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   332
qed
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   333
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   334
lemma (in product_prob_space) emeasure_PiM_emb:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   335
  assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   336
  shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   337
proof cases
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   338
  assume "J = {}"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   339
  moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^sub>M I M)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   340
    by (auto simp: space_PiM prod_emb_def)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   341
  ultimately show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   342
    by (simp add: space_PiM_empty P.emeasure_space_1)
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   343
next
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   344
  assume "J \<noteq> {}" with X show ?thesis
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   345
    by (subst emeasure_PiM_emb_not_empty) (auto simp: emeasure_PiM)
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   346
qed
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   347
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   348
lemma (in product_prob_space) emeasure_PiM_Collect:
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   349
  assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   350
  shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = (\<Prod> i\<in>J. emeasure (M i) (X i))"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   351
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   352
  have "{x\<in>space (Pi\<^sub>M I M). \<forall>i\<in>J. x i \<in> X i} = emb I J (Pi\<^sub>E J X)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   353
    unfolding prod_emb_def using assms by (auto simp: space_PiM Pi_iff)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   354
  with emeasure_PiM_emb[OF assms] show ?thesis by simp
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   355
qed
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   356
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   357
lemma (in product_prob_space) emeasure_PiM_Collect_single:
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   358
  assumes X: "i \<in> I" "A \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   359
  shows "emeasure (Pi\<^sub>M I M) {x\<in>space (Pi\<^sub>M I M). x i \<in> A} = emeasure (M i) A"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   360
  using emeasure_PiM_Collect[of "{i}" "\<lambda>i. A"] assms
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   361
  by simp
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   362
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   363
lemma (in product_prob_space) measure_PiM_emb:
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   364
  assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   365
  shows "measure (PiM I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   366
  using emeasure_PiM_emb[OF assms]
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   367
  unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal)
42865
36353d913424 add some lemmas for infinite product measure
hoelzl
parents: 42257
diff changeset
   368
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   369
lemma sets_Collect_single':
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   370
  "i \<in> I \<Longrightarrow> {x\<in>space (M i). P x} \<in> sets (M i) \<Longrightarrow> {x\<in>space (PiM I M). P (x i)} \<in> sets (PiM I M)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   371
  using sets_Collect_single[of i I "{x\<in>space (M i). P x}" M]
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50099
diff changeset
   372
  by (simp add: space_PiM PiE_iff cong: conj_cong)
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   373
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   374
lemma (in finite_product_prob_space) finite_measure_PiM_emb:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   375
  "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   376
  using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets.sets_into_space, of I A M]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   377
  by auto
42865
36353d913424 add some lemmas for infinite product measure
hoelzl
parents: 42257
diff changeset
   378
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   379
lemma (in product_prob_space) PiM_component:
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   380
  assumes "i \<in> I"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   381
  shows "distr (PiM I M) (M i) (\<lambda>\<omega>. \<omega> i) = M i"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   382
proof (rule measure_eqI[symmetric])
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   383
  fix A assume "A \<in> sets (M i)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   384
  moreover have "((\<lambda>\<omega>. \<omega> i) -` A \<inter> space (PiM I M)) = {x\<in>space (PiM I M). x i \<in> A}"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   385
    by auto
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   386
  ultimately show "emeasure (M i) A = emeasure (distr (PiM I M) (M i) (\<lambda>\<omega>. \<omega> i)) A"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   387
    by (auto simp: `i\<in>I` emeasure_distr measurable_component_singleton emeasure_PiM_Collect_single)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   388
qed simp
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   389
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   390
lemma (in product_prob_space) PiM_eq:
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   391
  assumes "I \<noteq> {}"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   392
  assumes "sets M' = sets (PiM I M)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   393
  assumes eq: "\<And>J F. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>j. j \<in> J \<Longrightarrow> F j \<in> sets (M j)) \<Longrightarrow>
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   394
    emeasure M' (prod_emb I M J (\<Pi>\<^sub>E j\<in>J. F j)) = (\<Prod>j\<in>J. emeasure (M j) (F j))"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   395
  shows "M' = (PiM I M)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   396
proof (rule measure_eqI_generator_eq[symmetric, OF Int_stable_prod_algebra prod_algebra_sets_into_space])
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   397
  show "sets (PiM I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   398
    by (rule sets_PiM)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   399
  then show "sets M' = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   400
    unfolding `sets M' = sets (PiM I M)` by simp
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   401
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   402
  def i \<equiv> "SOME i. i \<in> I"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   403
  with `I \<noteq> {}` have i: "i \<in> I"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   404
    by (auto intro: someI_ex)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   405
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   406
  def A \<equiv> "\<lambda>n::nat. prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. space (M i))"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   407
  then show "range A \<subseteq> prod_algebra I M"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   408
    by (auto intro!: prod_algebraI i)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   409
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   410
  have A_eq: "\<And>i. A i = space (PiM I M)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   411
    by (auto simp: prod_emb_def space_PiM Pi_iff A_def i)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   412
  show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   413
    unfolding A_eq by (auto simp: space_PiM)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   414
  show "\<And>i. emeasure (PiM I M) (A i) \<noteq> \<infinity>"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   415
    unfolding A_eq P.emeasure_space_1 by simp
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   416
next
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   417
  fix X assume X: "X \<in> prod_algebra I M"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   418
  then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   419
    and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   420
    by (force elim!: prod_algebraE)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   421
  from eq[OF J] have "emeasure M' X = (\<Prod>j\<in>J. emeasure (M j) (E j))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   422
    by (simp add: X)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   423
  also have "\<dots> = emeasure (PiM I M) X"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   424
    unfolding X using J by (intro emeasure_PiM_emb[symmetric]) auto
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   425
  finally show "emeasure (PiM I M) X = emeasure M' X" ..
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   426
qed
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   427
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   428
subsection {* Sequence space *}
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   429
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   430
definition comb_seq :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a)" where
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   431
  "comb_seq i \<omega> \<omega>' j = (if j < i then \<omega> j else \<omega>' (j - i))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   432
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   433
lemma split_comb_seq: "P (comb_seq i \<omega> \<omega>' j) \<longleftrightarrow> (j < i \<longrightarrow> P (\<omega> j)) \<and> (\<forall>k. j = i + k \<longrightarrow> P (\<omega>' k))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   434
  by (auto simp: comb_seq_def not_less)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   435
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   436
lemma split_comb_seq_asm: "P (comb_seq i \<omega> \<omega>' j) \<longleftrightarrow> \<not> ((j < i \<and> \<not> P (\<omega> j)) \<or> (\<exists>k. j = i + k \<and> \<not> P (\<omega>' k)))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   437
  by (auto simp: comb_seq_def)
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   438
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   439
lemma measurable_comb_seq:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   440
  "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> measurable ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) (\<Pi>\<^sub>M i\<in>UNIV. M)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   441
proof (rule measurable_PiM_single)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   442
  show "(\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)) \<rightarrow> (UNIV \<rightarrow>\<^sub>E space M)"
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50099
diff changeset
   443
    by (auto simp: space_pair_measure space_PiM PiE_iff split: split_comb_seq)
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   444
  fix j :: nat and A assume A: "A \<in> sets M"
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
   445
  then have *: "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). case_prod (comb_seq i) \<omega> j \<in> A} =
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   446
    (if j < i then {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> j \<in> A} \<times> space (\<Pi>\<^sub>M i\<in>UNIV. M)
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   447
              else space (\<Pi>\<^sub>M i\<in>UNIV. M) \<times> {\<omega> \<in> space (\<Pi>\<^sub>M i\<in>UNIV. M). \<omega> (j - i) \<in> A})"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   448
    by (auto simp: space_PiM space_pair_measure comb_seq_def dest: sets.sets_into_space)
55414
eab03e9cee8a renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents: 53374
diff changeset
   449
  show "{\<omega> \<in> space ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M)). case_prod (comb_seq i) \<omega> j \<in> A} \<in> sets ((\<Pi>\<^sub>M i\<in>UNIV. M) \<Otimes>\<^sub>M (\<Pi>\<^sub>M i\<in>UNIV. M))"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   450
    unfolding * by (auto simp: A intro!: sets_Collect_single)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   451
qed
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   452
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   453
lemma measurable_comb_seq'[measurable (raw)]:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   454
  assumes f: "f \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)" and g: "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   455
  shows "(\<lambda>x. comb_seq i (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   456
  using measurable_compose[OF measurable_Pair[OF f g] measurable_comb_seq] by simp
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   457
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   458
lemma comb_seq_0: "comb_seq 0 \<omega> \<omega>' = \<omega>'"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   459
  by (auto simp add: comb_seq_def)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   460
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   461
lemma comb_seq_Suc: "comb_seq (Suc n) \<omega> \<omega>' = comb_seq n \<omega> (case_nat (\<omega> n) \<omega>')"
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   462
  by (auto simp add: comb_seq_def not_less less_Suc_eq le_imp_diff_is_add intro!: ext split: nat.split)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   463
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   464
lemma comb_seq_Suc_0[simp]: "comb_seq (Suc 0) \<omega> = case_nat (\<omega> 0)"
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   465
  by (intro ext) (simp add: comb_seq_Suc comb_seq_0)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   466
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   467
lemma comb_seq_less: "i < n \<Longrightarrow> comb_seq n \<omega> \<omega>' i = \<omega> i"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   468
  by (auto split: split_comb_seq)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   469
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   470
lemma comb_seq_add: "comb_seq n \<omega> \<omega>' (i + n) = \<omega>' i"
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   471
  by (auto split: nat.split split_comb_seq)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   472
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   473
lemma case_nat_comb_seq: "case_nat s' (comb_seq n \<omega> \<omega>') (i + n) = case_nat (case_nat s' \<omega> n) \<omega>' i"
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   474
  by (auto split: nat.split split_comb_seq)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   475
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   476
lemma case_nat_comb_seq':
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   477
  "case_nat s (comb_seq i \<omega> \<omega>') = comb_seq (Suc i) (case_nat s \<omega>) \<omega>'"
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   478
  by (auto split: split_comb_seq nat.split)
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   479
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   480
locale sequence_space = product_prob_space "\<lambda>i. M" "UNIV :: nat set" for M
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   481
begin
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   482
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   483
abbreviation "S \<equiv> \<Pi>\<^sub>M i\<in>UNIV::nat set. M"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   484
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   485
lemma infprod_in_sets[intro]:
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   486
  fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   487
  shows "Pi UNIV E \<in> sets S"
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   488
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   489
  have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^sub>E j\<in>{..i}. E j))"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   490
    using E E[THEN sets.sets_into_space]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   491
    by (auto simp: prod_emb_def Pi_iff extensional_def) blast
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   492
  with E show ?thesis by auto
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   493
qed
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   494
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   495
lemma measure_PiM_countable:
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   496
  fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets M"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   497
  shows "(\<lambda>n. \<Prod>i\<le>n. measure M (E i)) ----> measure S (Pi UNIV E)"
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   498
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   499
  let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^sub>E {.. n} E)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   500
  have "\<And>n. (\<Prod>i\<le>n. measure M (E i)) = measure S (?E n)"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   501
    using E by (simp add: measure_PiM_emb)
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   502
  moreover have "Pi UNIV E = (\<Inter>n. ?E n)"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   503
    using E E[THEN sets.sets_into_space]
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   504
    by (auto simp: prod_emb_def extensional_def Pi_iff) blast
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   505
  moreover have "range ?E \<subseteq> sets S"
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   506
    using E by auto
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   507
  moreover have "decseq ?E"
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   508
    by (auto simp: prod_emb_def Pi_iff decseq_def)
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   509
  ultimately show ?thesis
47694
05663f75964c reworked Probability theory
hoelzl
parents: 46905
diff changeset
   510
    by (simp add: finite_Lim_measure_decseq)
42257
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   511
qed
08d717c82828 prove measurable_into_infprod_algebra and measure_infprod
hoelzl
parents: 42166
diff changeset
   512
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   513
lemma nat_eq_diff_eq: 
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   514
  fixes a b c :: nat
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   515
  shows "c \<le> b \<Longrightarrow> a = b - c \<longleftrightarrow> a + c = b"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   516
  by auto
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   517
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   518
lemma PiM_comb_seq:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   519
  "distr (S \<Otimes>\<^sub>M S) S (\<lambda>(\<omega>, \<omega>'). comb_seq i \<omega> \<omega>') = S" (is "?D = _")
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   520
proof (rule PiM_eq)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   521
  let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   522
  let "distr _ _ ?f" = "?D"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   523
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   524
  fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   525
  let ?X = "prod_emb ?I ?M J (\<Pi>\<^sub>E j\<in>J. E j)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   526
  have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   527
    using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   528
  with J have "?f -` ?X \<inter> space (S \<Otimes>\<^sub>M S) =
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   529
    (prod_emb ?I ?M (J \<inter> {..<i}) (PIE j:J \<inter> {..<i}. E j)) \<times>
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   530
    (prod_emb ?I ?M ((op + i) -` J) (PIE j:(op + i) -` J. E (i + j)))" (is "_ = ?E \<times> ?F")
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50099
diff changeset
   531
   by (auto simp: space_pair_measure space_PiM prod_emb_def all_conj_distrib PiE_iff
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   532
               split: split_comb_seq split_comb_seq_asm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   533
  then have "emeasure ?D ?X = emeasure (S \<Otimes>\<^sub>M S) (?E \<times> ?F)"
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   534
    by (subst emeasure_distr[OF measurable_comb_seq])
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   535
       (auto intro!: sets_PiM_I simp: split_beta' J)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   536
  also have "\<dots> = emeasure S ?E * emeasure S ?F"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   537
    using J by (intro P.emeasure_pair_measure_Times)  (auto intro!: sets_PiM_I finite_vimageI simp: inj_on_def)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   538
  also have "emeasure S ?F = (\<Prod>j\<in>(op + i) -` J. emeasure M (E (i + j)))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   539
    using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI inj_on_def)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   540
  also have "\<dots> = (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56996
diff changeset
   541
    by (rule setprod.reindex_cong [of "\<lambda>x. x - i"])
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   542
       (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   543
  also have "emeasure S ?E = (\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   544
    using J by (intro emeasure_PiM_emb) simp_all
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   545
  also have "(\<Prod>j\<in>J \<inter> {..<i}. emeasure M (E j)) * (\<Prod>j\<in>J - (J \<inter> {..<i}). emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   546
    by (subst mult.commute) (auto simp: J setprod.subset_diff[symmetric])
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   547
  finally show "emeasure ?D ?X = (\<Prod>j\<in>J. emeasure M (E j))" .
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   548
qed simp_all
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   549
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   550
lemma PiM_iter:
55415
05f5fdb8d093 renamed 'nat_{case,rec}' to '{case,rec}_nat'
blanchet
parents: 55414
diff changeset
   551
  "distr (M \<Otimes>\<^sub>M S) S (\<lambda>(s, \<omega>). case_nat s \<omega>) = S" (is "?D = _")
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   552
proof (rule PiM_eq)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   553
  let ?I = "UNIV::nat set" and ?M = "\<lambda>n. M"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   554
  let "distr _ _ ?f" = "?D"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   555
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   556
  fix J E assume J: "finite J" "J \<subseteq> ?I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets M"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   557
  let ?X = "prod_emb ?I ?M J (PIE j:J. E j)"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   558
  have "\<And>j x. j \<in> J \<Longrightarrow> x \<in> E j \<Longrightarrow> x \<in> space M"
50244
de72bbe42190 qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents: 50123
diff changeset
   559
    using J(3)[THEN sets.sets_into_space] by (auto simp: space_PiM Pi_iff subset_eq)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   560
  with J have "?f -` ?X \<inter> space (M \<Otimes>\<^sub>M S) = (if 0 \<in> J then E 0 else space M) \<times>
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   561
    (prod_emb ?I ?M (Suc -` J) (PIE j:Suc -` J. E (Suc j)))" (is "_ = ?E \<times> ?F")
50123
69b35a75caf3 merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents: 50099
diff changeset
   562
   by (auto simp: space_pair_measure space_PiM PiE_iff prod_emb_def all_conj_distrib
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   563
      split: nat.split nat.split_asm)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51351
diff changeset
   564
  then have "emeasure ?D ?X = emeasure (M \<Otimes>\<^sub>M S) (?E \<times> ?F)"
50099
a58bb401af80 measurability for nat_case and comb_seq
hoelzl
parents: 50095
diff changeset
   565
    by (subst emeasure_distr)
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   566
       (auto intro!: sets_PiM_I simp: split_beta' J)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   567
  also have "\<dots> = emeasure M ?E * emeasure S ?F"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   568
    using J by (intro P.emeasure_pair_measure_Times) (auto intro!: sets_PiM_I finite_vimageI)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   569
  also have "emeasure S ?F = (\<Prod>j\<in>Suc -` J. emeasure M (E (Suc j)))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   570
    using J by (intro emeasure_PiM_emb) (simp_all add: finite_vimageI)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   571
  also have "\<dots> = (\<Prod>j\<in>J - {0}. emeasure M (E j))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 56996
diff changeset
   572
    by (rule setprod.reindex_cong [of "\<lambda>x. x - 1"])
50000
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   573
       (auto simp: image_iff Bex_def not_less nat_eq_diff_eq ac_simps cong: conj_cong intro!: inj_onI)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   574
  also have "emeasure M ?E * (\<Prod>j\<in>J - {0}. emeasure M (E j)) = (\<Prod>j\<in>J. emeasure M (E j))"
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   575
    by (auto simp: M.emeasure_space_1 setprod.remove J)
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   576
  finally show "emeasure ?D ?X = (\<Prod>j\<in>J. emeasure M (E j))" .
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   577
qed simp_all
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   578
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   579
end
cfe8ee8a1371 infinite product measure is invariant under adding prefixes
hoelzl
parents: 49804
diff changeset
   580
42147
61d5d50ca74c add infinite product measure
hoelzl
parents:
diff changeset
   581
end