Theory Projective_Limit

Up to index of Isabelle/HOL/HOL-Multivariate_Analysis/HOL-Probability/Stochastic_Processes

theory Projective_Limit
imports Fin_Map
(* Author: Fabian Immler <immler@in.tum.de> *)

theory Projective_Limit
imports Probability Polish_Space Fin_Map
begin

section {* Projective Limit *}

text {* Formalization of the Daniell-Kolmogorov theorem. *}

subsection {* (Finite) Product of Measures *}

text {* TODO: unifiy with @{term PiM} *}

definition
"PiP I M P = extend_measure
E i∈I. space (M i))
{x. (domain x ≠ {} ∨ I = {}) ∧
finite (domain x) ∧ domain x ⊆ I ∧ (x)F ∈ (ΠE i∈(domain x). sets (M i))}
(λx. prod_emb I M (domain x) (PiE (domain x) (x)F))
(λx. emeasure (P (domain x)) (PiE (domain x) (x)F))"


definition proj_algebra where
"proj_algebra I M = (λx. prod_emb I M (domain x) (PiE (domain x) (x)F)) `
{x. (domain x ≠ {} ∨ I = {}) ∧
finite (domain x) ∧ domain x ⊆ I ∧ (x)F ∈ (ΠE i∈domain x. sets (M i))}"


lemma proj_algebra_eq_prod_algebra:
"proj_algebra I M = prod_algebra I M"
proof safe
case goal1 then obtain X where "x = prod_emb I M (domain X) (PiE (domain X) (X)F)"
"domain X ≠ {} ∨ I = {}" "finite (domain X)" "domain X ⊆ I"
"(X)F ∈ (ΠE i∈domain X. sets (M i))"
by (auto simp: proj_algebra_def)
thus ?case by (auto simp: prod_algebra_def intro!: image_eqI[where x="(domain X, (X)F)"])
next
case goal2 then obtain J X where "x = prod_emb I M J (PiE J X)"
"J ≠ {} ∨ I = {}" "finite J" "J ⊆ I" "X ∈ (Π j∈J. sets (M j))"
by (auto simp: prod_algebra_def)
thus ?case by (auto simp: Pi_def proj_algebra_def intro!: image_eqI[where x="finmap_of J X"])
qed

lemma
shows proj_algebra_eq:
"proj_algebra I M = {prod_emb I M J (PiE J F) |J F.
(J ≠ {} ∨ I = {}) ∧ finite J ∧ J ⊆ I ∧ (∀i ∈ J. F i ∈ sets (M i))}"

unfolding proj_algebra_def
proof (rule, blast, rule)
case goal1
then obtain J F where "x = prod_emb I M J (PiE J F)"
"J ≠ {} ∨ I = {}" "finite J" "J ⊆ I" "!!i. i∈J ==> F i ∈ sets (M i)"
by auto
thus ?case by (auto intro!: image_eqI[where x="finmap_of J F"] simp: Pi_def)
qed

lemma proj_algebra_eq':
assumes "I ≠ {}"
shows "proj_algebra I M =
{prod_emb I M J (PiE J F) |J F. J ≠ {} ∧ finite J ∧ J ⊆ I ∧ (∀i ∈ J. F i ∈ sets (M i))}"

unfolding proj_algebra_eq
proof (intro antisym subsetI)
case goal1
then obtain J F where JF: "x = prod_emb I M J (PiE J F)"
"J ≠ {} ∨ I = {}" "finite J" "J ⊆ I" "!!i. i∈J ==> F i ∈ sets (M i)"
by auto
show ?case using assms JF by (auto intro!: exI[where x=J] exI[where x=F])
qed auto

lemma space_PiP[simp]: "space (PiP I M P) = space (PiM I M)"
by (auto simp: PiP_def space_PiM prod_emb_def intro!: space_extend_measure)

lemma sets_PiP': "sets (PiP I M P) = sigma_sets (ΠE i∈I. space (M i)) (proj_algebra I M)"
using prod_algebra_sets_into_space[of I M, simplified proj_algebra_eq_prod_algebra[symmetric]]
unfolding PiP_def proj_algebra_def
by (intro sets_extend_measure) simp

lemma sets_PiP[simp]: "sets (PiP I M P) = sets (PiM I M)"
unfolding sets_PiP' sets_PiM proj_algebra_eq_prod_algebra ..

lemma measurable_PiP1[simp]: "measurable (PiP I M P) M' = measurable (ΠM i∈I. M i) M'"
unfolding measurable_def by auto

lemma measurable_PiP2[simp]: "measurable M' (PiP I M P) = measurable M' (ΠM i∈I. M i)"
unfolding measurable_def by auto

subsection {* Projective Family *}

locale projective_family =
fixes I::"'i set" and P::"'i set => ('i => 'a) measure" and M::"('i => 'a measure)"
assumes projective: "!!J H. J ⊆ H ==> H ⊆ I ==> finite H ==>
(P H) (prod_emb H M J X) = (P J) X"

assumes prob_space: "!!J. prob_space (P J)"
assumes proj_sets: "!!J. sets (P J) = sets (PiM J M)"
assumes proj_space: "!!J. space (P J) = space (PiM J M)"
assumes measure_space: "!!i. prob_space (M i)"
--{* TODO: generalize definitions from @{text "product_prob_space"} to
@{text "product_measure_space"} *}

begin

lemma measurable_P1[simp]: "measurable (P J) M' = measurable (ΠM i∈J. M i) M'"
unfolding measurable_def proj_sets proj_space by auto

lemma measurable_P2[simp]: "measurable M' (P J) = measurable M' (ΠM i∈J. M i)"
unfolding measurable_def proj_sets proj_space by auto

end

sublocale projective_family M: prob_space "M i" for i using measure_space .

sublocale projective_family prob_space: prob_space "P J" for J using prob_space .

sublocale projective_family MP: product_prob_space M ..

context projective_family begin

