src/HOL/Probability/Projective_Limit.thy
author wenzelm
Wed, 30 Dec 2015 11:21:54 +0100
changeset 61973 0c7e865fa7cb
parent 61969 e01015e49041
child 62397 5ae24f33d343
permissions -rw-r--r--
more symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50091
b3b5dc2350b7 corrected headers
immler
parents: 50090
diff changeset
     1
(*  Title:      HOL/Probability/Projective_Limit.thy
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
     2
    Author:     Fabian Immler, TU München
32d1795cc77a added projective limit;
immler
parents:
diff changeset
     3
*)
32d1795cc77a added projective limit;
immler
parents:
diff changeset
     4
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
     5
section \<open>Projective Limit\<close>
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
     6
32d1795cc77a added projective limit;
immler
parents:
diff changeset
     7
theory Projective_Limit
32d1795cc77a added projective limit;
immler
parents:
diff changeset
     8
  imports
32d1795cc77a added projective limit;
immler
parents:
diff changeset
     9
    Caratheodory
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    10
    Fin_Map
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    11
    Regularity
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    12
    Projective_Family
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    13
    Infinite_Product_Measure
50971
5e3d3d690975 tune prove compact_eq_totally_bounded
hoelzl
parents: 50884
diff changeset
    14
    "~~/src/HOL/Library/Diagonal_Subsequence"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    15
begin
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    16
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
    17