lemma emeasure_PiP:
assumes "finite J"
assumes "J ⊆ I"
assumes A: "!!i. i∈J ==> A i ∈ sets (M i)"
shows "emeasure (PiP J M P) (PiE J A) = emeasure (P J) (PiE J A)"
proof -
def f "finmap_of J A"
def μ' "P J"
have "PiE J (restrict A J) ⊆ (ΠE i∈J. space (M i))"
proof safe
fix x j assume "x ∈ Pi J (restrict A J)" "j ∈ J"
hence "x j ∈ restrict A J j" by (auto simp: Pi_def)
also have "… ⊆ space (M j)" using sets_into_space A `j ∈ J` by auto
finally show "x j ∈ space (M j)" .
qed
hence "emeasure (PiP J M P) (PiE J A) =
emeasure (PiP J M P) (emb J (domain f) (PiE (domain f) f))"

using assms(1-3) sets_into_space by (auto simp add: f_def prod_emb_id Pi_def)
also have "… = emeasure (P J) (PiE J A)"
proof (subst emeasure_extend_measure[OF PiP_def, of _ _ μ'])
show "positive (sets (PiP J M P)) μ'" unfolding μ'_def positive_def by auto
show "countably_additive (sets (PiP J M P)) μ'" unfolding μ'_def countably_additive_def
by (auto simp: suminf_emeasure proj_sets)
show "emeasure (P (domain f)) (PiE (domain f) f) = emeasure (P J) (PiE J A)"
using assms by (simp add: f_def Pi_def)
show "f ∈ {x. (domain x ≠ {} ∨ J = {}) ∧ finite (domain x) ∧ domain x ⊆ J ∧
(x)F ∈ (ΠE i∈domain x. sets (M i))}"

using assms by (auto simp: f_def)
show "(λx. emb J (domain x) (PiE (domain x) (x)F)) ` {x. (domain x ≠ {} ∨ J = {}) ∧
finite (domain x) ∧ domain x ⊆ J ∧ (x)F ∈ (PiE (domain x) M)} ⊆
Pow (ΠE i∈J. space (M i)) "
by (auto simp: prod_emb_def)
fix i::"'i =>F 'a set"
assume "i ∈ {x. (domain x ≠ {} ∨ J = {}) ∧ finite (domain x) ∧ domain x ⊆ J ∧
(x)F ∈ (ΠE i∈(domain x). sets (M i))}"

with assms have
"finite (domain i)" "domain i ⊆ J" "(i)F ∈ (Π i∈domain i. sets (M i))"
by auto
thus "μ' (emb J (domain i) (PiE (domain i) (i)F)) =
emeasure (P (domain i)) (PiE (domain i) (i)F)"