subsection \<open>Sequences of Finite Maps in Compact Sets\<close>
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    18
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    19
locale finmap_seqs_into_compact =
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
    20
  fixes K::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a::metric_space) set" and f::"nat \<Rightarrow> (nat \<Rightarrow>\<^sub>F 'a)" and M
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    21
  assumes compact: "\<And>n. compact (K n)"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    22
  assumes f_in_K: "\<And>n. K n \<noteq> {}"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    23
  assumes domain_K: "\<And>n. k \<in> K n \<Longrightarrow> domain k = domain (f n)"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    24
  assumes proj_in_K:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
    25
    "\<And>t n m. m \<ge> n \<Longrightarrow> t \<in> domain (f n) \<Longrightarrow> (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    26
begin
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    27
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
    28
lemma proj_in_K': "(\<exists>n. \<forall>m \<ge> n. (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n)"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    29
  using proj_in_K f_in_K
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    30
proof cases
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    31
  obtain k where "k \<in> K (Suc 0)" using f_in_K by auto
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    32
  assume "\<forall>n. t \<notin> domain (f n)"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    33
  thus ?thesis
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
    34
    by (auto intro!: exI[where x=1] image_eqI[OF _ \<open>k \<in> K (Suc 0)\<close>]
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
    35
      simp: domain_K[OF \<open>k \<in> K (Suc 0)\<close>])
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    36
qed blast
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    37
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    38
lemma proj_in_KE:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
    39
  obtains n where "\<And>m. m \<ge> n \<Longrightarrow> (f m)\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K n"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    40
  using proj_in_K' by blast
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    41
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    42
lemma compact_projset:
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
    43
  shows "compact ((\<lambda>k. (k)\<^sub>F i) ` K n)"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    44
  using continuous_proj compact by (rule compact_continuous_image)
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    45
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    46
end
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    47
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    48
lemma compactE':
50884
2b21b4e2d7cb differentiate (cover) compactness and sequential compactness
hoelzl
parents: 50252
diff changeset
    49
  fixes S :: "'a :: metric_space set"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    50
  assumes "compact S" "\<forall>n\<ge>m. f n \<in> S"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
    51
  obtains l r where "l \<in> S" "subseq r" "((f \<circ> r) \<longlongrightarrow> l) sequentially"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    52
proof atomize_elim
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    53
  have "subseq (op + m)" by (simp add: subseq_def)
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    54
  have "\<forall>n. (f o (\<lambda>i. m + i)) n \<in> S" using assms by auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
    55
  from seq_compactE[OF \<open>compact S\<close>[unfolded compact_eq_seq_compact_metric] this] guess l r .
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    56
  hence "l \<in> S" "subseq ((\<lambda>i. m + i) o r) \<and> (f \<circ> ((\<lambda>i. m + i) o r)) \<longlonglongrightarrow> l"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
    57
    using subseq_o[OF \<open>subseq (op + m)\<close> \<open>subseq r\<close>] by (auto simp: o_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    58
  thus "\<exists>l r. l \<in> S \<and> subseq r \<and> (f \<circ> r) \<longlonglongrightarrow> l" by blast
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    59
qed
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    60
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    61
sublocale finmap_seqs_into_compact \<subseteq> subseqs "\<lambda>n s. (\<exists>l. (\<lambda>i. ((f o s) i)\<^sub>F n) \<longlonglongrightarrow> l)"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    62
proof
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    63
  fix n s
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    64
  assume "subseq s"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    65
  from proj_in_KE[of n] guess n0 . note n0 = this
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
    66
  have "\<forall>i \<ge> n0. ((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    67
  proof safe
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    68
    fix i assume "n0 \<le> i"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    69
    also have "\<dots> \<le> s i" by (rule seq_suble) fact
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    70
    finally have "n0 \<le> s i" .
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
    71
    with n0 show "((f \<circ> s) i)\<^sub>F n \<in> (\<lambda>k. (k)\<^sub>F n) ` K n0 "
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    72
      by auto
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    73
  qed
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    74
  from compactE'[OF compact_projset this] guess ls rs .
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    75
  thus "\<exists>r'. subseq r' \<and> (\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r')) i)\<^sub>F n) \<longlonglongrightarrow> l)" by (auto simp: o_def)
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    76
qed
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    77
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    78
lemma (in finmap_seqs_into_compact) diagonal_tendsto: "\<exists>l. (\<lambda>i. (f (diagseq i))\<^sub>F n) \<longlonglongrightarrow> l"
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    79
proof -
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    80
  obtain l where "(\<lambda>i. ((f o (diagseq o op + (Suc n))) i)\<^sub>F n) \<longlonglongrightarrow> l"
52681
8cc7f76b827a tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
immler
parents: 51351
diff changeset
    81
  proof (atomize_elim, rule diagseq_holds)
8cc7f76b827a tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
immler
parents: 51351
diff changeset
    82
    fix r s n
8cc7f76b827a tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
immler
parents: 51351
diff changeset
    83
    assume "subseq r"
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    84
    assume "\<exists>l. (\<lambda>i. ((f \<circ> s) i)\<^sub>F n) \<longlonglongrightarrow> l"
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    85
    then obtain l where "((\<lambda>i. (f i)\<^sub>F n) o s) \<longlonglongrightarrow> l"
52681
8cc7f76b827a tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
immler
parents: 51351
diff changeset
    86
      by (auto simp: o_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    87
    hence "((\<lambda>i. (f i)\<^sub>F n) o s o r) \<longlonglongrightarrow> l" using \<open>subseq r\<close>
52681
8cc7f76b827a tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
immler
parents: 51351
diff changeset
    88
      by (rule LIMSEQ_subseq_LIMSEQ)
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    89
    thus "\<exists>l. (\<lambda>i. ((f \<circ> (s \<circ> r)) i)\<^sub>F n) \<longlonglongrightarrow> l" by (auto simp add: o_def)
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    90
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    91
  hence "(\<lambda>i. ((f (diagseq (i + Suc n))))\<^sub>F n) \<longlonglongrightarrow> l" by (simp add: ac_simps)
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
    92
  hence "(\<lambda>i. (f (diagseq i))\<^sub>F n) \<longlonglongrightarrow> l" by (rule LIMSEQ_offset)
52681
8cc7f76b827a tuned definition of seqseq; clarified usage of diagseq via diagseq_holds
immler
parents: 51351
diff changeset
    93
  thus ?thesis ..
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    94
qed
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    95
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
    96
subsection \<open>Daniell-Kolmogorov Theorem\<close>
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    97
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
    98
text \<open>Existence of Projective Limit\<close>
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
    99
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   100
locale polish_projective = projective_family I P "\<lambda>_. borel::'a::polish_space measure"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   101
  for I::"'i set" and P
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   102
begin
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   103
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   104
lemma emeasure_lim_emb:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   105
  assumes X: "J \<subseteq> I" "finite J" "X \<in> sets (\<Pi>\<^sub>M i\<in>J. borel)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   106
  shows "lim (emb I J X) = P J X"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   107
proof (rule emeasure_lim)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   108
  write mu_G ("\<mu>G")
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   109
  interpret generator: algebra "space (PiM I (\<lambda>i. borel))" generator
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   110
    by (rule algebra_generator)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   111
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   112
  fix J and B :: "nat \<Rightarrow> ('i \<Rightarrow> 'a) set"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   113
  assume J: "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. B n \<in> sets (\<Pi>\<^sub>M i\<in>J n. borel)" "incseq J"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   114
    and B: "decseq (\<lambda>n. emb I (J n) (B n))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   115
    and "0 < (INF i. P (J i) (B i))" (is "0 < ?a")
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   116
  moreover have "?a \<le> 1"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   117
    using J by (auto intro!: INF_lower2[of 0] prob_space_P[THEN prob_space.measure_le_1])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   118
  ultimately obtain r where r: "?a = ereal r" "0 < r" "r \<le> 1"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   119
    by (cases "?a") auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   120
  def Z \<equiv> "\<lambda>n. emb I (J n) (B n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   121
  have Z_mono: "n \<le> m \<Longrightarrow> Z m \<subseteq> Z n" for n m
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   122
    unfolding Z_def using B[THEN antimonoD, of n m] .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   123
  have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   124
    using \<open>incseq J\<close> by (force simp: incseq_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   125
  note [simp] = \<open>\<And>n. finite (J n)\<close>
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   126
  interpret prob_space "P (J i)" for i using J prob_space_P by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   127
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   128
  have P_eq[simp]:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   129
      "sets (P (J i)) = sets (\<Pi>\<^sub>M i\<in>J i. borel)" "space (P (J i)) = space (\<Pi>\<^sub>M i\<in>J i. borel)" for i
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   130
    using J by (auto simp: sets_P space_P)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   131
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   132
  have "Z i \<in> generator" for i
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   133
    unfolding Z_def by (auto intro!: generator.intros J)
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   134
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   135
  have countable_UN_J: "countable (\<Union>n. J n)" by (simp add: countable_finite)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   136
  def Utn \<equiv> "to_nat_on (\<Union>n. J n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   137
  interpret function_to_finmap "J n" Utn "from_nat_into (\<Union>n. J n)" for n
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   138
    by unfold_locales (auto simp: Utn_def intro: from_nat_into_to_nat_on[OF countable_UN_J])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   139
  have inj_on_Utn: "inj_on Utn (\<Union>n. J n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   140
    unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   141
  hence inj_on_Utn_J: "\<And>n. inj_on Utn (J n)" by (rule subset_inj_on) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   142
  def P' \<equiv> "\<lambda>n. mapmeasure n (P (J n)) (\<lambda>_. borel)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   143
  interpret P': prob_space "P' n" for n
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   144
    unfolding P'_def mapmeasure_def using J
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   145
    by (auto intro!: prob_space_distr fm_measurable simp: measurable_cong_sets[OF sets_P])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   146
  
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   147
  let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   148
  { fix n
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   149
    have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   150
      using J by (auto simp: P'_def mapmeasure_PiM space_P sets_P)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   151
    also
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   152
    have "\<dots> = ?SUP n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   153
    proof (rule inner_regular)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   154
      show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   155
    next
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   156
      show "fm n ` B n \<in> sets borel"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   157
        unfolding borel_eq_PiF_borel by (auto simp: P'_def fm_image_measurable_finite sets_P J(3))
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   158
    qed simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   159
    finally have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n \<noteq> \<infinity>" "?SUP n \<noteq> - \<infinity>" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   160
  } note R = this
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   161
  have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a \<and> compact K \<and> K \<subseteq> fm n ` B n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   162
  proof
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   163
    fix n show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ereal (2 powr - real n) * ?a \<and>
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   164
        compact K \<and> K \<subseteq> fm n ` B n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   165
      unfolding R[of n]
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   166
    proof (rule ccontr)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   167
      assume H: "\<not> (\<exists>K'. ?SUP n - emeasure (P' n) K' \<le> ereal (2 powr - real n)  * ?a \<and>
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   168
        compact K' \<and> K' \<subseteq> fm n ` B n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   169
      have "?SUP n \<le> ?SUP n - 2 powr (-n) * ?a"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   170
      proof (intro SUP_least)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   171
        fix K
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   172
        assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   173
        with H have "\<not> ?SUP n - emeasure (P' n) K \<le> 2 powr (-n) * ?a"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   174
          by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   175
        hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   176
          unfolding not_less[symmetric] by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   177
        hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   178
          using \<open>0 < ?a\<close> by (auto simp add: ereal_less_minus_iff ac_simps)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   179
        thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   180
      qed
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   181
      hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using \<open>0 < ?a\<close> by simp
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   182
      hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   183
      hence "0 \<le> - (2 powr (-n) * ?a)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   184
        using \<open>?SUP n \<noteq> \<infinity>\<close> \<open>?SUP n \<noteq> - \<infinity>\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   185
        by (subst (asm) ereal_add_le_add_iff) (auto simp:)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   186
      moreover have "ereal (2 powr - real n) * ?a > 0" using \<open>0 < ?a\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   187
        by (auto simp: ereal_zero_less_0_iff)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   188
      ultimately show False by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   189
    qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   190
  qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   191
  then obtain K' where K':
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   192
    "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   193
    "\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   194
    unfolding choice_iff by blast
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   195
  def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   196
  have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   197
    unfolding K_def
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   198
    using compact_imp_closed[OF \<open>compact (K' _)\<close>]
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   199
    by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   200
       (auto simp: borel_eq_PiF_borel[symmetric])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   201
  have K_B: "\<And>n. K n \<subseteq> B n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   202
  proof
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   203
    fix x n assume "x \<in> K n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   204
    then have fm_in: "fm n x \<in> fm n ` B n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   205
      using K' by (force simp: K_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   206
    show "x \<in> B n"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   207
      using \<open>x \<in> K n\<close> K_sets sets.sets_into_space J(1,2,3)[of n] inj_on_image_mem_iff[OF inj_on_fm]
61520
8f85bb443d33 Cauchy's integral formula, required lemmas, and a bit of reorganisation
paulson <lp15@cam.ac.uk>
parents: 61359
diff changeset
   208
    by (metis (no_types) Int_iff K_def fm_in space_borel)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   209
  qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   210
  def Z' \<equiv> "\<lambda>n. emb I (J n) (K n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   211
  have Z': "\<And>n. Z' n \<subseteq> Z n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   212
    unfolding Z'_def Z_def
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   213
  proof (rule prod_emb_mono, safe)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   214
    fix n x assume "x \<in> K n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   215
    hence "fm n x \<in> K' n" "x \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   216
      by (simp_all add: K_def space_P)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   217
    note this(1)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   218
    also have "K' n \<subseteq> fm n ` B n" by (simp add: K')
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   219
    finally have "fm n x \<in> fm n ` B n" .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   220
    thus "x \<in> B n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   221
    proof safe
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   222
      fix y assume y: "y \<in> B n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   223
      hence "y \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   224
        by (auto simp add: space_P sets_P)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   225
      assume "fm n x = fm n y"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   226
      note inj_onD[OF inj_on_fm[OF space_borel],
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   227
        OF \<open>fm n x = fm n y\<close> \<open>x \<in> space _\<close> \<open>y \<in> space _\<close>]
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   228
      with y show "x \<in> B n" by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   229
    qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   230
  qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   231
  have "\<And>n. Z' n \<in> generator" using J K'(2) unfolding Z'_def
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   232
    by (auto intro!: generator.intros measurable_sets[OF fm_measurable[of _ "Collect finite"]]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   233
             simp: K_def borel_eq_PiF_borel[symmetric] compact_imp_closed)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   234
  def Y \<equiv> "\<lambda>n. \<Inter>i\<in>{1..n}. Z' i"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   235
  hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   236
  hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   237
  have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   238
  hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   239
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   240
  have Y_notempty: "\<And>n. n \<ge> 1 \<Longrightarrow> (Y n) \<noteq> {}"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   241
  proof -
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   242
    fix n::nat assume "n \<ge> 1" hence "Y n \<subseteq> Z n" by fact
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   243
    have "Y n = (\<Inter>i\<in>{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   244
      by (auto simp: Y_def Z'_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   245
    also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. emb (J n) (J i) (K i))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   246
      using \<open>n \<ge> 1\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   247
      by (subst prod_emb_INT) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   248
    finally
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   249
    have Y_emb:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   250
      "Y n = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   251
    hence "Y n \<in> generator" using J J_mono K_sets \<open>n \<ge> 1\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   252
      by (auto simp del: prod_emb_INT intro!: generator.intros)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   253
    have *: "\<mu>G (Z n) = P (J n) (B n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   254
      unfolding Z_def using J by (intro mu_G_spec) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   255
    then have "\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   256
    note *
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   257
    moreover have *: "\<mu>G (Y n) = P (J n) (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   258
      unfolding Y_emb using J J_mono K_sets \<open>n \<ge> 1\<close> by (subst mu_G_spec) auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   259
    then have "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   260
    note *
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   261
    moreover have "\<mu>G (Z n - Y n) =
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   262
        P (J n) (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   263
      unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets \<open>n \<ge> 1\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   264
      by (subst mu_G_spec) (auto intro!: sets.Diff)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   265
    ultimately
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   266
    have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   267
      using J J_mono K_sets \<open>n \<ge> 1\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   268
      by (simp only: emeasure_eq_measure Z_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   269
        (auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] set_mp[OF K_B]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   270
          simp: extensional_restrict emeasure_eq_measure prod_emb_iff sets_P space_P)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   271
    also have subs: "Z n - Y n \<subseteq> (\<Union>i\<in>{1..n}. (Z i - Z' i))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   272
      using \<open>n \<ge> 1\<close> unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   273
    have "Z n - Y n \<in> generator" "(\<Union>i\<in>{1..n}. (Z i - Z' i)) \<in> generator"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   274
      using \<open>Z' _ \<in> generator\<close> \<open>Z _ \<in> generator\<close> \<open>Y _ \<in> generator\<close> by auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   275
    hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union>i\<in>{1..n}. (Z i - Z' i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   276
      using subs generator.additive_increasing[OF positive_mu_G additive_mu_G]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   277
      unfolding increasing_def by auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   278
    also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using \<open>Z _ \<in> generator\<close> \<open>Z' _ \<in> generator\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   279
      by (intro generator.subadditive[OF positive_mu_G additive_mu_G]) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   280
    also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   281
    proof (rule setsum_mono)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   282
      fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   283
      have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   284
        unfolding Z'_def Z_def by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   285
      also have "\<dots> = P (J i) (B i - K i)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   286
        using J K_sets by (subst mu_G_spec) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   287
      also have "\<dots> = P (J i) (B i) - P (J i) (K i)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   288
        using K_sets J \<open>K _ \<subseteq> B _\<close> by (simp add: emeasure_Diff)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   289
      also have "\<dots> = P (J i) (B i) - P' i (K' i)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   290
        unfolding K_def P'_def
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   291
        by (auto simp: mapmeasure_PiF borel_eq_PiF_borel[symmetric]
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   292
          compact_imp_closed[OF \<open>compact (K' _)\<close>] space_PiM PiE_def)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   293
      also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   294
      finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   295
    qed
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   296
    also have "\<dots> = ereal ((\<Sum> i\<in>{1..n}. (2 powr -real_of_ereal i)) * real_of_ereal ?a)"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   297
      by (simp add: setsum_left_distrib r)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   298
    also have "\<dots> < ereal (1 * real_of_ereal ?a)" unfolding less_ereal.simps
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   299
    proof (rule mult_strict_right_mono)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   300
      have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
   301
        by (rule setsum.cong) (auto simp: powr_realpow powr_divide power_divide powr_minus_divide)  
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   302
      also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   303
      also have "setsum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   304
        setsum (op ^ (1 / 2)) ({..<Suc n}) - 1" by (auto simp: setsum_diff1)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   305
      also have "\<dots> < 1" by (subst geometric_sum) auto
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61520
diff changeset
   306
      finally show "(\<Sum>i = 1..n. 2 powr - real_of_ereal i) < 1" by simp
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   307
    qed (auto simp: r ereal_less_real_iff zero_ereal_def[symmetric])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   308
    also have "\<dots> = ?a" by (auto simp: r)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   309
    also have "\<dots> \<le> \<mu>G (Z n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   310
      using J by (auto intro: INF_lower simp: Z_def mu_G_spec)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   311
    finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   312
    hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   313
      using \<open>\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>\<close> by (simp add: ereal_minus_less)
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   314
    have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   315
    also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   316
      apply (rule ereal_less_add[OF _ R]) using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   317
    finally have "\<mu>G (Y n) > 0"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   318
      using \<open>\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>\<close> by (auto simp: ac_simps zero_ereal_def[symmetric])
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   319
    thus "Y n \<noteq> {}" using positive_mu_G by (auto simp add: positive_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   320
  qed
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   321
  hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   322
  then obtain y where y: "\<And>n. n \<ge> 1 \<Longrightarrow> y n \<in> Y n" unfolding bchoice_iff by force
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   323
  {
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   324
    fix t and n m::nat
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   325
    assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   326
    from Y_mono[OF \<open>m \<ge> n\<close>] y[OF \<open>1 \<le> m\<close>] have "y m \<in> Y n" by auto
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   327
    also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF \<open>1 \<le> n\<close>] .
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   328
    finally
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   329
    have "fm n (restrict (y m) (J n)) \<in> K' n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   330
      unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   331
    moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   332
      using J by (simp add: fm_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   333
    ultimately have "fm n (y m) \<in> K' n" by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   334
  } note fm_in_K' = this
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   335
  interpret finmap_seqs_into_compact "\<lambda>n. K' (Suc n)" "\<lambda>k. fm (Suc k) (y (Suc k))" borel
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   336
  proof
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   337
    fix n show "compact (K' n)" by fact
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   338
  next
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   339
    fix n
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   340
    from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \<in> Y (Suc n)" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   341
    also have "\<dots> \<subseteq> Z' (Suc n)" using Y_Z' by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   342
    finally
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   343
    have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \<in> K' (Suc n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   344
      unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   345
    thus "K' (Suc n) \<noteq> {}" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   346
    fix k
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   347
    assume "k \<in> K' (Suc n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   348
    with K'[of "Suc n"] sets.sets_into_space have "k \<in> fm (Suc n) ` B (Suc n)" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   349
    then obtain b where "k = fm (Suc n) b" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   350
    thus "domain k = domain (fm (Suc n) (y (Suc n)))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   351
      by (simp_all add: fm_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   352
  next
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   353
    fix t and n m::nat
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   354
    assume "n \<le> m" hence "Suc n \<le> Suc m" by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   355
    assume "t \<in> domain (fm (Suc n) (y (Suc n)))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   356
    then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   357
    hence "j \<in> J (Suc m)" using J_mono[OF \<open>Suc n \<le> Suc m\<close>] by auto
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   358
    have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using \<open>n \<le> m\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   359
      by (intro fm_in_K') simp_all
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   360
    show "(fm (Suc m) (y (Suc m)))\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K' (Suc n)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   361
      apply (rule image_eqI[OF _ img])
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   362
      using \<open>j \<in> J (Suc n)\<close> \<open>j \<in> J (Suc m)\<close>
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   363
      unfolding j by (subst proj_fm, auto)+
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   364
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
   365
  have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> z"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   366
    using diagonal_tendsto ..
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   367
  then obtain z where z:
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
   368
    "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> z t"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   369
    unfolding choice_iff by blast
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   370
  {
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   371
    fix n :: nat assume "n \<ge> 1"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   372
    have "\<And>i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   373
      by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   374
    moreover
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   375
    {
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   376
      fix t
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   377
      assume t: "t \<in> domain (finmap_of (Utn ` J n) z)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   378
      hence "t \<in> Utn ` J n" by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   379
      then obtain j where j: "t = Utn j" "j \<in> J n" by auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
   380
      have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> z t"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   381
        apply (subst (2) tendsto_iff, subst eventually_sequentially)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   382
      proof safe
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   383
        fix e :: real assume "0 < e"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   384
        { fix i and x :: "'i \<Rightarrow> 'a" assume i: "i \<ge> n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   385
          assume "t \<in> domain (fm n x)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   386
          hence "t \<in> domain (fm i x)" using J_mono[OF \<open>i \<ge> n\<close>] by auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   387
          with i have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   388
            using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   389
        } note index_shift = this
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   390
        have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   391
          apply (rule le_SucI)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   392
          apply (rule order_trans) apply simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   393
          apply (rule seq_suble[OF subseq_diagseq])
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   394
          done
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   395
        from z
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   396
        have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   397
          unfolding tendsto_iff eventually_sequentially using \<open>0 < e\<close> by auto
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   398
        then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   399
          dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   400
        show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e "
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   401
        proof (rule exI[where x="max N n"], safe)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   402
          fix na assume "max N n \<le> na"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   403
          hence  "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) =
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   404
                  dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   405
            by (subst index_shift[OF I]) auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   406
          also have "\<dots> < e" using \<open>max N n \<le> na\<close> by (intro N) simp
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   407
          finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" .
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   408
        qed
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   409
      qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
   410
      hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) \<longlonglongrightarrow> (finmap_of (Utn ` J n) z)\<^sub>F t"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   411
        by (simp add: tendsto_intros)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   412
    } ultimately
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
   413
    have "(\<lambda>i. fm n (y (Suc (diagseq i)))) \<longlonglongrightarrow> finmap_of (Utn ` J n) z"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   414
      by (rule tendsto_finmap)
61969
e01015e49041 more symbols;
wenzelm
parents: 61808
diff changeset
   415
    hence "((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) \<longlonglongrightarrow> finmap_of (Utn ` J n) z"
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   416
      by (intro lim_subseq) (simp add: subseq_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   417
    moreover
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   418
    have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   419
      apply (auto simp add: o_def intro!: fm_in_K' \<open>1 \<le> n\<close> le_SucI)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   420
      apply (rule le_trans)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   421
      apply (rule le_add2)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   422
      using seq_suble[OF subseq_diagseq]
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   423
      apply auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   424
      done
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   425
    moreover
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61694
diff changeset
   426
    from \<open>compact (K' n)\<close> have "closed (K' n)" by (rule compact_imp_closed)
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   427
    ultimately
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   428
    have "finmap_of (Utn ` J n) z \<in> K' n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   429
      unfolding closed_sequential_limits by blast
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   430
    also have "finmap_of (Utn ` J n) z  = fm n (\<lambda>i. z (Utn i))"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   431
      unfolding finmap_eq_iff
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   432
    proof clarsimp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   433
      fix i assume i: "i \<in> J n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   434
      hence "from_nat_into (\<Union>n. J n) (Utn i) = i"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   435
        unfolding Utn_def
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   436
        by (subst from_nat_into_to_nat_on[OF countable_UN_J]) auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   437
      with i show "z (Utn i) = (fm n (\<lambda>i. z (Utn i)))\<^sub>F (Utn i)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   438
        by (simp add: finmap_eq_iff fm_def compose_def)
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   439
    qed
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   440
    finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   441
    moreover
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   442
    let ?J = "\<Union>n. J n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   443
    have "(?J \<inter> J n) = J n" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   444
    ultimately have "restrict (\<lambda>i. z (Utn i)) (?J \<inter> J n) \<in> K n"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   445
      unfolding K_def by (auto simp: space_P space_PiM)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   446
    hence "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z' n" unfolding Z'_def
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   447
      using J by (auto simp: prod_emb_def PiE_def extensional_def)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   448
    also have "\<dots> \<subseteq> Z n" using Z' by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   449
    finally have "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z n" .
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   450
  } note in_Z = this
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   451
  hence "(\<Inter>i\<in>{1..}. Z i) \<noteq> {}" by auto
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   452
  thus "(\<Inter>i. Z i) \<noteq> {}"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   453
    using INT_decseq_offset[OF antimonoI[OF Z_mono]] by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   454
qed fact+
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   455
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   456
lemma measure_lim_emb:
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   457
  "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> X \<in> sets (\<Pi>\<^sub>M i\<in>J. borel) \<Longrightarrow> measure lim (emb I J X) = measure (P J) X"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   458
  unfolding measure_def by (subst emeasure_lim_emb) auto
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   459
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   460
end
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   461
50090
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   462
hide_const (open) PiF
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 52681
diff changeset
   463
hide_const (open) Pi\<^sub>F
50090
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   464
hide_const (open) Pi'
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   465
hide_const (open) Abs_finmap
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   466
hide_const (open) Rep_finmap
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   467
hide_const (open) finmap_of
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   468
hide_const (open) proj
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   469
hide_const (open) domain
50245
dea9363887a6 based countable topological basis on Countable_Set
immler
parents: 50244
diff changeset
   470
hide_const (open) basis_finmap
50090
01203193dfa0 hide constants of auxiliary type finmap
immler
parents: 50088
diff changeset
   471
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61520
diff changeset
   472
sublocale polish_projective \<subseteq> P: prob_space lim
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   473
proof
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   474
  have *: "emb I {} {\<lambda>x. undefined} = space (\<Pi>\<^sub>M i\<in>I. borel)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   475
    by (auto simp: prod_emb_def space_PiM)
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   476
  interpret prob_space "P {}" 
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   477
    using prob_space_P by simp
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   478
  show "emeasure lim (space lim) = 1"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   479
    using emeasure_lim_emb[of "{}" "{\<lambda>x. undefined}"] emeasure_space_1
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   480
    by (simp add: * PiM_empty space_P)
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   481
qed
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   482
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   483
locale polish_product_prob_space =
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   484
  product_prob_space "\<lambda>_. borel::('a::polish_space) measure" I for I::"'i set"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   485
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   486
sublocale polish_product_prob_space \<subseteq> P: polish_projective I "\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)"
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   487
proof qed
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   488
61359
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   489
lemma (in polish_product_prob_space) limP_eq_PiM: "lim = PiM I (\<lambda>_. borel)"
e985b52c3eb3 cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
hoelzl
parents: 60809
diff changeset
   490
  by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_lim_emb)
50088
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   491
32d1795cc77a added projective limit;
immler
parents:
diff changeset
   492
end