using assms by (auto simp: projective μ'_def)
qed
finally show ?thesis .
qed

lemma PiP_finite:
assumes "finite J"
assumes "J ⊆ I"
shows "PiP J M P = P J" (is "?P = _")
proof (rule measure_eqI_generator_eq)
interpret J: finite_product_prob_space M J proof qed fact
let ?J = "{PiE J E | E. ∀i∈J. E i ∈ sets (M i)}"
let ?F = "λi. ΠE k∈J. space (M k)"
let = "(ΠE k∈J. space (M k))"
show "Int_stable ?J"
by (rule Int_stable_PiE)
show "emeasure ?P (?F _) ≠ ∞" using assms `finite J` by (auto simp: emeasure_PiP)
show "?J ⊆ Pow ?Ω" by (auto simp: Pi_iff dest: sets_into_space)
show "sets (PiP J M P) = sigma_sets ?Ω ?J" "sets (P J) = sigma_sets ?Ω ?J"
using `finite J` proj_sets by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
fix X assume "X ∈ ?J"
then obtain E where [simp]: "X = PiE J E" and E: "∀i∈J. E i ∈ sets (M i)" by auto
with `finite J` have X: "X ∈ sets (PiP J M P)" by auto
show "emeasure (PiP J M P) X = emeasure (P J) X" using assms `finite J` E
by (auto simp: emeasure_PiP)
qed (insert `finite J`, auto intro!: prod_algebraI_finite)

lemma emeasure_fun_emb[simp]:
assumes L: "J ⊆ L" "finite L" "L ⊆ I" and X: "X ∈ sets (PiP J M P)"
shows "emeasure (PiP L M P) (emb L J X) = emeasure (PiP J M P) X"
using assms
by (subst PiP_finite) (auto simp: PiP_finite finite_subset projective)

lemma distr_restrict:
assumes "J ⊆ K" "finite K" "K ⊆ I"
shows "(PiP J M P) = distr (PiP K M P) (PiP J M P) (λf. restrict f J)" (is "?P = ?D")
proof (rule measure_eqI)
show "sets (PiP J M P) = sets (distr (PiP K M P) (PiP J M P) (λf. restrict f J))" by simp
fix A assume "A ∈ sets (PiP J M P)"
with assms show "emeasure (PiP J M P) A =
emeasure (distr (PiP K M P) (PiP J M P) (λf. restrict f J)) A"

by (auto simp: emeasure_distr measurable_restrict_subset space_PiM prod_emb_def[symmetric])
qed

subsection {* Content on Generator *}

definition
"μG' A =
(THE x. ∀J. J ≠ {} --> finite J --> J ⊆ I -->
(∀X∈sets (PiP J M P). A = emb I J X --> x = emeasure (PiP J M P) X))"


lemma μG'_spec:
assumes J: "J ≠ {}" "finite J" "J ⊆ I" "A = emb I J X" "X ∈ sets (PiP J M P)"
shows "μG' A = emeasure (PiP J M P) X"
unfolding μG'_def
proof (intro the_equality allI impI ballI)
fix K Y assume K: "K ≠ {}" "finite K" "K ⊆ I" "A = emb I K Y" "Y ∈ sets (PiP K M P)"
have "emeasure (PiP K M P) Y = emeasure (PiP (K ∪ J) M P) (emb (K ∪ J) K Y)"
using K J by simp
also have "emb (K ∪ J) K Y = emb (K ∪ J) J X"
using K J by (simp add: prod_emb_injective[of "K ∪ J" I])
also have "emeasure (PiP (K ∪ J) M P) (emb (K ∪ J) J X) = emeasure (PiP J M P) X"
using K J by simp
finally show "emeasure (PiP J M P) X = emeasure (PiP K M P) Y" ..
qed (insert J, force)

lemma μG'_eq:
"J ≠ {} ==> finite J ==> J ⊆ I ==> X ∈ sets (PiP J M P) ==>
μG' (emb I J X) = emeasure (PiP J M P) X"

by (intro μG'_spec) auto

lemma generator_Ex':
assumes *: "A ∈ generator"
shows "∃J X. J ≠ {} ∧ finite J ∧ J ⊆ I ∧ X ∈ sets (ΠM i∈J. M i) ∧ A = emb I J X ∧
μG' A = emeasure (PiP J M P) X"

proof -
from * obtain J X where J: "J ≠ {}" "finite J" "J ⊆ I" "A = emb I J X" "X ∈ sets (PiP J M P)"
unfolding generator_def by auto
with μG'_spec[OF this] show ?thesis by auto
qed

lemma generatorE':
assumes A: "A ∈ generator"
obtains J X where "J ≠ {}" "finite J" "J ⊆ I" "X ∈ sets (PiP J M P)" "emb I J X = A"
"μG' A = emeasure (PiP J M P) X"
proof -
from generator_Ex'[OF A] obtain X J where "J ≠ {}" "finite J" "J ⊆ I" "X ∈ sets (PiP J M P)"
"emb I J X = A" "μG' A = emeasure (PiP J M P) X"
by auto
then show thesis by (intro that) auto
qed

lemma positive_μG':
assumes "I ≠ {}"
shows "positive generator μG'"
proof -
interpret G!: algebra E i∈I. space (M i)" generator by (rule algebra_generator) fact
show ?thesis
proof (intro positive_def[THEN iffD2] conjI ballI)
from generatorE'[OF G.empty_sets] guess J X . note this[simplified, simp]
interpret J: finite_product_sigma_finite M J by default fact
have "X = {}"
by (rule prod_emb_injective[of J I]) simp_all
then show "μG' {} = 0" by simp
next
fix A assume "A ∈ generator"
from generatorE'[OF this] guess J X . note this[simp]
interpret J: finite_product_sigma_finite M J by default fact
show "0 ≤ μG' A" by (simp add: emeasure_nonneg)
qed
qed

lemma additive_μG':
assumes "I ≠ {}"
shows "additive generator μG'"
proof -
interpret G!: algebra E i∈I. space (M i)" generator by (rule algebra_generator) fact
show ?thesis
proof (intro additive_def[THEN iffD2] ballI impI)
fix A assume "A ∈ generator" with generatorE' guess J X . note J = this
fix B assume "B ∈ generator" with generatorE' guess K Y . note K = this
assume "A ∩ B = {}"
have JK: "J ∪ K ≠ {}" "J ∪ K ⊆ I" "finite (J ∪ K)"
using J K by auto
interpret JK: finite_product_sigma_finite M "J ∪ K" by default fact
have JK_disj: "emb (J ∪ K) J X ∩ emb (J ∪ K) K Y = {}"
apply (rule prod_emb_injective[of "J ∪ K" I])
apply (insert `A ∩ B = {}` JK J K)
apply (simp_all add: Int prod_emb_Int)
done
have AB: "A = emb I (J ∪ K) (emb (J ∪ K) J X)" "B = emb I (J ∪ K) (emb (J ∪ K) K Y)"
using J K by simp_all
then have "μG' (A ∪ B) = μG' (emb I (J ∪ K) (emb (J ∪ K) J X ∪ emb (J ∪ K) K Y))"
by simp
also have "… = emeasure (PiP (J ∪ K) M P) (emb (J ∪ K) J X ∪ emb (J ∪ K) K Y)"
using JK J(1, 4) K(1, 4) by (simp add: μG'_eq Un del: prod_emb_Un)
also have "… = μG' A + μG' B"
using J K JK_disj by (simp add: plus_emeasure[symmetric])
finally show "μG' (A ∪ B) = μG' A + μG' B" .
qed
qed

end

subsection {* Sequences of Finite Maps in Compact Sets *}

locale finmap_seqs_into_compact =
fixes K::"nat => (nat =>F 'a::metric_space) set" and f::"nat => (nat =>F 'a)" and M
assumes compact: "!!n. compact (K n)"
assumes f_in_K: "!!n. K n ≠ {}"
assumes domain_K: "!!n. k ∈ K n ==> domain k = domain (f n)"
assumes proj_in_K:
"!!t n m. m ≥ n ==> t ∈ domain (f n) ==> (f m)F t ∈ (λk. (k)F t) ` K n"
begin

lemma proj_in_K': "(∃n. ∀m ≥ n. (f m)F t ∈ (λk. (k)F t) ` K n)"
using proj_in_K f_in_K
proof cases
obtain k where "k ∈ K (Suc 0)" using f_in_K by auto
assume "∀n. t ∉ domain (f n)"
thus ?thesis
by (auto intro!: exI[where x=1] image_eqI[OF _ `k ∈ K (Suc 0)`]
simp: domain_K[OF `k ∈ K (Suc 0)`])
qed blast

lemma proj_in_KE:
obtains n where "!!m. m ≥ n ==> (f m)F t ∈ (λk. (k)F t) ` K n"
using proj_in_K' by blast

lemma compact_projset:
shows "compact ((λk. (k)F i) ` K n)"
using continuous_proj compact by (rule compact_continuous_image)

end

sublocale finmap_seqs_into_compact subseqs "λn s r. (∃l. (λi. ((f o s o r) i)F n) ----> l)"
proof
fix n s
assume "subseq s"
from proj_in_KE[of n] guess n0 . note n0 = this
have "∀i ≥ n0. ((f o s) i)F n ∈ (λk. (k)F n) ` K n0"
proof safe
fix i assume "n0 ≤ i"
also have "… ≤ s i" by (rule seq_suble) fact
finally have "n0 ≤ s i" .
with n0 show "((f o s) i)F n ∈ (λk. (k)F n) ` K n0 "
by auto
qed
from compactE'[OF compact_projset this] guess ls rs .
thus "∃r'. subseq r' ∧ (∃l. (λi. ((f o s o r') i)F n) ----> l)" by (auto simp: o_def)
qed

lemma (in finmap_seqs_into_compact)
diagonal_tendsto: "∃l. (λi. (f (diagseq i))F n) ----> l"
proof -
have "!!i n0. (f o seqseq i) i = f (diagseq i)" unfolding diagseq_def by simp
from reducer_reduces obtain l where l: "(λi. ((f o seqseq (Suc n)) i)F n) ----> l"
unfolding seqseq_reducer
by auto
have "(λi. (f (diagseq (i + Suc n)))F n) =
(λi. ((f o (diagseq o (op + (Suc n)))) i)F n)"
by (simp add: add_commute)
also have "… =
(λi. ((f o ((seqseq (Suc n) o (λx. fold_reduce (Suc n) x (Suc n + x))))) i)F n)"

unfolding diagseq_seqseq by simp
also have "… = (λi. ((f o ((seqseq (Suc n)))) i)F n) o (λx. fold_reduce (Suc n) x (Suc n + x))"
by (simp add: o_def)
also have "… ----> l"
proof (rule LIMSEQ_subseq_LIMSEQ[OF _ subseq_diagonal_rest], rule tendstoI)
fix e::real assume "0 < e"
from tendstoD[OF l `0 < e`]
show "eventually (λx. dist (((f o seqseq (Suc n)) x)F n) l < e)
sequentially"
.
qed
finally show ?thesis by (intro exI) (rule LIMSEQ_offset)
qed

subsection {* The Daniell-Kolmogorov theorem *}

locale polish_projective = projective_family I P "λ_. borel::'a::polish_space measure"
for I::"'i set" and P
begin

abbreviation "PiB ≡ (λJ P. PiP J (λ_. borel) P)"

lemma
emeasure_PiB_emb_not_empty:
assumes "I ≠ {}"
assumes X: "J ≠ {}" "J ⊆ I" "finite J" "∀i∈J. B i ∈ sets borel"
shows "emeasure (PiB I P) (emb I J (PiE J B)) = emeasure (PiB J P) (PiE J B)"
proof -
let = E i∈I. space borel"
let ?G = generator
interpret G!: algebra generator by (intro algebra_generator) fact
note μG'_mono =
G.additive_increasing[OF positive_μG'[OF `I ≠ {}`] additive_μG'[OF `I ≠ {}`], THEN increasingD]
have "∃μ. (∀s∈?G. μ s = μG' s) ∧ measure_space ?Ω (sigma_sets ?Ω ?G) μ"
proof (rule G.caratheodory_empty_continuous[OF positive_μG' additive_μG',
OF `I ≠ {}`, OF `I ≠ {}`])
fix A assume "A ∈ ?G"
with generatorE' guess J X .
thus "μG' A ≠ ∞" by (simp add: PiP_finite)
next
fix Z assume Z: "range Z ⊆ ?G" "decseq Z" "(\<Inter>i. Z i) = {}"
then have "decseq (λi. μG' (Z i))"
by (auto intro!: μG'_mono simp: decseq_def)
moreover
have "(INF i. μG' (Z i)) = 0"
proof (rule ccontr)
assume "(INF i. μG' (Z i)) ≠ 0" (is "?a ≠ 0")
moreover have "0 ≤ ?a"
using Z positive_μG'[OF `I ≠ {}`] by (auto intro!: INF_greatest simp: positive_def)
ultimately have "0 < ?a" by auto
hence "?a ≠ -∞" by auto
have "∀n. ∃J B. J ≠ {} ∧ finite J ∧ J ⊆ I ∧ B ∈ sets (PiM J (λ_. borel)) ∧
Z n = emb I J B ∧ μG' (Z n) = emeasure (PiB J P) B"

using Z by (intro allI generator_Ex') auto
then obtain J' B' where J': "!!n. J' n ≠ {}" "!!n. finite (J' n)" "!!n. J' n ⊆ I"
"!!n. B' n ∈ sets (ΠM i∈J' n. borel)"
and Z_emb: "!!n. Z n = emb I (J' n) (B' n)"
unfolding choice_iff by blast
moreover def J "λn. (\<Union>i≤n. J' i)"
moreover def B "λn. emb (J n) (J' n) (B' n)"
ultimately have J: "!!n. J n ≠ {}" "!!n. finite (J n)" "!!n. J n ⊆ I"
"!!n. B n ∈ sets (ΠM i∈J n. borel)"
by auto
have J_mono: "!!n m. n ≤ m ==> J n ⊆ J m"
unfolding J_def by force
have "∀n. ∃j. j ∈ J n" using J by blast
then obtain j where j: "!!n. j n ∈ J n"
unfolding choice_iff by blast
note [simp] = `!!n. finite (J n)`
from J Z_emb have Z_eq: "!!n. Z n = emb I (J n) (B n)" "!!n. Z n ∈ ?G"
unfolding J_def B_def by (subst prod_emb_trans) (insert Z, auto)

have "?a ≤ μG' (Z 0)" by (auto intro: INF_lower)
also have "… < ∞" using J by (auto simp: Z_eq μG'_eq PiP_finite proj_sets)
finally have "?a ≠ ∞" by simp
have "!!n. ¦μG' (Z n)¦ ≠ ∞" unfolding Z_eq using J J_mono
by (subst μG'_eq) (auto simp: PiP_finite proj_sets μG'_eq)

interpret finite_set_sequence J by unfold_locales simp
def Utn Un_to_nat
interpret function_to_finmap "J n" Utn "inv_into (J n) Utn" for n
by unfold_locales (auto simp: Utn_def)
def P' "λn. mapmeasure n (P (J n)) (λ_. borel)"
let ?SUP = "λn. SUP K : {K. K ⊆ fm n ` (B n) ∧ compact K}. emeasure (P' n) K"
{
fix n
interpret finite_measure "P (J n)" by unfold_locales
have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"
using J
by (auto simp: P'_def mapmeasure_PiM proj_space proj_sets)
also
have "… = ?SUP n"
proof (rule inner_regular)
show "emeasure (P' n) (space (P' n)) ≠ ∞"
unfolding P'_def
by (auto simp: P'_def mapmeasure_PiF fm_measurable proj_space proj_sets)
show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def)
next
show "fm n ` B n ∈ sets borel"
unfolding borel_eq_PiF_borel
by (auto simp del: J(2) simp: P'_def fm_image_measurable_finite proj_sets J)
qed
finally
have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n ≠ ∞" "?SUP n ≠ - ∞" by auto
} note R = this
have "∀n. ∃K. emeasure (P (J n)) (B n) - emeasure (P' n) K ≤ 2 powr (-n) * ?a
∧ compact K ∧ K ⊆ fm n ` B n"

proof
fix n
have "emeasure (P' n) (space (P' n)) ≠ ∞"
by (simp add: mapmeasure_PiF P'_def proj_space proj_sets)
then interpret finite_measure "P' n" ..
show "∃K. emeasure (P (J n)) (B n) - emeasure (P' n) K ≤ ereal (2 powr - real n) * ?a ∧
compact K ∧ K ⊆ fm n ` B n"

unfolding R
proof (rule ccontr)
assume H: "¬ (∃K'. ?SUP n - emeasure (P' n) K' ≤ ereal (2 powr - real n) * ?a ∧
compact K' ∧ K' ⊆ fm n ` B n)"

have "?SUP n ≤ ?SUP n - 2 powr (-n) * ?a"
proof (intro SUP_least)
fix K
assume "K ∈ {K. K ⊆ fm n ` B n ∧ compact K}"
with H have "¬ ?SUP n - emeasure (P' n) K ≤ 2 powr (-n) * ?a"
by auto
hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a"
unfolding not_less[symmetric] by simp
hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K"
using `0 < ?a` by (auto simp add: ereal_less_minus_iff ac_simps)
thus "?SUP n - 2 powr (-n) * ?a ≥ emeasure (P' n) K" by simp
qed
hence "?SUP n + 0 ≤ ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp
hence "?SUP n + 0 ≤ ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
hence "0 ≤ - (2 powr (-n) * ?a)"
using `?SUP _ ≠ ∞` `?SUP _ ≠ - ∞`
by (subst (asm) ereal_add_le_add_iff) (auto simp:)
moreover have "ereal (2 powr - real n) * ?a > 0" using `0 < ?a`
by (auto simp: ereal_zero_less_0_iff)
ultimately show False by simp
qed
qed
then obtain K' where K':
"!!n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) ≤ ereal (2 powr - real n) * ?a"
"!!n. compact (K' n)" "!!n. K' n ⊆ fm n ` B n"
unfolding choice_iff by blast
def K "λn. fm n -` K' n ∩ space (P (J n))"
have K_sets: "!!n. K n ∈ sets (PiM (J n) (λ_. borel))"
unfolding K_def proj_space
using compact_imp_closed[OF `compact (K' _)`]
by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
(auto simp: borel_eq_PiF_borel[symmetric])
have "!!n. K n ⊆ B n"
proof
fix x n
assume "x ∈ K n" hence fm_in: "fm n x ∈ fm n ` B n"
using K' by (force simp: K_def)
show "x ∈ B n"
apply (rule inj_on_image_mem_iff[OF inj_on_fm _ fm_in])
using `x ∈ K n` K_sets J[of n] sets_into_space
apply (auto simp: proj_space)
using J[of n] sets_into_space apply auto
done
qed
def Z' "λn. emb I (J n) (K n)"
have Z': "!!n. Z' n ⊆ Z n"
unfolding Z_eq unfolding Z'_def
proof (rule prod_emb_subsetI, safe)
fix n x assume "x ∈ K n"
hence "fm n x ∈ K' n" "x ∈ space (PiM (J n) (λ_. borel))"
by (simp_all add: K_def proj_space)
note this(1)
also have "K' n ⊆ fm n ` B n" by (simp add: K')
finally have "fm n x ∈ fm n ` B n" .
thus "x ∈ B n"
proof safe
fix y assume "y ∈ B n"
moreover
hence "y ∈ space (PiM (J n) (λ_. borel))" using J sets_into_space[of "B n" "P (J n)"]
by (auto simp add: proj_space proj_sets)
assume "fm n x = fm n y"
note inj_onD[OF inj_on_fm[OF space_borel],
OF `fm n x = fm n y` `x ∈ space _` `y ∈ space _`]
ultimately show "x ∈ B n" by simp
qed
qed
{ fix n
have "Z' n ∈ ?G" using K' unfolding Z'_def
apply (intro generatorI'[OF J(1-3)])
unfolding K_def proj_space
apply (rule measurable_sets[OF fm_measurable[of _ "Collect finite"]])
apply (auto simp add: P'_def borel_eq_PiF_borel[symmetric] compact_imp_closed)
done
}
def Y "λn. \<Inter>i∈{1..n}. Z' i"
hence "!!n k. Y (n + k) ⊆ Y n" by (induct_tac k) (auto simp: Y_def)
hence Y_mono: "!!n m. n ≤ m ==> Y m ⊆ Y n" by (auto simp: le_iff_add)
have Y_Z': "!!n. n ≥ 1 ==> Y n ⊆ Z' n" by (auto simp: Y_def)
hence Y_Z: "!!n. n ≥ 1 ==> Y n ⊆ Z n" using Z' by auto
have Y_notempty: "!!n. n ≥ 1 ==> (Y n) ≠ {}"
proof -
fix n::nat assume "n ≥ 1" hence "Y n ⊆ Z n" by fact
have "Y n = (\<Inter> i∈{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
by (auto simp: Y_def Z'_def)
also have "… = prod_emb I (λ_. borel) (J n) (\<Inter> i∈{1..n}. emb (J n) (J i) (K i))"
using `n ≥ 1`
by (subst prod_emb_INT) auto
finally
have Y_emb:
"Y n = prod_emb I (λ_. borel) (J n)
(\<Inter> i∈{1..n}. prod_emb (J n) (λ_. borel) (J i) (K i))"
.
hence "Y n ∈ ?G" using J J_mono K_sets `n ≥ 1` by (intro generatorI[OF _ _ _ _ Y_emb]) auto
hence "¦μG' (Y n)¦ ≠ ∞" unfolding Y_emb using J J_mono K_sets `n ≥ 1`
by (subst μG'_eq) (auto simp: PiP_finite proj_sets μG'_eq)
interpret finite_measure "(PiP (J n) (λ_. borel) P)"
proof
have "emeasure (PiP (J n) (λ_. borel) P) (J n ->E space borel) ≠ ∞"
using J by (subst emeasure_PiP) auto
thus "emeasure (PiP (J n) (λ_. borel) P) (space (PiP (J n) (λ_. borel) P)) ≠ ∞"
by (simp add: space_PiM)
qed
have "μG' (Z n) - μG' (Y n) = μG' (Z n - Y n)" using J J_mono K_sets `n ≥ 1`
apply (intro G.subtractive[OF positive_μG' additive_μG',
OF `I ≠ {}` `I ≠ {}` `Y n ∈ ?G` `Z n ∈ ?G` `Y n ⊆ Z n`, symmetric])
apply (subst μG'_spec[OF `J n ≠ {}` `finite (J n)` `J n ⊆ I` Y_emb])
apply auto done
also have subs: "Z n - Y n ⊆ (\<Union> i∈{1..n}. (Z i - Z' i))" using Z' Z `n ≥ 1`
unfolding Y_def
apply (auto simp: decseq_def Y_def)
proof -
case goal1 hence "x ∈ Z xa" by (metis set_mp)
with goal1 show "x ∈ Z' xa" by auto
qed
have "Z n - Y n ∈ ?G" "(\<Union> i∈{1..n}. (Z i - Z' i)) ∈ ?G"
using `Z' _ ∈ ?G` `Z _ ∈ ?G` `Y _ ∈ ?G` by auto
hence "μG' (Z n - Y n) ≤ μG' (\<Union> i∈{1..n}. (Z i - Z' i))"
using subs G.additive_increasing[OF positive_μG'[OF `I ≠ {}`] additive_μG'[OF `I ≠ {}`]]
unfolding increasing_def by auto
also have "… ≤ (∑ i∈{1..n}. μG' (Z i - Z' i))" using `Z _ ∈ ?G` `Z' _ ∈ ?G`
by (intro G.subadditive[OF positive_μG' additive_μG', OF `I ≠ {}` `I ≠ {}`]) auto
also have "… ≤ (∑ i∈{1..n}. 2 powr -real i * ?a)"
proof (rule setsum_mono)
fix i assume "i ∈ {1..n}" hence "i ≤ n" by simp
have "μG' (Z i - Z' i) = μG' (prod_emb I (λ_. borel) (J i) (B i - K i))"
unfolding Z'_def Z_eq by simp
also have "… = P (J i) (B i - K i)"
apply (subst μG'_eq) using J K_sets apply auto
apply (subst PiP_finite) apply auto
done
also have "… = P (J i) (B i) - P (J i) (K i)"
apply (subst emeasure_Diff) using K_sets J `K _ ⊆ B _` apply (auto simp: proj_sets)
done
also have "… = P (J i) (B i) - P' i (K' i)"
unfolding K_def P'_def
by (auto simp: mapmeasure_PiF proj_space proj_sets borel_eq_PiF_borel[symmetric]
compact_imp_closed[OF `compact (K' _)`] space_PiM)
also have "… ≤ ereal (2 powr - real i) * ?a" using K'(1)[of i] .
finally show "μG' (Z i - Z' i) ≤ (2 powr - real i) * ?a" .
qed
also have "… = (∑ i∈{1..n}. ereal (2 powr -real i) * ereal(real ?a))"
using `?a ≠ ∞` `?a ≠ - ∞` by (subst ereal_real') auto
also have "… = ereal (∑ i∈{1..n}. (2 powr -real i) * (real ?a))" by simp
also have "… = ereal ((∑ i∈{1..n}. (2 powr -real i)) * real ?a)"
by (simp add: setsum_left_distrib)
also have "… < ereal (1 * real ?a)" unfolding less_ereal.simps
proof (rule mult_strict_right_mono)
have "(∑i∈{1..n}. 2 powr - real i) = (∑i∈{1..<Suc n}. (1/2) ^ i)"
by (rule setsum_cong)
(auto simp: powr_realpow[symmetric] powr_minus powr_divide inverse_eq_divide)
also have "{1..<Suc n} = {0..<Suc n} - {0}" by auto
also have "setsum (op ^ (1 / 2::real)) ({0..<Suc n} - {0}) =
setsum (op ^ (1 / 2)) ({0..<Suc n}) - 1"
by (auto simp: setsum_diff1)
also have "… < 1" by (subst sumr_geometric) auto
finally show "(∑i = 1..n. 2 powr - real i) < 1" .
qed (auto simp:
`0 < ?a` `?a ≠ ∞` `?a ≠ - ∞` ereal_less_real_iff zero_ereal_def[symmetric])
also have "… = ?a" using `0 < ?a` `?a ≠ ∞` by (auto simp: ereal_real')
also have "… ≤ μG' (Z n)" by (auto intro: INF_lower)
finally have "μG' (Z n) - μG' (Y n) < μG' (Z n)" .
hence R: "μG' (Z n) < μG' (Z n) + μG' (Y n)"
using `¦μG' (Y n)¦ ≠ ∞` by (simp add: ereal_minus_less)
have "0 ≤ (- μG' (Z n)) + μG' (Z n)" using `¦μG' (Z n)¦ ≠ ∞` by auto
also have "… < (- μG' (Z n)) + (μG' (Z n) + μG' (Y n))"
apply (rule ereal_less_add[OF _ R]) using `¦μG' (Z n)¦ ≠ ∞` by auto
finally have "μG' (Y n) > 0"
using `¦μG' (Z n)¦ ≠ ∞` by (auto simp: ac_simps zero_ereal_def[symmetric])
thus "Y n ≠ {}" using positive_μG' `I ≠ {}` by (auto simp add: positive_def)
qed
hence "∀n∈{1..}. ∃y. y ∈ Y n" by auto
then obtain y where y: "!!n. n ≥ 1 ==> y n ∈ Y n" unfolding bchoice_iff by force
{
fix t and n m::nat
assume "1 ≤ n" "n ≤ m" hence "1 ≤ m" by simp
from Y_mono[OF `m ≥ n`] y[OF `1 ≤ m`] have "y m ∈ Y n" by auto
also have "… ⊆ Z' n" using Y_Z'[OF `1 ≤ n`] .
finally
have "fm n (restrict (y m) (J n)) ∈ K' n"
unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)"
using J by (simp add: fm_def)
ultimately have "fm n (y m) ∈ K' n" by simp
} note fm_in_K' = this
interpret finmap_seqs_into_compact "λn. K' (Suc n)" "λk. fm (Suc k) (y (Suc k))" borel
proof
fix n show "compact (K' n)" by fact
next
fix n
from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) ∈ Y (Suc n)" by auto
also have "… ⊆ Z' (Suc n)" using Y_Z' by auto
finally
have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) ∈ K' (Suc n)"
unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
thus "K' (Suc n) ≠ {}" by auto
fix k
assume "k ∈ K' (Suc n)"
with K'[of "Suc n"] sets_into_space have "k ∈ fm (Suc n) ` B (Suc n)" by auto
then obtain b where "k = fm (Suc n) b" by auto
thus "domain k = domain (fm (Suc n) (y (Suc n)))"
by (simp_all add: fm_def)
next
fix t and n m::nat
assume "n ≤ m" hence "Suc n ≤ Suc m" by simp
assume "t ∈ domain (fm (Suc n) (y (Suc n)))"
then obtain j where j: "t = Utn j" "j ∈ J (Suc n)" by auto
hence "j ∈ J (Suc m)" using J_mono[OF `Suc n ≤ Suc m`] by auto
have img: "fm (Suc n) (y (Suc m)) ∈ K' (Suc n)" using `n ≤ m`
by (intro fm_in_K') simp_all
show "(fm (Suc m) (y (Suc m)))F t ∈ (λk. (k)F t) ` K' (Suc n)"
apply (rule image_eqI[OF _ img])
using `j ∈ J (Suc n)` `j ∈ J (Suc m)`
unfolding j by (subst proj_fm, auto)+
qed
have "∀t. ∃z. (λi. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))F t) ----> z"
using diagonal_tendsto ..
then obtain z where z:
"!!t. (λi. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))F t) ----> z t"
unfolding choice_iff by blast
{
fix n :: nat assume "n ≥ 1"
have "!!i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)"
by simp
moreover
{
fix t
assume t: "t ∈ domain (finmap_of (Utn ` J n) z)"
hence "t ∈ Utn ` J n" by simp
then obtain j where j: "t = Utn j" "j ∈ J n" by auto
have "(λi. (fm n (y (Suc (diagseq i))))F t) ----> z t"
apply (subst (2) tendsto_iff, subst eventually_sequentially)
proof safe
fix e :: real assume "0 < e"
{ fix i x assume "i ≥ n" "t ∈ domain (fm n x)"
moreover
hence "t ∈ domain (fm i x)" using J_mono[OF `i ≥ n`] by auto
ultimately have "(fm i x)F t = (fm n x)F t"
using j by (auto simp: proj_fm dest!:
Un_to_nat_injectiveD[simplified Utn_def[symmetric]])
} note index_shift = this
have I: "!!i. i ≥ n ==> Suc (diagseq i) ≥ n"
apply (rule le_SucI)
apply (rule order_trans) apply simp
apply (rule seq_suble[OF subseq_diagseq])
done
from z
have "∃N. ∀i≥N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))F t) (z t) < e"
unfolding tendsto_iff eventually_sequentially using `0 < e` by auto
then obtain N where N: "!!i. i ≥ N ==>
dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))F t) (z t) < e"
by auto
show "∃N. ∀na≥N. dist ((fm n (y (Suc (diagseq na))))F t) (z t) < e "
proof (rule exI[where x="max N n"], safe)
fix na assume "max N n ≤ na"
hence "dist ((fm n (y (Suc (diagseq na))))F t) (z t) =
dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))F t) (z t)"
using t
by (subst index_shift[OF I]) auto
also have "… < e" using `max N n ≤ na` by (intro N) simp
finally show "dist ((fm n (y (Suc (diagseq na))))F t) (z t) < e" .
qed
qed
hence "(λi. (fm n (y (Suc (diagseq i))))F t) ----> (finmap_of (Utn ` J n) z)F t"
by (simp add: tendsto_intros)
} ultimately
have "(λi. fm n (y (Suc (diagseq i)))) ----> finmap_of (Utn ` J n) z"
by (rule tendsto_finmap)
hence "((λi. fm n (y (Suc (diagseq i)))) o (λi. i + n)) ----> finmap_of (Utn ` J n) z"
by (intro lim_subseq) (simp add: subseq_def)
moreover
have "(∀i. ((λi. fm n (y (Suc (diagseq i)))) o (λi. i + n)) i ∈ K' n)"
apply (auto simp add: o_def intro!: fm_in_K' `1 ≤ n` le_SucI)
apply (rule le_trans)
apply (rule le_add2)
using seq_suble[OF subseq_diagseq]
apply auto
done
moreover
from `compact (K' n)` have "closed (K' n)" by (rule compact_imp_closed)
ultimately
have "finmap_of (Utn ` J n) z ∈ K' n"
unfolding closed_sequential_limits by blast
also have "finmap_of (Utn ` J n) z = fm n (λi. z (Utn i))"
by (auto simp: finmap_eq_iff fm_def compose_def f_inv_into_f)
finally have "fm n (λi. z (Utn i)) ∈ K' n" .
moreover
let ?J = "\<Union>n. J n"
have "(?J ∩ J n) = J n" by auto
ultimately have "restrict (λi. z (Utn i)) (?J ∩ J n) ∈ K n"
unfolding K_def by (auto simp: proj_space space_PiM)
hence "restrict (λi. z (Utn i)) ?J ∈ Z' n" unfolding Z'_def
using J by (auto simp: prod_emb_def extensional_def)
also have "… ⊆ Z n" using Z' by simp
finally have "restrict (λi. z (Utn i)) ?J ∈ Z n" .
} note in_Z = this
hence "(\<Inter>i∈{1..}. Z i) ≠ {}" by auto
hence "(\<Inter>i. Z i) ≠ {}" using Z INT_decseq_offset[OF `decseq Z`] by simp
thus False using Z by simp
qed
ultimately show "(λi. μG' (Z i)) ----> 0"
using LIMSEQ_ereal_INFI[of "λi. μG' (Z i)"] by simp
qed
then guess μ .. note μ = this
def f "finmap_of J B"
have "emeasure (PiB I P) (emb I J (PiE J B)) =
emeasure (PiB I P) (emb I (domain f) (PiE (domain f) (f)F))"

using assms sets_into_space
by (simp add: f_def Pi_def)
also have "… = emeasure (PiB J P) (PiE J B)"
proof (subst emeasure_extend_measure[OF PiP_def, of I "λ_. borel" μ])
show "positive (sets (PiB I P)) μ" "countably_additive (sets (PiB I P)) μ"
using μ unfolding sets_PiP sets_PiM_generator[OF `I ≠ {}`] by (auto simp: measure_space_def)
next
show "f ∈ {x. (domain x ≠ {} ∨ I = {}) ∧ finite (domain x) ∧ domain x ⊆ I ∧
(x)F ∈ (ΠE i∈domain x. sets borel)}"

using assms by (auto simp: f_def)
next
show "(λx. emb I (domain x) (PiE (domain x) (x)F)) `
{x. (domain x ≠ {} ∨ I = {}) ∧ finite (domain x) ∧ domain x ⊆ I ∧
(x)F ∈ (ΠE i∈domain x. sets borel)}
⊆ Pow (ΠE i∈I. space borel)"
by (auto simp: prod_emb_def)
next
fix i::"'i =>F 'a set"
assume i: "i ∈ {x. (domain x ≠ {} ∨ I = {}) ∧ finite (domain x) ∧ domain x ⊆ I ∧
(x)F ∈ (ΠE i∈domain x. sets borel)}"

hence "emb I (domain i) (PiE (domain i) (i)F) ∈ generator"
using assms by (auto intro!: generatorI')
hence "μ (emb I (domain i) (PiE (domain i) (i)F)) =
μG' (emb I (domain i) (PiE (domain i) (i)F))"

using μ by simp
also have "… = emeasure (P (domain i)) (PiE (domain i) (i)F)"
using i assms proj_sets by (subst μG'_eq) (auto simp: μG'_eq PiP_finite)
finally show "μ (emb I (domain i) (PiE (domain i) (i)F)) =
emeasure (P (domain i)) (PiE (domain i) (i)F)"
.
next
show "emeasure (P (domain f)) (PiE (domain f) (f)F) = emeasure (PiB J P) (PiE J B)"
using assms by (simp add: f_def PiP_finite Pi_def)
qed
finally show ?thesis .
qed

end

sublocale polish_projective P: prob_space "(PiB I P)"
proof
show "emeasure (PiB I P) (space (PiB I P)) = 1"
proof cases
assume "I = {}" then show ?thesis
by (simp add: space_PiM_empty PiP_finite prob_space.emeasure_space_1)
next
assume "I ≠ {}"
then obtain i where "i ∈ I" by auto
moreover then have R: "(space (PiB I P)) = (emb I {i} (PiE {i} (λ_. space borel)))"
by (auto simp: prod_emb_def space_PiM)
moreover have "extensional {i} = space (P {i})" by (simp add: proj_space space_PiM)
ultimately show ?thesis
apply (subst R)
apply (subst emeasure_PiB_emb_not_empty)
apply (auto simp: PiP_finite prob_space.emeasure_space_1)
done
qed
qed

context polish_projective begin

lemma emeasure_PiB_emb:
assumes X: "J ⊆ I" "finite J" "∀i∈J. B i ∈ sets borel"
shows "emeasure (PiB I P) (emb I J (PiE J B)) = emeasure (P J) (PiE J B)"
proof cases
assume "J = {}"
moreover have "emb I {} {λx. undefined} = space (PiB I P)"
by (auto simp: space_PiM prod_emb_def)
moreover have "{λx. undefined} = space (PiB {} P)"
by (auto simp: space_PiM prod_emb_def)
ultimately show ?thesis
by (simp add: P.emeasure_space_1 PiP_finite prob_space.emeasure_space_1 del: space_PiP)
next
assume "J ≠ {}" with X show ?thesis
by (subst emeasure_PiB_emb_not_empty) (auto simp: PiP_finite)
qed

lemma measure_PiB_emb:
assumes "J ⊆ I" "finite J" "∀i∈J. B i ∈ sets borel"
shows "measure (PiB I P) (emb I J (PiE J B)) = measure (P J) (PiE J B)"
using emeasure_PiB_emb[OF assms]
unfolding emeasure_eq_measure PiP_finite[OF `finite J` `J ⊆ I`] prob_space.emeasure_eq_measure
by simp

end

end