| 
50088
 | 
     1  | 
(*  Title:      HOL/Probability/Projective_Family.thy
  | 
| 
 | 
     2  | 
    Author:     Fabian Immler, TU München
  | 
| 
 | 
     3  | 
*)
  | 
| 
 | 
     4  | 
  | 
| 
 | 
     5  | 
theory Fin_Map
  | 
| 
 | 
     6  | 
imports Finite_Product_Measure
  | 
| 
 | 
     7  | 
begin
  | 
| 
 | 
     8  | 
  | 
| 
 | 
     9  | 
section {* Finite Maps *}
 | 
| 
 | 
    10  | 
  | 
| 
 | 
    11  | 
text {* Auxiliary type that is instantiated to @{class polish_space}, needed for the proof of
 | 
| 
 | 
    12  | 
  projective limit. @{const extensional} functions are used for the representation in order to
 | 
| 
 | 
    13  | 
  stay close to the developments of (finite) products @{const Pi\<^isub>E} and their sigma-algebra
 | 
| 
 | 
    14  | 
  @{const Pi\<^isub>M}. *}
 | 
| 
 | 
    15  | 
  | 
| 
 | 
    16  | 
typedef ('i, 'a) finmap ("(_ \<Rightarrow>\<^isub>F /_)" [22, 21] 21) =
 | 
| 
 | 
    17  | 
  "{(I::'i set, f::'i \<Rightarrow> 'a). finite I \<and> f \<in> extensional I}" by auto
 | 
| 
 | 
    18  | 
  | 
| 
 | 
    19  | 
subsection {* Domain and Application *}
 | 
| 
 | 
    20  | 
  | 
| 
 | 
    21  | 
definition domain where "domain P = fst (Rep_finmap P)"
  | 
| 
 | 
    22  | 
  | 
| 
 | 
    23  | 
lemma finite_domain[simp, intro]: "finite (domain P)"
  | 
| 
 | 
    24  | 
  by (cases P) (auto simp: domain_def Abs_finmap_inverse)
  | 
| 
 | 
    25  | 
  | 
| 
 | 
    26  | 
definition proj ("_\<^isub>F" [1000] 1000) where "proj P i = snd (Rep_finmap P) i"
 | 
| 
 | 
    27  | 
  | 
| 
 | 
    28  | 
declare [[coercion proj]]
  | 
| 
 | 
    29  | 
  | 
| 
 | 
    30  | 
lemma extensional_proj[simp, intro]: "(P)\<^isub>F \<in> extensional (domain P)"
  | 
| 
 | 
    31  | 
  by (cases P) (auto simp: domain_def Abs_finmap_inverse proj_def[abs_def])
  | 
| 
 | 
    32  | 
  | 
| 
 | 
    33  | 
lemma proj_undefined[simp, intro]: "i \<notin> domain P \<Longrightarrow> P i = undefined"
  | 
| 
 | 
    34  | 
  using extensional_proj[of P] unfolding extensional_def by auto
  | 
| 
 | 
    35  | 
  | 
| 
 | 
    36  | 
lemma finmap_eq_iff: "P = Q \<longleftrightarrow> (domain P = domain Q \<and> (\<forall>i\<in>domain P. P i = Q i))"
  | 
| 
 | 
    37  | 
  by (cases P, cases Q)
  | 
| 
 | 
    38  | 
     (auto simp add: Abs_finmap_inject extensional_def domain_def proj_def Abs_finmap_inverse
  | 
| 
 | 
    39  | 
              intro: extensionalityI)
  | 
| 
 | 
    40  | 
  | 
| 
 | 
    41  | 
subsection {* Countable Finite Maps *}
 | 
| 
 | 
    42  | 
  | 
| 
 | 
    43  | 
instance finmap :: (countable, countable) countable
  | 
| 
 | 
    44  | 
proof
  | 
| 
 | 
    45  | 
  obtain mapper where mapper: "\<And>fm :: 'a \<Rightarrow>\<^isub>F 'b. set (mapper fm) = domain fm"
  | 
| 
 | 
    46  | 
    by (metis finite_list[OF finite_domain])
  | 
| 
 | 
    47  | 
  have "inj (\<lambda>fm. map (\<lambda>i. (i, (fm)\<^isub>F i)) (mapper fm))" (is "inj ?F")
  | 
| 
 | 
    48  | 
  proof (rule inj_onI)
  | 
| 
 | 
    49  | 
    fix f1 f2 assume "?F f1 = ?F f2"
  | 
| 
 | 
    50  | 
    then have "map fst (?F f1) = map fst (?F f2)" by simp
  | 
| 
 | 
    51  | 
    then have "mapper f1 = mapper f2" by (simp add: comp_def)
  | 
| 
 | 
    52  | 
    then have "domain f1 = domain f2" by (simp add: mapper[symmetric])
  | 
| 
 | 
    53  | 
    with `?F f1 = ?F f2` show "f1 = f2"
  | 
| 
 | 
    54  | 
      unfolding `mapper f1 = mapper f2` map_eq_conv mapper
  | 
| 
 | 
    55  | 
      by (simp add: finmap_eq_iff)
  | 
| 
 | 
    56  | 
  qed
  | 
| 
 | 
    57  | 
  then show "\<exists>to_nat :: 'a \<Rightarrow>\<^isub>F 'b \<Rightarrow> nat. inj to_nat"
  | 
| 
 | 
    58  | 
    by (intro exI[of _ "to_nat \<circ> ?F"] inj_comp) auto
  | 
| 
 | 
    59  | 
qed
  | 
| 
 | 
    60  | 
  | 
| 
 | 
    61  | 
subsection {* Constructor of Finite Maps *}
 | 
| 
 | 
    62  | 
  | 
| 
 | 
    63  | 
definition "finmap_of inds f = Abs_finmap (inds, restrict f inds)"
  | 
| 
 | 
    64  | 
  | 
| 
 | 
    65  | 
lemma proj_finmap_of[simp]:
  | 
| 
 | 
    66  | 
  assumes "finite inds"
  | 
| 
 | 
    67  | 
  shows "(finmap_of inds f)\<^isub>F = restrict f inds"
  | 
| 
 | 
    68  | 
  using assms
  | 
| 
 | 
    69  | 
  by (auto simp: Abs_finmap_inverse finmap_of_def proj_def)
  | 
| 
 | 
    70  | 
  | 
| 
 | 
    71  | 
lemma domain_finmap_of[simp]:
  | 
| 
 | 
    72  | 
  assumes "finite inds"
  | 
| 
 | 
    73  | 
  shows "domain (finmap_of inds f) = inds"
  | 
| 
 | 
    74  | 
  using assms
  | 
| 
 | 
    75  | 
  by (auto simp: Abs_finmap_inverse finmap_of_def domain_def)
  | 
| 
 | 
    76  | 
  | 
| 
 | 
    77  | 
lemma finmap_of_eq_iff[simp]:
  | 
| 
 | 
    78  | 
  assumes "finite i" "finite j"
  | 
| 
 | 
    79  | 
  shows "finmap_of i m = finmap_of j n \<longleftrightarrow> i = j \<and> restrict m i = restrict n i"
  | 
| 
 | 
    80  | 
  using assms
  | 
| 
 | 
    81  | 
  apply (auto simp: finmap_eq_iff restrict_def) by metis
  | 
| 
 | 
    82  | 
  | 
| 
 | 
    83  | 
lemma
  | 
| 
 | 
    84  | 
  finmap_of_inj_on_extensional_finite:
  | 
| 
 | 
    85  | 
  assumes "finite K"
  | 
| 
 | 
    86  | 
  assumes "S \<subseteq> extensional K"
  | 
| 
 | 
    87  | 
  shows "inj_on (finmap_of K) S"
  | 
| 
 | 
    88  | 
proof (rule inj_onI)
  | 
| 
 | 
    89  | 
  fix x y::"'a \<Rightarrow> 'b"
  | 
| 
 | 
    90  | 
  assume "finmap_of K x = finmap_of K y"
  | 
| 
 | 
    91  | 
  hence "(finmap_of K x)\<^isub>F = (finmap_of K y)\<^isub>F" by simp
  | 
| 
 | 
    92  | 
  moreover
  | 
| 
 | 
    93  | 
  assume "x \<in> S" "y \<in> S" hence "x \<in> extensional K" "y \<in> extensional K" using assms by auto
  | 
| 
 | 
    94  | 
  ultimately
  | 
| 
 | 
    95  | 
  show "x = y" using assms by (simp add: extensional_restrict)
  | 
| 
 | 
    96  | 
qed
  | 
| 
 | 
    97  | 
  | 
| 
 | 
    98  | 
lemma finmap_choice:
  | 
| 
 | 
    99  | 
  assumes *: "\<And>i. i \<in> I \<Longrightarrow> \<exists>x. P i x" and I: "finite I"
  | 
| 
 | 
   100  | 
  shows "\<exists>fm. domain fm = I \<and> (\<forall>i\<in>I. P i (fm i))"
  | 
| 
 | 
   101  | 
proof -
  | 
| 
 | 
   102  | 
  have "\<exists>f. \<forall>i\<in>I. P i (f i)"
  | 
| 
 | 
   103  | 
    unfolding bchoice_iff[symmetric] using * by auto
  | 
| 
 | 
   104  | 
  then guess f ..
  | 
| 
 | 
   105  | 
  with I show ?thesis
  | 
| 
 | 
   106  | 
    by (intro exI[of _ "finmap_of I f"]) auto
  | 
| 
 | 
   107  | 
qed
  | 
| 
 | 
   108  | 
  | 
| 
 | 
   109  | 
subsection {* Product set of Finite Maps *}
 | 
| 
 | 
   110  | 
  | 
| 
 | 
   111  | 
text {* This is @{term Pi} for Finite Maps, most of this is copied *}
 | 
| 
 | 
   112  | 
  | 
| 
 | 
   113  | 
definition Pi' :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a set) \<Rightarrow> ('i \<Rightarrow>\<^isub>F 'a) set" where
 | 
| 
 | 
   114  | 
  "Pi' I A = { P. domain P = I \<and> (\<forall>i. i \<in> I \<longrightarrow> (P)\<^isub>F i \<in> A i) } "
 | 
| 
 | 
   115  | 
  | 
| 
 | 
   116  | 
syntax
  | 
| 
 | 
   117  | 
  "_Pi'"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PI' _:_./ _)" 10)
 | 
| 
 | 
   118  | 
  | 
| 
 | 
   119  | 
syntax (xsymbols)
  | 
| 
 | 
   120  | 
  "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>' _\<in>_./ _)"   10)
 | 
| 
 | 
   121  | 
  | 
| 
 | 
   122  | 
syntax (HTML output)
  | 
| 
 | 
   123  | 
  "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>' _\<in>_./ _)"   10)
 | 
| 
 | 
   124  | 
  | 
| 
 | 
   125  | 
translations
  | 
| 
 | 
   126  | 
  "PI' x:A. B" == "CONST Pi' A (%x. B)"
  | 
| 
 | 
   127  | 
  | 
| 
 | 
   128  | 
abbreviation
  | 
| 
 | 
   129  | 
  finmapset :: "['a set, 'b set] => ('a \<Rightarrow>\<^isub>F 'b) set"
 | 
| 
 | 
   130  | 
    (infixr "~>" 60) where
  | 
| 
 | 
   131  | 
  "A ~> B \<equiv> Pi' A (%_. B)"
  | 
| 
 | 
   132  | 
  | 
| 
 | 
   133  | 
notation (xsymbols)
  | 
| 
 | 
   134  | 
  finmapset  (infixr "\<leadsto>" 60)
  | 
| 
 | 
   135  | 
  | 
| 
 | 
   136  | 
subsubsection{*Basic Properties of @{term Pi'}*}
 | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
lemma Pi'_I[intro!]: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B x) \<Longrightarrow> f \<in> Pi' A B"
  | 
| 
 | 
   139  | 
  by (simp add: Pi'_def)
  | 
| 
 | 
   140  | 
  | 
| 
 | 
   141  | 
lemma Pi'_I'[simp]: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<longrightarrow> f x \<in> B x) \<Longrightarrow> f \<in> Pi' A B"
  | 
| 
 | 
   142  | 
  by (simp add:Pi'_def)
  | 
| 
 | 
   143  | 
  | 
| 
 | 
   144  | 
lemma finmapsetI: "domain f = A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> f \<in> A \<leadsto> B"
  | 
| 
 | 
   145  | 
  by (simp add: Pi_def)
  | 
| 
 | 
   146  | 
  | 
| 
 | 
   147  | 
lemma Pi'_mem: "f\<in> Pi' A B \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<in> B x"
  | 
| 
 | 
   148  | 
  by (simp add: Pi'_def)
  | 
| 
 | 
   149  | 
  | 
| 
 | 
   150  | 
lemma Pi'_iff: "f \<in> Pi' I X \<longleftrightarrow> domain f = I \<and> (\<forall>i\<in>I. f i \<in> X i)"
  | 
| 
 | 
   151  | 
  unfolding Pi'_def by auto
  | 
| 
 | 
   152  | 
  | 
| 
 | 
   153  | 
lemma Pi'E [elim]:
  | 
| 
 | 
   154  | 
  "f \<in> Pi' A B \<Longrightarrow> (f x \<in> B x \<Longrightarrow> domain f = A \<Longrightarrow> Q) \<Longrightarrow> (x \<notin> A \<Longrightarrow> Q) \<Longrightarrow> Q"
  | 
| 
 | 
   155  | 
  by(auto simp: Pi'_def)
  | 
| 
 | 
   156  | 
  | 
| 
 | 
   157  | 
lemma in_Pi'_cong:
  | 
| 
 | 
   158  | 
  "domain f = domain g \<Longrightarrow> (\<And> w. w \<in> A \<Longrightarrow> f w = g w) \<Longrightarrow> f \<in> Pi' A B \<longleftrightarrow> g \<in> Pi' A B"
  | 
| 
 | 
   159  | 
  by (auto simp: Pi'_def)
  | 
| 
 | 
   160  | 
  | 
| 
 | 
   161  | 
lemma funcset_mem: "[|f \<in> A \<leadsto> B; x \<in> A|] ==> f x \<in> B"
  | 
| 
 | 
   162  | 
  by (simp add: Pi'_def)
  | 
| 
 | 
   163  | 
  | 
| 
 | 
   164  | 
lemma funcset_image: "f \<in> A \<leadsto> B ==> f ` A \<subseteq> B"
  | 
| 
 | 
   165  | 
by auto
  | 
| 
 | 
   166  | 
  | 
| 
 | 
   167  | 
lemma Pi'_eq_empty[simp]:
  | 
| 
 | 
   168  | 
  assumes "finite A" shows "(Pi' A B) = {} \<longleftrightarrow> (\<exists>x\<in>A. B x = {})"
 | 
| 
 | 
   169  | 
  using assms
  | 
| 
 | 
   170  | 
  apply (simp add: Pi'_def, auto)
  | 
| 
 | 
   171  | 
  apply (drule_tac x = "finmap_of A (\<lambda>u. SOME y. y \<in> B u)" in spec, auto)
  | 
| 
 | 
   172  | 
  apply (cut_tac P= "%y. y \<in> B i" in some_eq_ex, auto)
  | 
| 
 | 
   173  | 
  done
  | 
| 
 | 
   174  | 
  | 
| 
 | 
   175  | 
lemma Pi'_mono: "(\<And>x. x \<in> A \<Longrightarrow> B x \<subseteq> C x) \<Longrightarrow> Pi' A B \<subseteq> Pi' A C"
  | 
| 
 | 
   176  | 
  by (auto simp: Pi'_def)
  | 
| 
 | 
   177  | 
  | 
| 
 | 
   178  | 
lemma Pi_Pi': "finite A \<Longrightarrow> (Pi\<^isub>E A B) = proj ` Pi' A B"
  | 
| 
 | 
   179  | 
  apply (auto simp: Pi'_def Pi_def extensional_def)
  | 
| 
 | 
   180  | 
  apply (rule_tac x = "finmap_of A (restrict x A)" in image_eqI)
  | 
| 
 | 
   181  | 
  apply auto
  | 
| 
 | 
   182  | 
  done
  | 
| 
 | 
   183  | 
  | 
| 
 | 
   184  | 
subsection {* Metric Space of Finite Maps *}
 | 
| 
 | 
   185  | 
  | 
| 
 | 
   186  | 
instantiation finmap :: (type, metric_space) metric_space
  | 
| 
 | 
   187  | 
begin
  | 
| 
 | 
   188  | 
  | 
| 
 | 
   189  | 
definition dist_finmap where
  | 
| 
 | 
   190  | 
  "dist P Q = (\<Sum>i\<in>domain P \<union> domain Q. dist ((P)\<^isub>F i) ((Q)\<^isub>F i)) +
  | 
| 
 | 
   191  | 
    card ((domain P - domain Q) \<union> (domain Q - domain P))"
  | 
| 
 | 
   192  | 
  | 
| 
 | 
   193  | 
lemma dist_finmap_extend:
  | 
| 
 | 
   194  | 
  assumes "finite X"
  | 
| 
 | 
   195  | 
  shows "dist P Q = (\<Sum>i\<in>domain P \<union> domain Q \<union> X. dist ((P)\<^isub>F i) ((Q)\<^isub>F i)) +
  | 
| 
 | 
   196  | 
    card ((domain P - domain Q) \<union> (domain Q - domain P))"
  | 
| 
 | 
   197  | 
    unfolding dist_finmap_def add_right_cancel
  | 
| 
 | 
   198  | 
    using assms extensional_arb[of "(P)\<^isub>F"] extensional_arb[of "(Q)\<^isub>F" "domain Q"]
  | 
| 
 | 
   199  | 
    by (intro setsum_mono_zero_cong_left) auto
  | 
| 
 | 
   200  | 
  | 
| 
 | 
   201  | 
definition open_finmap :: "('a \<Rightarrow>\<^isub>F 'b) set \<Rightarrow> bool" where
 | 
| 
 | 
   202  | 
  "open_finmap S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
  | 
| 
 | 
   203  | 
  | 
| 
 | 
   204  | 
lemma add_eq_zero_iff[simp]:
  | 
| 
 | 
   205  | 
  fixes a b::real
  | 
| 
 | 
   206  | 
  assumes "a \<ge> 0" "b \<ge> 0"
  | 
| 
 | 
   207  | 
  shows "a + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
  | 
| 
 | 
   208  | 
using assms by auto
  | 
| 
 | 
   209  | 
  | 
| 
 | 
   210  | 
lemma dist_le_1_imp_domain_eq:
  | 
| 
 | 
   211  | 
  assumes "dist P Q < 1"
  | 
| 
 | 
   212  | 
  shows "domain P = domain Q"
  | 
| 
 | 
   213  | 
proof -
  | 
| 
 | 
   214  | 
  have "0 \<le> (\<Sum>i\<in>domain P \<union> domain Q. dist (P i) (Q i))"
  | 
| 
 | 
   215  | 
    by (simp add: setsum_nonneg)
  | 
| 
 | 
   216  | 
  with assms have "card (domain P - domain Q \<union> (domain Q - domain P)) = 0"
  | 
| 
 | 
   217  | 
    unfolding dist_finmap_def by arith
  | 
| 
 | 
   218  | 
  thus "domain P = domain Q" by auto
  | 
| 
 | 
   219  | 
qed
  | 
| 
 | 
   220  | 
  | 
| 
 | 
   221  | 
lemma dist_proj:
  | 
| 
 | 
   222  | 
  shows "dist ((x)\<^isub>F i) ((y)\<^isub>F i) \<le> dist x y"
  | 
| 
 | 
   223  | 
proof -
  | 
| 
 | 
   224  | 
  have "dist (x i) (y i) = (\<Sum>i\<in>{i}. dist (x i) (y i))" by simp
 | 
| 
 | 
   225  | 
  also have "\<dots> \<le> (\<Sum>i\<in>domain x \<union> domain y \<union> {i}. dist (x i) (y i))"
 | 
| 
 | 
   226  | 
    by (intro setsum_mono2) auto
  | 
| 
 | 
   227  | 
  also have "\<dots> \<le> dist x y" by (simp add: dist_finmap_extend[of "{i}"])
 | 
| 
 | 
   228  | 
  finally show ?thesis by simp
  | 
| 
 | 
   229  | 
qed
  | 
| 
 | 
   230  | 
  | 
| 
 | 
   231  | 
lemma open_Pi'I:
  | 
| 
 | 
   232  | 
  assumes open_component: "\<And>i. i \<in> I \<Longrightarrow> open (A i)"
  | 
| 
 | 
   233  | 
  shows "open (Pi' I A)"
  | 
| 
 | 
   234  | 
proof (subst open_finmap_def, safe)
  | 
| 
 | 
   235  | 
  fix x assume x: "x \<in> Pi' I A"
  | 
| 
 | 
   236  | 
  hence dim_x: "domain x = I" by (simp add: Pi'_def)
  | 
| 
 | 
   237  | 
  hence [simp]: "finite I" unfolding dim_x[symmetric] by simp
  | 
| 
 | 
   238  | 
  have "\<exists>ei. \<forall>i\<in>I. 0 < ei i \<and> (\<forall>y. dist y (x i) < ei i \<longrightarrow> y \<in> A i)"
  | 
| 
 | 
   239  | 
  proof (safe intro!: bchoice)
  | 
| 
 | 
   240  | 
    fix i assume i: "i \<in> I"
  | 
| 
 | 
   241  | 
    moreover with open_component have "open (A i)" by simp
  | 
| 
 | 
   242  | 
    moreover have "x i \<in> A i" using x i
  | 
| 
 | 
   243  | 
      by (auto simp: proj_def)
  | 
| 
 | 
   244  | 
    ultimately show "\<exists>e>0. \<forall>y. dist y (x i) < e \<longrightarrow> y \<in> A i"
  | 
| 
 | 
   245  | 
      using x by (auto simp: open_dist Ball_def)
  | 
| 
 | 
   246  | 
  qed
  | 
| 
 | 
   247  | 
  then guess ei .. note ei = this
  | 
| 
 | 
   248  | 
  def es \<equiv> "ei ` I"
  | 
| 
 | 
   249  | 
  def e \<equiv> "if es = {} then 0.5 else min 0.5 (Min es)"
 | 
| 
 | 
   250  | 
  from ei have "e > 0" using x
  | 
| 
 | 
   251  | 
    by (auto simp add: e_def es_def Pi'_def Ball_def)
  | 
| 
 | 
   252  | 
  moreover have "\<forall>y. dist y x < e \<longrightarrow> y \<in> Pi' I A"
  | 
| 
 | 
   253  | 
  proof (intro allI impI)
  | 
| 
 | 
   254  | 
    fix y
  | 
| 
 | 
   255  | 
    assume "dist y x < e"
  | 
| 
 | 
   256  | 
    also have "\<dots> < 1" by (auto simp: e_def)
  | 
| 
 | 
   257  | 
    finally have "domain y = domain x" by (rule dist_le_1_imp_domain_eq)
  | 
| 
 | 
   258  | 
    with dim_x have dims: "domain y = domain x" "domain x = I" by auto
  | 
| 
 | 
   259  | 
    show "y \<in> Pi' I A"
  | 
| 
 | 
   260  | 
    proof
  | 
| 
 | 
   261  | 
      show "domain y = I" using dims by simp
  | 
| 
 | 
   262  | 
    next
  | 
| 
 | 
   263  | 
      fix i
  | 
| 
 | 
   264  | 
      assume "i \<in> I"
  | 
| 
 | 
   265  | 
      have "dist (y i) (x i) \<le> dist y x" using dims `i \<in> I`
  | 
| 
 | 
   266  | 
        by (auto intro: dist_proj)
  | 
| 
 | 
   267  | 
      also have "\<dots> < e" using `dist y x < e` dims
  | 
| 
 | 
   268  | 
        by (simp add: dist_finmap_def)
  | 
| 
 | 
   269  | 
      also have "e \<le> Min (ei ` I)" using dims `i \<in> I`
  | 
| 
 | 
   270  | 
        by (auto simp: e_def es_def)
  | 
| 
 | 
   271  | 
      also have "\<dots> \<le> ei i" using `i \<in> I` by (simp add: e_def)
  | 
| 
 | 
   272  | 
      finally have "dist (y i) (x i) < ei i" .
  | 
| 
 | 
   273  | 
      with ei `i \<in> I` show "y i \<in> A  i" by simp
  | 
| 
 | 
   274  | 
    qed
  | 
| 
 | 
   275  | 
  qed
  | 
| 
 | 
   276  | 
  ultimately
  | 
| 
 | 
   277  | 
  show "\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> Pi' I A" by blast
  | 
| 
 | 
   278  | 
qed
  | 
| 
 | 
   279  | 
  | 
| 
 | 
   280  | 
instance
  | 
| 
 | 
   281  | 
proof
  | 
| 
 | 
   282  | 
  fix S::"('a \<Rightarrow>\<^isub>F 'b) set"
 | 
| 
 | 
   283  | 
  show "open S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)"
  | 
| 
 | 
   284  | 
    unfolding open_finmap_def ..
  | 
| 
 | 
   285  | 
next
  | 
| 
 | 
   286  | 
  fix P Q::"'a \<Rightarrow>\<^isub>F 'b"
  | 
| 
 | 
   287  | 
  show "dist P Q = 0 \<longleftrightarrow> P = Q"
  | 
| 
 | 
   288  | 
    by (auto simp: finmap_eq_iff dist_finmap_def setsum_nonneg setsum_nonneg_eq_0_iff)
  | 
| 
 | 
   289  | 
next
  | 
| 
 | 
   290  | 
  fix P Q R::"'a \<Rightarrow>\<^isub>F 'b"
  | 
| 
 | 
   291  | 
  let ?symdiff = "\<lambda>a b. domain a - domain b \<union> (domain b - domain a)"
  | 
| 
 | 
   292  | 
  def E \<equiv> "domain P \<union> domain Q \<union> domain R"
  | 
| 
 | 
   293  | 
  hence "finite E" by (simp add: E_def)
  | 
| 
 | 
   294  | 
  have "card (?symdiff P Q) \<le> card (?symdiff P R \<union> ?symdiff Q R)"
  | 
| 
 | 
   295  | 
    by (auto intro: card_mono)
  | 
| 
 | 
   296  | 
  also have "\<dots> \<le> card (?symdiff P R) + card (?symdiff Q R)"
  | 
| 
 | 
   297  | 
    by (subst card_Un_Int) auto
  | 
| 
 | 
   298  | 
  finally have "dist P Q \<le> (\<Sum>i\<in>E. dist (P i) (R i) + dist (Q i) (R i)) +
  | 
| 
 | 
   299  | 
    real (card (?symdiff P R) + card (?symdiff Q R))"
  | 
| 
 | 
   300  | 
    unfolding dist_finmap_extend[OF `finite E`]
  | 
| 
 | 
   301  | 
    by (intro add_mono) (auto simp: E_def intro: setsum_mono dist_triangle_le)
  | 
| 
 | 
   302  | 
  also have "\<dots> \<le> dist P R + dist Q R"
  | 
| 
 | 
   303  | 
    unfolding dist_finmap_extend[OF `finite E`] by (simp add: ac_simps E_def setsum_addf[symmetric])
  | 
| 
 | 
   304  | 
  finally show "dist P Q \<le> dist P R + dist Q R" by simp
  | 
| 
 | 
   305  | 
qed
  | 
| 
 | 
   306  | 
  | 
| 
 | 
   307  | 
end
  | 
| 
 | 
   308  | 
  | 
| 
 | 
   309  | 
lemma open_restricted_space:
  | 
| 
 | 
   310  | 
  shows "open {m. P (domain m)}"
 | 
| 
 | 
   311  | 
proof -
  | 
| 
 | 
   312  | 
  have "{m. P (domain m)} = (\<Union>i \<in> Collect P. {m. domain m = i})" by auto
 | 
| 
 | 
   313  | 
  also have "open \<dots>"
  | 
| 
 | 
   314  | 
  proof (rule, safe, cases)
  | 
| 
 | 
   315  | 
    fix i::"'a set"
  | 
| 
 | 
   316  | 
    assume "finite i"
  | 
| 
 | 
   317  | 
    hence "{m. domain m = i} = Pi' i (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
 | 
| 
 | 
   318  | 
    also have "open \<dots>" by (auto intro: open_Pi'I simp: `finite i`)
  | 
| 
 | 
   319  | 
    finally show "open {m. domain m = i}" .
 | 
| 
 | 
   320  | 
  next
  | 
| 
 | 
   321  | 
    fix i::"'a set"
  | 
| 
 | 
   322  | 
    assume "\<not> finite i" hence "{m. domain m = i} = {}" by auto
 | 
| 
 | 
   323  | 
    also have "open \<dots>" by simp
  | 
| 
 | 
   324  | 
    finally show "open {m. domain m = i}" .
 | 
| 
 | 
   325  | 
  qed
  | 
| 
 | 
   326  | 
  finally show ?thesis .
  | 
| 
 | 
   327  | 
qed
  | 
| 
 | 
   328  | 
  | 
| 
 | 
   329  | 
lemma closed_restricted_space:
  | 
| 
 | 
   330  | 
  shows "closed {m. P (domain m)}"
 | 
| 
 | 
   331  | 
proof -
  | 
| 
 | 
   332  | 
  have "{m. P (domain m)} = - (\<Union>i \<in> - Collect P. {m. domain m = i})" by auto
 | 
| 
 | 
   333  | 
  also have "closed \<dots>"
  | 
| 
 | 
   334  | 
  proof (rule, rule, rule, cases)
  | 
| 
 | 
   335  | 
    fix i::"'a set"
  | 
| 
 | 
   336  | 
    assume "finite i"
  | 
| 
 | 
   337  | 
    hence "{m. domain m = i} = Pi' i (\<lambda>_. UNIV)" by (auto simp: Pi'_def)
 | 
| 
 | 
   338  | 
    also have "open \<dots>" by (auto intro: open_Pi'I simp: `finite i`)
  | 
| 
 | 
   339  | 
    finally show "open {m. domain m = i}" .
 | 
| 
 | 
   340  | 
  next
  | 
| 
 | 
   341  | 
    fix i::"'a set"
  | 
| 
 | 
   342  | 
    assume "\<not> finite i" hence "{m. domain m = i} = {}" by auto
 | 
| 
 | 
   343  | 
    also have "open \<dots>" by simp
  | 
| 
 | 
   344  | 
    finally show "open {m. domain m = i}" .
 | 
| 
 | 
   345  | 
  qed
  | 
| 
 | 
   346  | 
  finally show ?thesis .
  | 
| 
 | 
   347  | 
qed
  | 
| 
 | 
   348  | 
  | 
| 
 | 
   349  | 
lemma continuous_proj:
  | 
| 
 | 
   350  | 
  shows "continuous_on s (\<lambda>x. (x)\<^isub>F i)"
  | 
| 
 | 
   351  | 
  unfolding continuous_on_topological
  | 
| 
 | 
   352  | 
proof safe
  | 
| 
 | 
   353  | 
  fix x B assume "x \<in> s" "open B" "x i \<in> B"
  | 
| 
 | 
   354  | 
  let ?A = "Pi' (domain x) (\<lambda>j. if i = j then B else UNIV)"
  | 
| 
 | 
   355  | 
  have "open ?A" using `open B` by (auto intro: open_Pi'I)
  | 
| 
 | 
   356  | 
  moreover have "x \<in> ?A" using `x i \<in> B` by auto
  | 
| 
 | 
   357  | 
  moreover have "(\<forall>y\<in>s. y \<in> ?A \<longrightarrow> y i \<in> B)"
  | 
| 
 | 
   358  | 
  proof (cases, safe)
  | 
| 
 | 
   359  | 
    fix y assume "y \<in> s"
  | 
| 
 | 
   360  | 
    assume "i \<notin> domain x" hence "undefined \<in> B" using `x i \<in> B`
  | 
| 
 | 
   361  | 
      by simp
  | 
| 
 | 
   362  | 
    moreover
  | 
| 
 | 
   363  | 
    assume "y \<in> ?A" hence "domain y = domain x" by (simp add: Pi'_def)
  | 
| 
 | 
   364  | 
    hence "y i = undefined" using `i \<notin> domain x` by simp
  | 
| 
 | 
   365  | 
    ultimately
  | 
| 
 | 
   366  | 
    show "y i \<in> B" by simp
  | 
| 
 | 
   367  | 
  qed force
  | 
| 
 | 
   368  | 
  ultimately
  | 
| 
 | 
   369  | 
  show "\<exists>A. open A \<and> x \<in> A \<and> (\<forall>y\<in>s. y \<in> A \<longrightarrow> y i \<in> B)" by blast
  | 
| 
 | 
   370  | 
qed
  | 
| 
 | 
   371  | 
  | 
| 
 | 
   372  | 
subsection {* Complete Space of Finite Maps *}
 | 
| 
 | 
   373  | 
  | 
| 
 | 
   374  | 
lemma tendsto_dist_zero:
  | 
| 
 | 
   375  | 
  assumes "(\<lambda>i. dist (f i) g) ----> 0"
  | 
| 
 | 
   376  | 
  shows "f ----> g"
  | 
| 
 | 
   377  | 
  using assms by (auto simp: tendsto_iff dist_real_def)
  | 
| 
 | 
   378  | 
  | 
| 
 | 
   379  | 
lemma tendsto_dist_zero':
  | 
| 
 | 
   380  | 
  assumes "(\<lambda>i. dist (f i) g) ----> x"
  | 
| 
 | 
   381  | 
  assumes "0 = x"
  | 
| 
 | 
   382  | 
  shows "f ----> g"
  | 
| 
 | 
   383  | 
  using assms tendsto_dist_zero by simp
  | 
| 
 | 
   384  | 
  | 
| 
 | 
   385  | 
lemma tendsto_finmap:
  | 
| 
 | 
   386  | 
  fixes f::"nat \<Rightarrow> ('i \<Rightarrow>\<^isub>F ('a::metric_space))"
 | 
| 
 | 
   387  | 
  assumes ind_f:  "\<And>n. domain (f n) = domain g"
  | 
| 
 | 
   388  | 
  assumes proj_g:  "\<And>i. i \<in> domain g \<Longrightarrow> (\<lambda>n. (f n) i) ----> g i"
  | 
| 
 | 
   389  | 
  shows "f ----> g"
  | 
| 
 | 
   390  | 
  apply (rule tendsto_dist_zero')
  | 
| 
 | 
   391  | 
  unfolding dist_finmap_def assms
  | 
| 
 | 
   392  | 
  apply (rule tendsto_intros proj_g | simp)+
  | 
| 
 | 
   393  | 
  done
  | 
| 
 | 
   394  | 
  | 
| 
 | 
   395  | 
instance finmap :: (type, complete_space) complete_space
  | 
| 
 | 
   396  | 
proof
  | 
| 
 | 
   397  | 
  fix P::"nat \<Rightarrow> 'a \<Rightarrow>\<^isub>F 'b"
  | 
| 
 | 
   398  | 
  assume "Cauchy P"
  | 
| 
 | 
   399  | 
  then obtain Nd where Nd: "\<And>n. n \<ge> Nd \<Longrightarrow> dist (P n) (P Nd) < 1"
  | 
| 
 | 
   400  | 
    by (force simp: cauchy)
  | 
| 
 | 
   401  | 
  def d \<equiv> "domain (P Nd)"
  | 
| 
 | 
   402  | 
  with Nd have dim: "\<And>n. n \<ge> Nd \<Longrightarrow> domain (P n) = d" using dist_le_1_imp_domain_eq by auto
  | 
| 
 | 
   403  | 
  have [simp]: "finite d" unfolding d_def by simp
  | 
| 
 | 
   404  | 
  def p \<equiv> "\<lambda>i n. (P n) i"
  | 
| 
 | 
   405  | 
  def q \<equiv> "\<lambda>i. lim (p i)"
  | 
| 
 | 
   406  | 
  def Q \<equiv> "finmap_of d q"
  | 
| 
 | 
   407  | 
  have q: "\<And>i. i \<in> d \<Longrightarrow> q i = Q i" by (auto simp add: Q_def Abs_finmap_inverse)
  | 
| 
 | 
   408  | 
  {
 | 
| 
 | 
   409  | 
    fix i assume "i \<in> d"
  | 
| 
 | 
   410  | 
    have "Cauchy (p i)" unfolding cauchy p_def
  | 
| 
 | 
   411  | 
    proof safe
  | 
| 
 | 
   412  | 
      fix e::real assume "0 < e"
  | 
| 
 | 
   413  | 
      with `Cauchy P` obtain N where N: "\<And>n. n\<ge>N \<Longrightarrow> dist (P n) (P N) < min e 1"
  | 
| 
 | 
   414  | 
        by (force simp: cauchy min_def)
  | 
| 
 | 
   415  | 
      hence "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = domain (P N)" using dist_le_1_imp_domain_eq by auto
  | 
| 
 | 
   416  | 
      with dim have dim: "\<And>n. n \<ge> N \<Longrightarrow> domain (P n) = d" by (metis nat_le_linear)
  | 
| 
 | 
   417  | 
      show "\<exists>N. \<forall>n\<ge>N. dist ((P n) i) ((P N) i) < e"
  | 
| 
 | 
   418  | 
      proof (safe intro!: exI[where x="N"])
  | 
| 
 | 
   419  | 
        fix n assume "N \<le> n" have "N \<le> N" by simp
  | 
| 
 | 
   420  | 
        have "dist ((P n) i) ((P N) i) \<le> dist (P n) (P N)"
  | 
| 
 | 
   421  | 
          using dim[OF `N \<le> n`]  dim[OF `N \<le> N`] `i \<in> d`
  | 
| 
 | 
   422  | 
          by (auto intro!: dist_proj)
  | 
| 
 | 
   423  | 
        also have "\<dots> < e" using N[OF `N \<le> n`] by simp
  | 
| 
 | 
   424  | 
        finally show "dist ((P n) i) ((P N) i) < e" .
  | 
| 
 | 
   425  | 
      qed
  | 
| 
 | 
   426  | 
    qed
  | 
| 
 | 
   427  | 
    hence "convergent (p i)" by (metis Cauchy_convergent_iff)
  | 
| 
 | 
   428  | 
    hence "p i ----> q i" unfolding q_def convergent_def by (metis limI)
  | 
| 
 | 
   429  | 
  } note p = this
  | 
| 
 | 
   430  | 
  have "P ----> Q"
  | 
| 
 | 
   431  | 
  proof (rule metric_LIMSEQ_I)
  | 
| 
 | 
   432  | 
    fix e::real assume "0 < e"
  | 
| 
 | 
   433  | 
    def e' \<equiv> "min 1 (e / (card d + 1))"
  | 
| 
 | 
   434  | 
    hence "0 < e'" using `0 < e` by (auto simp: e'_def intro: divide_pos_pos)
  | 
| 
 | 
   435  | 
    have "\<exists>ni. \<forall>i\<in>d. \<forall>n\<ge>ni i. dist (p i n) (q i) < e'"
  | 
| 
 | 
   436  | 
    proof (safe intro!: bchoice)
  | 
| 
 | 
   437  | 
      fix i assume "i \<in> d"
  | 
| 
 | 
   438  | 
      from p[OF `i \<in> d`, THEN metric_LIMSEQ_D, OF `0 < e'`]
  | 
| 
 | 
   439  | 
      show "\<exists>no. \<forall>n\<ge>no. dist (p i n) (q i) < e'" .
  | 
| 
 | 
   440  | 
    qed then guess ni .. note ni = this
  | 
| 
 | 
   441  | 
    def N \<equiv> "max Nd (Max (ni ` d))"
  | 
| 
 | 
   442  | 
    show "\<exists>N. \<forall>n\<ge>N. dist (P n) Q < e"
  | 
| 
 | 
   443  | 
    proof (safe intro!: exI[where x="N"])
  | 
| 
 | 
   444  | 
      fix n assume "N \<le> n"
  | 
| 
 | 
   445  | 
      hence "domain (P n) = d" "domain Q = d" "domain (P n) = domain Q"
  | 
| 
 | 
   446  | 
        using dim by (simp_all add: N_def Q_def dim_def Abs_finmap_inverse)
  | 
| 
 | 
   447  | 
      hence "dist (P n) Q = (\<Sum>i\<in>d. dist ((P n) i) (Q i))" by (simp add: dist_finmap_def)
  | 
| 
 | 
   448  | 
      also have "\<dots> \<le> (\<Sum>i\<in>d. e')"
  | 
| 
 | 
   449  | 
      proof (intro setsum_mono less_imp_le)
  | 
| 
 | 
   450  | 
        fix i assume "i \<in> d"
  | 
| 
 | 
   451  | 
        hence "ni i \<le> Max (ni ` d)" by simp
  | 
| 
 | 
   452  | 
        also have "\<dots> \<le> N" by (simp add: N_def)
  | 
| 
 | 
   453  | 
        also have "\<dots> \<le> n" using `N \<le> n` .
  | 
| 
 | 
   454  | 
        finally
  | 
| 
 | 
   455  | 
        show "dist ((P n) i) (Q i) < e'"
  | 
| 
 | 
   456  | 
          using ni `i \<in> d` by (auto simp: p_def q N_def)
  | 
| 
 | 
   457  | 
      qed
  | 
| 
 | 
   458  | 
      also have "\<dots> = card d * e'" by (simp add: real_eq_of_nat)
  | 
| 
 | 
   459  | 
      also have "\<dots> < e" using `0 < e` by (simp add: e'_def field_simps min_def)
  | 
| 
 | 
   460  | 
      finally show "dist (P n) Q < e" .
  | 
| 
 | 
   461  | 
    qed
  | 
| 
 | 
   462  | 
  qed
  | 
| 
 | 
   463  | 
  thus "convergent P" by (auto simp: convergent_def)
  | 
| 
 | 
   464  | 
qed
  | 
| 
 | 
   465  | 
  | 
| 
 | 
   466  | 
subsection {* Polish Space of Finite Maps *}
 | 
| 
 | 
   467  | 
  | 
| 
 | 
   468  | 
instantiation finmap :: (countable, polish_space) polish_space
  | 
| 
 | 
   469  | 
begin
  | 
| 
 | 
   470  | 
  | 
| 
 | 
   471  | 
definition enum_basis_finmap :: "nat \<Rightarrow> ('a \<Rightarrow>\<^isub>F 'b) set" where
 | 
| 
 | 
   472  | 
  "enum_basis_finmap n =
  | 
| 
 | 
   473  | 
  (let m = from_nat n::('a \<Rightarrow>\<^isub>F nat) in Pi' (domain m) (enum_basis o (m)\<^isub>F))"
 | 
| 
 | 
   474  | 
  | 
| 
 | 
   475  | 
lemma range_enum_basis_eq:
  | 
| 
 | 
   476  | 
  "range enum_basis_finmap = {Pi' I S|I S. finite I \<and> (\<forall>i \<in> I. S i \<in> range enum_basis)}"
 | 
| 
 | 
   477  | 
proof (auto simp: enum_basis_finmap_def[abs_def])
  | 
| 
 | 
   478  | 
  fix S::"('a \<Rightarrow> 'b set)" and I
 | 
| 
 | 
   479  | 
  assume "\<forall>i\<in>I. S i \<in> range enum_basis"
  | 
| 
 | 
   480  | 
  hence "\<forall>i\<in>I. \<exists>n. S i = enum_basis n" by auto
  | 
| 
 | 
   481  | 
  then obtain n where n: "\<forall>i\<in>I. S i = enum_basis (n i)"
  | 
| 
 | 
   482  | 
    unfolding bchoice_iff by blast
  | 
| 
 | 
   483  | 
  assume [simp]: "finite I"
  | 
| 
 | 
   484  | 
  have "\<exists>fm. domain fm = I \<and> (\<forall>i\<in>I. n i = (fm i))"
  | 
| 
 | 
   485  | 
    by (rule finmap_choice) auto
  | 
| 
 | 
   486  | 
  then obtain m where "Pi' I S = Pi' (domain m) (enum_basis o m)"
  | 
| 
 | 
   487  | 
    using n by (auto simp: Pi'_def)
  | 
| 
 | 
   488  | 
  hence "Pi' I S = (let m = from_nat (to_nat m) in Pi' (domain m) (enum_basis \<circ> m))"
  | 
| 
 | 
   489  | 
    by simp
  | 
| 
 | 
   490  | 
  thus "Pi' I S \<in> range (\<lambda>n. let m = from_nat n in Pi' (domain m) (enum_basis \<circ> m))"
  | 
| 
 | 
   491  | 
    by blast
  | 
| 
 | 
   492  | 
qed (metis finite_domain o_apply rangeI)
  | 
| 
 | 
   493  | 
  | 
| 
 | 
   494  | 
lemma in_enum_basis_finmapI:
  | 
| 
 | 
   495  | 
  assumes "finite I" assumes "\<And>i. i \<in> I \<Longrightarrow> S i \<in> range enum_basis"
  | 
| 
 | 
   496  | 
  shows "Pi' I S \<in> range enum_basis_finmap"
  | 
| 
 | 
   497  | 
  using assms unfolding range_enum_basis_eq by auto
  | 
| 
 | 
   498  | 
  | 
| 
 | 
   499  | 
lemma finmap_topological_basis:
  | 
| 
 | 
   500  | 
  "topological_basis (range (enum_basis_finmap))"
  | 
| 
 | 
   501  | 
proof (subst topological_basis_iff, safe)
  | 
| 
 | 
   502  | 
  fix n::nat
  | 
| 
 | 
   503  | 
  show "open (enum_basis_finmap n::('a \<Rightarrow>\<^isub>F 'b) set)" using enumerable_basis
 | 
| 
 | 
   504  | 
    by (auto intro!: open_Pi'I simp: topological_basis_def enum_basis_finmap_def Let_def)
  | 
| 
 | 
   505  | 
next
  | 
| 
 | 
   506  | 
  fix O'::"('a \<Rightarrow>\<^isub>F 'b) set" and x
 | 
| 
 | 
   507  | 
  assume "open O'" "x \<in> O'"
  | 
| 
 | 
   508  | 
  then obtain e where e: "e > 0" "\<And>y. dist y x < e \<Longrightarrow> y \<in> O'"  unfolding open_dist by blast
  | 
| 
 | 
   509  | 
  def e' \<equiv> "e / (card (domain x) + 1)"
  | 
| 
 | 
   510  | 
  | 
| 
 | 
   511  | 
  have "\<exists>B.
  | 
| 
 | 
   512  | 
    (\<forall>i\<in>domain x. x i \<in> enum_basis (B i) \<and> enum_basis (B i) \<subseteq> ball (x i) e')"
  | 
| 
 | 
   513  | 
  proof (rule bchoice, safe)
  | 
| 
 | 
   514  | 
    fix i assume "i \<in> domain x"
  | 
| 
 | 
   515  | 
    have "open (ball (x i) e')" "x i \<in> ball (x i) e'" using e
  | 
| 
 | 
   516  | 
      by (auto simp add: e'_def intro!: divide_pos_pos)
  | 
| 
 | 
   517  | 
    from enumerable_basisE[OF this] guess b' .
  | 
| 
 | 
   518  | 
    thus "\<exists>y. x i \<in> enum_basis y \<and>
  | 
| 
 | 
   519  | 
            enum_basis y \<subseteq> ball (x i) e'" by auto
  | 
| 
 | 
   520  | 
  qed
  | 
| 
 | 
   521  | 
  then guess B .. note B = this
  | 
| 
 | 
   522  | 
  def B' \<equiv> "Pi' (domain x) (\<lambda>i. enum_basis (B i)::'b set)"
  | 
| 
 | 
   523  | 
  hence "B' \<in> range enum_basis_finmap" unfolding B'_def
  | 
| 
 | 
   524  | 
    by (intro in_enum_basis_finmapI) auto
  | 
| 
 | 
   525  | 
  moreover have "x \<in> B'" unfolding B'_def using B by auto
  | 
| 
 | 
   526  | 
  moreover have "B' \<subseteq> O'"
  | 
| 
 | 
   527  | 
  proof
  | 
| 
 | 
   528  | 
    fix y assume "y \<in> B'" with B have "domain y = domain x" unfolding B'_def
  | 
| 
 | 
   529  | 
      by (simp add: Pi'_def)
  | 
| 
 | 
   530  | 
    show "y \<in> O'"
  | 
| 
 | 
   531  | 
    proof (rule e)
  | 
| 
 | 
   532  | 
      have "dist y x = (\<Sum>i \<in> domain x. dist (y i) (x i))"
  | 
| 
 | 
   533  | 
        using `domain y = domain x` by (simp add: dist_finmap_def)
  | 
| 
 | 
   534  | 
      also have "\<dots> \<le> (\<Sum>i \<in> domain x. e')"
  | 
| 
 | 
   535  | 
      proof (rule setsum_mono)
  | 
| 
 | 
   536  | 
        fix i assume "i \<in> domain x"
  | 
| 
 | 
   537  | 
        with `y \<in> B'` B have "y i \<in> enum_basis (B i)"
  | 
| 
 | 
   538  | 
          by (simp add: Pi'_def B'_def)
  | 
| 
 | 
   539  | 
        hence "y i \<in> ball (x i) e'" using B `domain y = domain x` `i \<in> domain x`
  | 
| 
 | 
   540  | 
          by force
  | 
| 
 | 
   541  | 
        thus "dist (y i) (x i) \<le> e'" by (simp add: dist_commute)
  | 
| 
 | 
   542  | 
      qed
  | 
| 
 | 
   543  | 
      also have "\<dots> = card (domain x) * e'" by (simp add: real_eq_of_nat)
  | 
| 
 | 
   544  | 
      also have "\<dots> < e" using e by (simp add: e'_def field_simps)
  | 
| 
 | 
   545  | 
      finally show "dist y x < e" .
  | 
| 
 | 
   546  | 
    qed
  | 
| 
 | 
   547  | 
  qed
  | 
| 
 | 
   548  | 
  ultimately
  | 
| 
 | 
   549  | 
  show "\<exists>B'\<in>range enum_basis_finmap. x \<in> B' \<and> B' \<subseteq> O'" by blast
  | 
| 
 | 
   550  | 
qed
  | 
| 
 | 
   551  | 
  | 
| 
 | 
   552  | 
lemma range_enum_basis_finmap_imp_open:
  | 
| 
 | 
   553  | 
  assumes "x \<in> range enum_basis_finmap"
  | 
| 
 | 
   554  | 
  shows "open x"
  | 
| 
 | 
   555  | 
  using finmap_topological_basis assms by (auto simp: topological_basis_def)
  | 
| 
 | 
   556  | 
  | 
| 
 | 
   557  | 
lemma
  | 
| 
 | 
   558  | 
  open_imp_ex_UNION_of_enum:
  | 
| 
 | 
   559  | 
  fixes X::"('a \<Rightarrow>\<^isub>F 'b) set"
 | 
| 
 | 
   560  | 
  assumes "open X" assumes "X \<noteq> {}"
 | 
| 
 | 
   561  | 
  shows "\<exists>A::nat\<Rightarrow>'a set. \<exists>B::nat\<Rightarrow>('a \<Rightarrow> 'b set) . X = UNION UNIV (\<lambda>i. Pi' (A i) (B i)) \<and>
 | 
| 
 | 
   562  | 
    (\<forall>n. \<forall>i\<in>A n. (B n) i \<in> range enum_basis) \<and> (\<forall>n. finite (A n))"
  | 
| 
 | 
   563  | 
proof -
  | 
| 
 | 
   564  | 
  from `open X` obtain B' where B': "B'\<subseteq>range enum_basis_finmap" "\<Union>B' = X"
  | 
| 
 | 
   565  | 
    using finmap_topological_basis by (force simp add: topological_basis_def)
  | 
| 
 | 
   566  | 
  then obtain B where B: "B' = enum_basis_finmap ` B" by (auto simp: subset_image_iff)
  | 
| 
 | 
   567  | 
  show ?thesis
  | 
| 
 | 
   568  | 
  proof cases
  | 
| 
 | 
   569  | 
    assume "B = {}" with B have "B' = {}" by simp hence False using B' assms by simp
 | 
| 
 | 
   570  | 
    thus ?thesis by simp
  | 
| 
 | 
   571  | 
  next
  | 
| 
 | 
   572  | 
    assume "B \<noteq> {}" then obtain b where b: "b \<in> B" by auto
 | 
| 
 | 
   573  | 
    def NA \<equiv> "\<lambda>n::nat. if n \<in> B
  | 
| 
 | 
   574  | 
      then domain ((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) n)
  | 
| 
 | 
   575  | 
      else domain ((from_nat::_\<Rightarrow>'a\<Rightarrow>\<^isub>F nat) b)"
  | 
| 
 | 
   576  | 
    def NB \<equiv> "\<lambda>n::nat. if n \<in> B
  | 
| 
 | 
   577  | 
      then (\<lambda>i. (enum_basis::nat\<Rightarrow>'b set) (((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) n) i))
  | 
| 
 | 
   578  | 
      else (\<lambda>i. (enum_basis::nat\<Rightarrow>'b set) (((from_nat::_\<Rightarrow>'a \<Rightarrow>\<^isub>F nat) b) i))"
  | 
| 
 | 
   579  | 
    have "X = UNION UNIV (\<lambda>i. Pi' (NA i) (NB i))" unfolding B'(2)[symmetric] using b
  | 
| 
 | 
   580  | 
      unfolding B
  | 
| 
 | 
   581  | 
      by safe
  | 
| 
 | 
   582  | 
         (auto simp add: NA_def NB_def enum_basis_finmap_def Let_def o_def split: split_if_asm)
  | 
| 
 | 
   583  | 
    moreover
  | 
| 
 | 
   584  | 
    have "(\<forall>n. \<forall>i\<in>NA n. (NB n) i \<in> range enum_basis)"
  | 
| 
 | 
   585  | 
      using enumerable_basis by (auto simp: topological_basis_def NA_def NB_def)
  | 
| 
 | 
   586  | 
    moreover have "(\<forall>n. finite (NA n))" by (simp add: NA_def)
  | 
| 
 | 
   587  | 
    ultimately show ?thesis by auto
  | 
| 
 | 
   588  | 
  qed
  | 
| 
 | 
   589  | 
qed
  | 
| 
 | 
   590  | 
  | 
| 
 | 
   591  | 
lemma
  | 
| 
 | 
   592  | 
  open_imp_ex_UNION:
  | 
| 
 | 
   593  | 
  fixes X::"('a \<Rightarrow>\<^isub>F 'b) set"
 | 
| 
 | 
   594  | 
  assumes "open X" assumes "X \<noteq> {}"
 | 
| 
 | 
   595  | 
  shows "\<exists>A::nat\<Rightarrow>'a set. \<exists>B::nat\<Rightarrow>('a \<Rightarrow> 'b set) . X = UNION UNIV (\<lambda>i. Pi' (A i) (B i)) \<and>
 | 
| 
 | 
   596  | 
    (\<forall>n. \<forall>i\<in>A n. open ((B n) i)) \<and> (\<forall>n. finite (A n))"
  | 
| 
 | 
   597  | 
  using open_imp_ex_UNION_of_enum[OF assms]
  | 
| 
 | 
   598  | 
  apply auto
  | 
| 
 | 
   599  | 
  apply (rule_tac x = A in exI)
  | 
| 
 | 
   600  | 
  apply (rule_tac x = B in exI)
  | 
| 
 | 
   601  | 
  apply (auto simp: open_enum_basis)
  | 
| 
 | 
   602  | 
  done
  | 
| 
 | 
   603  | 
  | 
| 
 | 
   604  | 
lemma
  | 
| 
 | 
   605  | 
  open_basisE:
  | 
| 
 | 
   606  | 
  assumes "open X" assumes "X \<noteq> {}"
 | 
| 
 | 
   607  | 
  obtains A::"nat\<Rightarrow>'a set" and B::"nat\<Rightarrow>('a \<Rightarrow> 'b set)" where
 | 
| 
 | 
   608  | 
  "X = UNION UNIV (\<lambda>i. Pi' (A i) (B i))" "\<And>n i. i\<in>A n \<Longrightarrow> open ((B n) i)" "\<And>n. finite (A n)"
  | 
| 
 | 
   609  | 
using open_imp_ex_UNION[OF assms] by auto
  | 
| 
 | 
   610  | 
  | 
| 
 | 
   611  | 
lemma
  | 
| 
 | 
   612  | 
  open_basis_of_enumE:
  | 
| 
 | 
   613  | 
  assumes "open X" assumes "X \<noteq> {}"
 | 
| 
 | 
   614  | 
  obtains A::"nat\<Rightarrow>'a set" and B::"nat\<Rightarrow>('a \<Rightarrow> 'b set)" where
 | 
| 
 | 
   615  | 
  "X = UNION UNIV (\<lambda>i. Pi' (A i) (B i))" "\<And>n i. i\<in>A n \<Longrightarrow> (B n) i \<in> range enum_basis"
  | 
| 
 | 
   616  | 
  "\<And>n. finite (A n)"
  | 
| 
 | 
   617  | 
using open_imp_ex_UNION_of_enum[OF assms] by auto
  | 
| 
 | 
   618  | 
  | 
| 
 | 
   619  | 
instance proof qed (blast intro: finmap_topological_basis)
  | 
| 
 | 
   620  | 
  | 
| 
 | 
   621  | 
end
  | 
| 
 | 
   622  | 
  | 
| 
 | 
   623  | 
subsection {* Product Measurable Space of Finite Maps *}
 | 
| 
 | 
   624  | 
  | 
| 
 | 
   625  | 
definition "PiF I M \<equiv>
  | 
| 
 | 
   626  | 
  sigma
  | 
| 
 | 
   627  | 
    (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))
  | 
| 
 | 
   628  | 
    {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
 | 
| 
 | 
   629  | 
  | 
| 
 | 
   630  | 
abbreviation
  | 
| 
 | 
   631  | 
  "Pi\<^isub>F I M \<equiv> PiF I M"
  | 
| 
 | 
   632  | 
  | 
| 
 | 
   633  | 
syntax
  | 
| 
 | 
   634  | 
  "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3PIF _:_./ _)" 10)
 | 
| 
 | 
   635  | 
  | 
| 
 | 
   636  | 
syntax (xsymbols)
  | 
| 
 | 
   637  | 
  "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>F _\<in>_./ _)"  10)
 | 
| 
 | 
   638  | 
  | 
| 
 | 
   639  | 
syntax (HTML output)
  | 
| 
 | 
   640  | 
  "_PiF" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>F _\<in>_./ _)"  10)
 | 
| 
 | 
   641  | 
  | 
| 
 | 
   642  | 
translations
  | 
| 
 | 
   643  | 
  "PIF x:I. M" == "CONST PiF I (%x. M)"
  | 
| 
 | 
   644  | 
  | 
| 
 | 
   645  | 
lemma PiF_gen_subset: "{(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))} \<subseteq>
 | 
| 
 | 
   646  | 
    Pow (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
  | 
| 
 | 
   647  | 
  by (auto simp: Pi'_def) (blast dest: sets_into_space)
  | 
| 
 | 
   648  | 
  | 
| 
 | 
   649  | 
lemma space_PiF: "space (PiF I M) = (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))"
  | 
| 
 | 
   650  | 
  unfolding PiF_def using PiF_gen_subset by (rule space_measure_of)
  | 
| 
 | 
   651  | 
  | 
| 
 | 
   652  | 
lemma sets_PiF:
  | 
| 
 | 
   653  | 
  "sets (PiF I M) = sigma_sets (\<Union>J \<in> I. (\<Pi>' j\<in>J. space (M j)))
  | 
| 
 | 
   654  | 
    {(\<Pi>' j\<in>J. X j) |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
 | 
| 
 | 
   655  | 
  unfolding PiF_def using PiF_gen_subset by (rule sets_measure_of)
  | 
| 
 | 
   656  | 
  | 
| 
 | 
   657  | 
lemma sets_PiF_singleton:
  | 
| 
 | 
   658  | 
  "sets (PiF {I} M) = sigma_sets (\<Pi>' j\<in>I. space (M j))
 | 
| 
 | 
   659  | 
    {(\<Pi>' j\<in>I. X j) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
 | 
| 
 | 
   660  | 
  unfolding sets_PiF by simp
  | 
| 
 | 
   661  | 
  | 
| 
 | 
   662  | 
lemma in_sets_PiFI:
  | 
| 
 | 
   663  | 
  assumes "X = (Pi' J S)" "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
  | 
| 
 | 
   664  | 
  shows "X \<in> sets (PiF I M)"
  | 
| 
 | 
   665  | 
  unfolding sets_PiF
  | 
| 
 | 
   666  | 
  using assms by blast
  | 
| 
 | 
   667  | 
  | 
| 
 | 
   668  | 
lemma product_in_sets_PiFI:
  | 
| 
 | 
   669  | 
  assumes "J \<in> I" "\<And>i. i\<in>J \<Longrightarrow> S i \<in> sets (M i)"
  | 
| 
 | 
   670  | 
  shows "(Pi' J S) \<in> sets (PiF I M)"
  | 
| 
 | 
   671  | 
  unfolding sets_PiF
  | 
| 
 | 
   672  | 
  using assms by blast
  | 
| 
 | 
   673  | 
  | 
| 
 | 
   674  | 
lemma singleton_space_subset_in_sets:
  | 
| 
 | 
   675  | 
  fixes J
  | 
| 
 | 
   676  | 
  assumes "J \<in> I"
  | 
| 
 | 
   677  | 
  assumes "finite J"
  | 
| 
 | 
   678  | 
  shows "space (PiF {J} M) \<in> sets (PiF I M)"
 | 
| 
 | 
   679  | 
  using assms
  | 
| 
 | 
   680  | 
  by (intro in_sets_PiFI[where J=J and S="\<lambda>i. space (M i)"])
  | 
| 
 | 
   681  | 
      (auto simp: product_def space_PiF)
  | 
| 
 | 
   682  | 
  | 
| 
 | 
   683  | 
lemma singleton_subspace_set_in_sets:
  | 
| 
 | 
   684  | 
  assumes A: "A \<in> sets (PiF {J} M)"
 | 
| 
 | 
   685  | 
  assumes "finite J"
  | 
| 
 | 
   686  | 
  assumes "J \<in> I"
  | 
| 
 | 
   687  | 
  shows "A \<in> sets (PiF I M)"
  | 
| 
 | 
   688  | 
  using A[unfolded sets_PiF]
  | 
| 
 | 
   689  | 
  apply (induct A)
  | 
| 
 | 
   690  | 
  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
  | 
| 
 | 
   691  | 
  using assms
  | 
| 
 | 
   692  | 
  by (auto intro: in_sets_PiFI intro!: singleton_space_subset_in_sets)
  | 
| 
 | 
   693  | 
  | 
| 
 | 
   694  | 
lemma
  | 
| 
 | 
   695  | 
  finite_measurable_singletonI:
  | 
| 
 | 
   696  | 
  assumes "finite I"
  | 
| 
 | 
   697  | 
  assumes "\<And>J. J \<in> I \<Longrightarrow> finite J"
  | 
| 
 | 
   698  | 
  assumes MN: "\<And>J. J \<in> I \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
 | 
| 
 | 
   699  | 
  shows "A \<in> measurable (PiF I M) N"
  | 
| 
 | 
   700  | 
  unfolding measurable_def
  | 
| 
 | 
   701  | 
proof safe
  | 
| 
 | 
   702  | 
  fix y assume "y \<in> sets N"
  | 
| 
 | 
   703  | 
  have "A -` y \<inter> space (PiF I M) = (\<Union>J\<in>I. A -` y \<inter> space (PiF {J} M))"
 | 
| 
 | 
   704  | 
    by (auto simp: space_PiF)
  | 
| 
 | 
   705  | 
  also have "\<dots> \<in> sets (PiF I M)"
  | 
| 
 | 
   706  | 
  proof
  | 
| 
 | 
   707  | 
    show "finite I" by fact
  | 
| 
 | 
   708  | 
    fix J assume "J \<in> I"
  | 
| 
 | 
   709  | 
    with assms have "finite J" by simp
  | 
| 
 | 
   710  | 
    show "A -` y \<inter> space (PiF {J} M) \<in> sets (PiF I M)"
 | 
| 
 | 
   711  | 
      by (rule singleton_subspace_set_in_sets[OF measurable_sets[OF assms(3)]]) fact+
  | 
| 
 | 
   712  | 
  qed
  | 
| 
 | 
   713  | 
  finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
  | 
| 
 | 
   714  | 
next
  | 
| 
 | 
   715  | 
  fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
  | 
| 
 | 
   716  | 
    using MN[of "domain x"]
  | 
| 
 | 
   717  | 
    by (auto simp: space_PiF measurable_space Pi'_def)
  | 
| 
 | 
   718  | 
qed
  | 
| 
 | 
   719  | 
  | 
| 
 | 
   720  | 
lemma
  | 
| 
 | 
   721  | 
  countable_finite_comprehension:
  | 
| 
 | 
   722  | 
  fixes f :: "'a::countable set \<Rightarrow> _"
  | 
| 
 | 
   723  | 
  assumes "\<And>s. P s \<Longrightarrow> finite s"
  | 
| 
 | 
   724  | 
  assumes "\<And>s. P s \<Longrightarrow> f s \<in> sets M"
  | 
| 
 | 
   725  | 
  shows "\<Union>{f s|s. P s} \<in> sets M"
 | 
| 
 | 
   726  | 
proof -
  | 
| 
 | 
   727  | 
  have "\<Union>{f s|s. P s} = (\<Union>n::nat. let s = set (from_nat n) in if P s then f s else {})"
 | 
| 
 | 
   728  | 
  proof safe
  | 
| 
 | 
   729  | 
    fix x X s assume "x \<in> f s" "P s"
  | 
| 
 | 
   730  | 
    moreover with assms obtain l where "s = set l" using finite_list by blast
  | 
| 
 | 
   731  | 
    ultimately show "x \<in> (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" using `P s`
 | 
| 
 | 
   732  | 
      by (auto intro!: exI[where x="to_nat l"])
  | 
| 
 | 
   733  | 
  next
  | 
| 
 | 
   734  | 
    fix x n assume "x \<in> (let s = set (from_nat n) in if P s then f s else {})"
 | 
| 
 | 
   735  | 
    thus "x \<in> \<Union>{f s|s. P s}" using assms by (auto simp: Let_def split: split_if_asm)
 | 
| 
 | 
   736  | 
  qed
  | 
| 
 | 
   737  | 
  hence "\<Union>{f s|s. P s} = (\<Union>n. let s = set (from_nat n) in if P s then f s else {})" by simp
 | 
| 
 | 
   738  | 
  also have "\<dots> \<in> sets M" using assms by (auto simp: Let_def)
  | 
| 
 | 
   739  | 
  finally show ?thesis .
  | 
| 
 | 
   740  | 
qed
  | 
| 
 | 
   741  | 
  | 
| 
 | 
   742  | 
lemma space_subset_in_sets:
  | 
| 
 | 
   743  | 
  fixes J::"'a::countable set set"
  | 
| 
 | 
   744  | 
  assumes "J \<subseteq> I"
  | 
| 
 | 
   745  | 
  assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
  | 
| 
 | 
   746  | 
  shows "space (PiF J M) \<in> sets (PiF I M)"
  | 
| 
 | 
   747  | 
proof -
  | 
| 
 | 
   748  | 
  have "space (PiF J M) = \<Union>{space (PiF {j} M)|j. j \<in> J}"
 | 
| 
 | 
   749  | 
    unfolding space_PiF by blast
  | 
| 
 | 
   750  | 
  also have "\<dots> \<in> sets (PiF I M)" using assms
  | 
| 
 | 
   751  | 
    by (intro countable_finite_comprehension) (auto simp: singleton_space_subset_in_sets)
  | 
| 
 | 
   752  | 
  finally show ?thesis .
  | 
| 
 | 
   753  | 
qed
  | 
| 
 | 
   754  | 
  | 
| 
 | 
   755  | 
lemma subspace_set_in_sets:
  | 
| 
 | 
   756  | 
  fixes J::"'a::countable set set"
  | 
| 
 | 
   757  | 
  assumes A: "A \<in> sets (PiF J M)"
  | 
| 
 | 
   758  | 
  assumes "J \<subseteq> I"
  | 
| 
 | 
   759  | 
  assumes "\<And>j. j \<in> J \<Longrightarrow> finite j"
  | 
| 
 | 
   760  | 
  shows "A \<in> sets (PiF I M)"
  | 
| 
 | 
   761  | 
  using A[unfolded sets_PiF]
  | 
| 
 | 
   762  | 
  apply (induct A)
  | 
| 
 | 
   763  | 
  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
  | 
| 
 | 
   764  | 
  using assms
  | 
| 
 | 
   765  | 
  by (auto intro: in_sets_PiFI intro!: space_subset_in_sets)
  | 
| 
 | 
   766  | 
  | 
| 
 | 
   767  | 
lemma
  | 
| 
 | 
   768  | 
  countable_measurable_PiFI:
  | 
| 
 | 
   769  | 
  fixes I::"'a::countable set set"
  | 
| 
 | 
   770  | 
  assumes MN: "\<And>J. J \<in> I \<Longrightarrow> finite J \<Longrightarrow> A \<in> measurable (PiF {J} M) N"
 | 
| 
 | 
   771  | 
  shows "A \<in> measurable (PiF I M) N"
  | 
| 
 | 
   772  | 
  unfolding measurable_def
  | 
| 
 | 
   773  | 
proof safe
  | 
| 
 | 
   774  | 
  fix y assume "y \<in> sets N"
  | 
| 
 | 
   775  | 
  have "A -` y = (\<Union>{A -` y \<inter> {x. domain x = J}|J. finite J})" by auto
 | 
| 
 | 
   776  | 
  hence "A -` y \<inter> space (PiF I M) = (\<Union>n. A -` y \<inter> space (PiF ({set (from_nat n)}\<inter>I) M))"
 | 
| 
 | 
   777  | 
    apply (auto simp: space_PiF Pi'_def)
  | 
| 
 | 
   778  | 
  proof -
  | 
| 
 | 
   779  | 
    case goal1
  | 
| 
 | 
   780  | 
    from finite_list[of "domain x"] obtain xs where "set xs = domain x" by auto
  | 
| 
 | 
   781  | 
    thus ?case
  | 
| 
 | 
   782  | 
      apply (intro exI[where x="to_nat xs"])
  | 
| 
 | 
   783  | 
      apply auto
  | 
| 
 | 
   784  | 
      done
  | 
| 
 | 
   785  | 
  qed
  | 
| 
 | 
   786  | 
  also have "\<dots> \<in> sets (PiF I M)"
  | 
| 
 | 
   787  | 
    apply (intro Int countable_nat_UN subsetI, safe)
  | 
| 
 | 
   788  | 
    apply (case_tac "set (from_nat i) \<in> I")
  | 
| 
 | 
   789  | 
    apply simp_all
  | 
| 
 | 
   790  | 
    apply (rule singleton_subspace_set_in_sets[OF measurable_sets[OF MN]])
  | 
| 
 | 
   791  | 
    using assms `y \<in> sets N`
  | 
| 
 | 
   792  | 
    apply (auto simp: space_PiF)
  | 
| 
 | 
   793  | 
    done
  | 
| 
 | 
   794  | 
  finally show "A -` y \<inter> space (PiF I M) \<in> sets (PiF I M)" .
  | 
| 
 | 
   795  | 
next
  | 
| 
 | 
   796  | 
  fix x assume "x \<in> space (PiF I M)" thus "A x \<in> space N"
  | 
| 
 | 
   797  | 
    using MN[of "domain x"] by (auto simp: space_PiF measurable_space Pi'_def)
  | 
| 
 | 
   798  | 
qed
  | 
| 
 | 
   799  | 
  | 
| 
 | 
   800  | 
lemma measurable_PiF:
  | 
| 
 | 
   801  | 
  assumes f: "\<And>x. x \<in> space N \<Longrightarrow> domain (f x) \<in> I \<and> (\<forall>i\<in>domain (f x). (f x) i \<in> space (M i))"
  | 
| 
 | 
   802  | 
  assumes S: "\<And>J S. J \<in> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> S i \<in> sets (M i)) \<Longrightarrow>
  | 
| 
 | 
   803  | 
    f -` (Pi' J S) \<inter> space N \<in> sets N"
  | 
| 
 | 
   804  | 
  shows "f \<in> measurable N (PiF I M)"
  | 
| 
 | 
   805  | 
  unfolding PiF_def
  | 
| 
 | 
   806  | 
  using PiF_gen_subset
  | 
| 
 | 
   807  | 
  apply (rule measurable_measure_of)
  | 
| 
 | 
   808  | 
  using f apply force
  | 
| 
 | 
   809  | 
  apply (insert S, auto)
  | 
| 
 | 
   810  | 
  done
  | 
| 
 | 
   811  | 
  | 
| 
 | 
   812  | 
lemma
  | 
| 
 | 
   813  | 
  restrict_sets_measurable:
  | 
| 
 | 
   814  | 
  assumes A: "A \<in> sets (PiF I M)" and "J \<subseteq> I"
  | 
| 
 | 
   815  | 
  shows "A \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
 | 
| 
 | 
   816  | 
  using A[unfolded sets_PiF]
  | 
| 
 | 
   817  | 
  apply (induct A)
  | 
| 
 | 
   818  | 
  unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric]
  | 
| 
 | 
   819  | 
proof -
  | 
| 
 | 
   820  | 
  fix a assume "a \<in> {Pi' J X |X J. J \<in> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
 | 
| 
 | 
   821  | 
  then obtain K S where S: "a = Pi' K S" "K \<in> I" "(\<forall>i\<in>K. S i \<in> sets (M i))"
  | 
| 
 | 
   822  | 
    by auto
  | 
| 
 | 
   823  | 
  show "a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
 | 
| 
 | 
   824  | 
  proof cases
  | 
| 
 | 
   825  | 
    assume "K \<in> J"
  | 
| 
 | 
   826  | 
    hence "a \<inter> {m. domain m \<in> J} \<in> {Pi' K X |X K. K \<in> J \<and> X \<in> (\<Pi> j\<in>K. sets (M j))}" using S
 | 
| 
 | 
   827  | 
      by (auto intro!: exI[where x=K] exI[where x=S] simp: Pi'_def)
  | 
| 
 | 
   828  | 
    also have "\<dots> \<subseteq> sets (PiF J M)" unfolding sets_PiF by auto
  | 
| 
 | 
   829  | 
    finally show ?thesis .
  | 
| 
 | 
   830  | 
  next
  | 
| 
 | 
   831  | 
    assume "K \<notin> J"
  | 
| 
 | 
   832  | 
    hence "a \<inter> {m. domain m \<in> J} = {}" using S by (auto simp: Pi'_def)
 | 
| 
 | 
   833  | 
    also have "\<dots> \<in> sets (PiF J M)" by simp
  | 
| 
 | 
   834  | 
    finally show ?thesis .
  | 
| 
 | 
   835  | 
  qed
  | 
| 
 | 
   836  | 
next
  | 
| 
 | 
   837  | 
  show "{} \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" by simp
 | 
| 
 | 
   838  | 
next
  | 
| 
 | 
   839  | 
  fix a :: "nat \<Rightarrow> _"
  | 
| 
 | 
   840  | 
  assume a: "(\<And>i. a i \<inter> {m. domain m \<in> J} \<in> sets (PiF J M))"
 | 
| 
 | 
   841  | 
  have "UNION UNIV a \<inter> {m. domain m \<in> J} = (\<Union>i. (a i \<inter> {m. domain m \<in> J}))"
 | 
| 
 | 
   842  | 
    by simp
  | 
| 
 | 
   843  | 
  also have "\<dots> \<in> sets (PiF J M)" using a by (intro countable_nat_UN) auto
  | 
| 
 | 
   844  | 
  finally show "UNION UNIV a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" .
 | 
| 
 | 
   845  | 
next
  | 
| 
 | 
   846  | 
  fix a assume a: "a \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)"
 | 
| 
 | 
   847  | 
  have "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} = (space (PiF J M) - (a \<inter> {m. domain m \<in> J}))"
 | 
| 
 | 
   848  | 
    using `J \<subseteq> I` by (auto simp: space_PiF Pi'_def)
  | 
| 
 | 
   849  | 
  also have "\<dots> \<in> sets (PiF J M)" using a by auto
  | 
| 
 | 
   850  | 
  finally show "(space (PiF I M) - a) \<inter> {m. domain m \<in> J} \<in> sets (PiF J M)" .
 | 
| 
 | 
   851  | 
qed
  | 
| 
 | 
   852  | 
  | 
| 
 | 
   853  | 
lemma measurable_finmap_of:
  | 
| 
 | 
   854  | 
  assumes f: "\<And>i. (\<exists>x \<in> space N. i \<in> J x) \<Longrightarrow> (\<lambda>x. f x i) \<in> measurable N (M i)"
  | 
| 
 | 
   855  | 
  assumes J: "\<And>x. x \<in> space N \<Longrightarrow> J x \<in> I" "\<And>x. x \<in> space N \<Longrightarrow> finite (J x)"
  | 
| 
 | 
   856  | 
  assumes JN: "\<And>S. {x. J x = S} \<inter> space N \<in> sets N"
 | 
| 
 | 
   857  | 
  shows "(\<lambda>x. finmap_of (J x) (f x)) \<in> measurable N (PiF I M)"
  | 
| 
 | 
   858  | 
proof (rule measurable_PiF)
  | 
| 
 | 
   859  | 
  fix x assume "x \<in> space N"
  | 
| 
 | 
   860  | 
  with J[of x] measurable_space[OF f]
  | 
| 
 | 
   861  | 
  show "domain (finmap_of (J x) (f x)) \<in> I \<and>
  | 
| 
 | 
   862  | 
        (\<forall>i\<in>domain (finmap_of (J x) (f x)). (finmap_of (J x) (f x)) i \<in> space (M i))"
  | 
| 
 | 
   863  | 
    by auto
  | 
| 
 | 
   864  | 
next
  | 
| 
 | 
   865  | 
  fix K S assume "K \<in> I" and *: "\<And>i. i \<in> K \<Longrightarrow> S i \<in> sets (M i)"
  | 
| 
 | 
   866  | 
  with J have eq: "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N =
  | 
| 
 | 
   867  | 
    (if \<exists>x \<in> space N. K = J x \<and> finite K then if K = {} then {x \<in> space N. J x = K}
 | 
| 
 | 
   868  | 
      else (\<Inter>i\<in>K. (\<lambda>x. f x i) -` S i \<inter> {x \<in> space N. J x = K}) else {})"
 | 
| 
 | 
   869  | 
    by (auto simp: Pi'_def)
  | 
| 
 | 
   870  | 
  have r: "{x \<in> space N. J x = K} = space N \<inter> ({x. J x = K} \<inter> space N)" by auto
 | 
| 
 | 
   871  | 
  show "(\<lambda>x. finmap_of (J x) (f x)) -` Pi' K S \<inter> space N \<in> sets N"
  | 
| 
 | 
   872  | 
    unfolding eq r
  | 
| 
 | 
   873  | 
    apply (simp del: INT_simps add: )
  | 
| 
 | 
   874  | 
    apply (intro conjI impI finite_INT JN Int[OF top])
  | 
| 
 | 
   875  | 
    apply simp apply assumption
  | 
| 
 | 
   876  | 
    apply (subst Int_assoc[symmetric])
  | 
| 
 | 
   877  | 
    apply (rule Int)
  | 
| 
 | 
   878  | 
    apply (intro measurable_sets[OF f] *) apply force apply assumption
  | 
| 
 | 
   879  | 
    apply (intro JN)
  | 
| 
 | 
   880  | 
    done
  | 
| 
 | 
   881  | 
qed
  | 
| 
 | 
   882  | 
  | 
| 
 | 
   883  | 
lemma measurable_PiM_finmap_of:
  | 
| 
 | 
   884  | 
  assumes "finite J"
  | 
| 
 | 
   885  | 
  shows "finmap_of J \<in> measurable (Pi\<^isub>M J M) (PiF {J} M)"
 | 
| 
 | 
   886  | 
  apply (rule measurable_finmap_of)
  | 
| 
 | 
   887  | 
  apply (rule measurable_component_singleton)
  | 
| 
 | 
   888  | 
  apply simp
  | 
| 
 | 
   889  | 
  apply rule
  | 
| 
 | 
   890  | 
  apply (rule `finite J`)
  | 
| 
 | 
   891  | 
  apply simp
  | 
| 
 | 
   892  | 
  done
  | 
| 
 | 
   893  | 
  | 
| 
 | 
   894  | 
lemma proj_measurable_singleton:
  | 
| 
 | 
   895  | 
  assumes "A \<in> sets (M i)" "finite I"
  | 
| 
 | 
   896  | 
  shows "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) \<in> sets (PiF {I} M)"
 | 
| 
 | 
   897  | 
proof cases
  | 
| 
 | 
   898  | 
  assume "i \<in> I"
  | 
| 
 | 
   899  | 
  hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
 | 
| 
 | 
   900  | 
    Pi' I (\<lambda>x. if x = i then A else space (M x))"
  | 
| 
 | 
   901  | 
    using sets_into_space[OF ] `A \<in> sets (M i)` assms
  | 
| 
 | 
   902  | 
    by (auto simp: space_PiF Pi'_def)
  | 
| 
 | 
   903  | 
  thus ?thesis  using assms `A \<in> sets (M i)`
  | 
| 
 | 
   904  | 
    by (intro in_sets_PiFI) auto
  | 
| 
 | 
   905  | 
next
  | 
| 
 | 
   906  | 
  assume "i \<notin> I"
  | 
| 
 | 
   907  | 
  hence "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space (PiF {I} M) =
 | 
| 
 | 
   908  | 
    (if undefined \<in> A then space (PiF {I} M) else {})" by (auto simp: space_PiF Pi'_def)
 | 
| 
 | 
   909  | 
  thus ?thesis by simp
  | 
| 
 | 
   910  | 
qed
  | 
| 
 | 
   911  | 
  | 
| 
 | 
   912  | 
lemma measurable_proj_singleton:
  | 
| 
 | 
   913  | 
  fixes I
  | 
| 
 | 
   914  | 
  assumes "finite I" "i \<in> I"
  | 
| 
 | 
   915  | 
  shows "(\<lambda>x. (x)\<^isub>F i) \<in> measurable (PiF {I} M) (M i)"
 | 
| 
 | 
   916  | 
proof (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms)
  | 
| 
 | 
   917  | 
qed (insert `i \<in> I`, auto simp: space_PiF)
  | 
| 
 | 
   918  | 
  | 
| 
 | 
   919  | 
lemma measurable_proj_countable:
  | 
| 
 | 
   920  | 
  fixes I::"'a::countable set set"
  | 
| 
 | 
   921  | 
  assumes "y \<in> space (M i)"
  | 
| 
 | 
   922  | 
  shows "(\<lambda>x. if i \<in> domain x then (x)\<^isub>F i else y) \<in> measurable (PiF I M) (M i)"
  | 
| 
 | 
   923  | 
proof (rule countable_measurable_PiFI)
  | 
| 
 | 
   924  | 
  fix J assume "J \<in> I" "finite J"
  | 
| 
 | 
   925  | 
  show "(\<lambda>x. if i \<in> domain x then x i else y) \<in> measurable (PiF {J} M) (M i)"
 | 
| 
 | 
   926  | 
    unfolding measurable_def
  | 
| 
 | 
   927  | 
  proof safe
  | 
| 
 | 
   928  | 
    fix z assume "z \<in> sets (M i)"
  | 
| 
 | 
   929  | 
    have "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) =
 | 
| 
 | 
   930  | 
      (\<lambda>x. if i \<in> J then (x)\<^isub>F i else y) -` z \<inter> space (PiF {J} M)"
 | 
| 
 | 
   931  | 
      by (auto simp: space_PiF Pi'_def)
  | 
| 
 | 
   932  | 
    also have "\<dots> \<in> sets (PiF {J} M)" using `z \<in> sets (M i)` `finite J`
 | 
| 
 | 
   933  | 
      by (cases "i \<in> J") (auto intro!: measurable_sets[OF measurable_proj_singleton])
  | 
| 
 | 
   934  | 
    finally show "(\<lambda>x. if i \<in> domain x then x i else y) -` z \<inter> space (PiF {J} M) \<in>
 | 
| 
 | 
   935  | 
      sets (PiF {J} M)" .
 | 
| 
 | 
   936  | 
  qed (insert `y \<in> space (M i)`, auto simp: space_PiF Pi'_def)
  | 
| 
 | 
   937  | 
qed
  | 
| 
 | 
   938  | 
  | 
| 
 | 
   939  | 
lemma measurable_restrict_proj:
  | 
| 
 | 
   940  | 
  assumes "J \<in> II" "finite J"
  | 
| 
 | 
   941  | 
  shows "finmap_of J \<in> measurable (PiM J M) (PiF II M)"
  | 
| 
 | 
   942  | 
  using assms
  | 
| 
 | 
   943  | 
  by (intro measurable_finmap_of measurable_component_singleton) auto
  | 
| 
 | 
   944  | 
  | 
| 
 | 
   945  | 
lemma
  | 
| 
 | 
   946  | 
  measurable_proj_PiM:
  | 
| 
 | 
   947  | 
  fixes J K ::"'a::countable set" and I::"'a set set"
  | 
| 
 | 
   948  | 
  assumes "finite J" "J \<in> I"
  | 
| 
 | 
   949  | 
  assumes "x \<in> space (PiM J M)"
  | 
| 
 | 
   950  | 
  shows "proj \<in>
  | 
| 
 | 
   951  | 
    measurable (PiF {J} M) (PiM J M)"
 | 
| 
 | 
   952  | 
proof (rule measurable_PiM_single)
  | 
| 
 | 
   953  | 
  show "proj \<in> space (PiF {J} M) \<rightarrow> (\<Pi>\<^isub>E i \<in> J. space (M i))"
 | 
| 
 | 
   954  | 
    using assms by (auto simp add: space_PiM space_PiF extensional_def sets_PiF Pi'_def)
  | 
| 
 | 
   955  | 
next
  | 
| 
 | 
   956  | 
  fix A i assume A: "i \<in> J" "A \<in> sets (M i)"
  | 
| 
 | 
   957  | 
  show "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} \<in> sets (PiF {J} M)"
 | 
| 
 | 
   958  | 
  proof
  | 
| 
 | 
   959  | 
    have "{\<omega> \<in> space (PiF {J} M). (\<omega>)\<^isub>F i \<in> A} =
 | 
| 
 | 
   960  | 
      (\<lambda>\<omega>. (\<omega>)\<^isub>F i) -` A \<inter> space (PiF {J} M)" by auto
 | 
| 
 | 
   961  | 
    also have "\<dots> \<in> sets (PiF {J} M)"
 | 
| 
 | 
   962  | 
      using assms A by (auto intro: measurable_sets[OF measurable_proj_singleton] simp: space_PiM)
  | 
| 
 | 
   963  | 
    finally show ?thesis .
  | 
| 
 | 
   964  | 
  qed simp
  | 
| 
 | 
   965  | 
qed
  | 
| 
 | 
   966  | 
  | 
| 
 | 
   967  | 
lemma sets_subspaceI:
  | 
| 
 | 
   968  | 
  assumes "A \<inter> space M \<in> sets M"
  | 
| 
 | 
   969  | 
  assumes "B \<in> sets M"
  | 
| 
 | 
   970  | 
  shows "A \<inter> B \<in> sets M" using assms
  | 
| 
 | 
   971  | 
proof -
  | 
| 
 | 
   972  | 
  have "A \<inter> B = (A \<inter> space M) \<inter> B"
  | 
| 
 | 
   973  | 
    using assms sets_into_space by auto
  | 
| 
 | 
   974  | 
  thus ?thesis using assms by auto
  | 
| 
 | 
   975  | 
qed
  | 
| 
 | 
   976  | 
  | 
| 
 | 
   977  | 
lemma space_PiF_singleton_eq_product:
  | 
| 
 | 
   978  | 
  assumes "finite I"
  | 
| 
 | 
   979  | 
  shows "space (PiF {I} M) = (\<Pi>' i\<in>I. space (M i))"
 | 
| 
 | 
   980  | 
  by (auto simp: product_def space_PiF assms)
  | 
| 
 | 
   981  | 
  | 
| 
 | 
   982  | 
text {* adapted from @{thm sets_PiM_single} *}
 | 
| 
 | 
   983  | 
  | 
| 
 | 
   984  | 
lemma sets_PiF_single:
  | 
| 
 | 
   985  | 
  assumes "finite I" "I \<noteq> {}"
 | 
| 
 | 
   986  | 
  shows "sets (PiF {I} M) =
 | 
| 
 | 
   987  | 
    sigma_sets (\<Pi>' i\<in>I. space (M i))
  | 
| 
 | 
   988  | 
      {{f\<in>\<Pi>' i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
 | 
| 
 | 
   989  | 
    (is "_ = sigma_sets ?\<Omega> ?R")
  | 
| 
 | 
   990  | 
  unfolding sets_PiF_singleton
  | 
| 
 | 
   991  | 
proof (rule sigma_sets_eqI)
  | 
| 
 | 
   992  | 
  interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
  | 
| 
 | 
   993  | 
  fix A assume "A \<in> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
 | 
| 
 | 
   994  | 
  then obtain X where X: "A = Pi' I X" "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
  | 
| 
 | 
   995  | 
  show "A \<in> sigma_sets ?\<Omega> ?R"
  | 
| 
 | 
   996  | 
  proof -
  | 
| 
 | 
   997  | 
    from `I \<noteq> {}` X have "A = (\<Inter>j\<in>I. {f\<in>space (PiF {I} M). f j \<in> X j})"
 | 
| 
 | 
   998  | 
      using sets_into_space
  | 
| 
 | 
   999  | 
      by (auto simp: space_PiF product_def) blast
  | 
| 
 | 
  1000  | 
    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
  | 
| 
 | 
  1001  | 
      using X `I \<noteq> {}` assms by (intro R.finite_INT) (auto simp: space_PiF)
 | 
| 
 | 
  1002  | 
    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
  | 
| 
 | 
  1003  | 
  qed
  | 
| 
 | 
  1004  | 
next
  | 
| 
 | 
  1005  | 
  fix A assume "A \<in> ?R"
  | 
| 
 | 
  1006  | 
  then obtain i B where A: "A = {f\<in>\<Pi>' i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
 | 
| 
 | 
  1007  | 
    by auto
  | 
| 
 | 
  1008  | 
  then have "A = (\<Pi>' j \<in> I. if j = i then B else space (M j))"
  | 
| 
 | 
  1009  | 
    using sets_into_space[OF A(3)]
  | 
| 
 | 
  1010  | 
    apply (auto simp: Pi'_iff split: split_if_asm)
  | 
| 
 | 
  1011  | 
    apply blast
  | 
| 
 | 
  1012  | 
    done
  | 
| 
 | 
  1013  | 
  also have "\<dots> \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}"
 | 
| 
 | 
  1014  | 
    using A
  | 
| 
 | 
  1015  | 
    by (intro sigma_sets.Basic )
  | 
| 
 | 
  1016  | 
       (auto intro: exI[where x="\<lambda>j. if j = i then B else space (M j)"])
  | 
| 
 | 
  1017  | 
  finally show "A \<in> sigma_sets ?\<Omega> {Pi' I X |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" .
 | 
| 
 | 
  1018  | 
qed
  | 
| 
 | 
  1019  | 
  | 
| 
 | 
  1020  | 
text {* adapted from @{thm PiE_cong} *}
 | 
| 
 | 
  1021  | 
  | 
| 
 | 
  1022  | 
lemma Pi'_cong:
  | 
| 
 | 
  1023  | 
  assumes "finite I"
  | 
| 
 | 
  1024  | 
  assumes "\<And>i. i \<in> I \<Longrightarrow> f i = g i"
  | 
| 
 | 
  1025  | 
  shows "Pi' I f = Pi' I g"
  | 
| 
 | 
  1026  | 
using assms by (auto simp: Pi'_def)
  | 
| 
 | 
  1027  | 
  | 
| 
 | 
  1028  | 
text {* adapted from @{thm Pi_UN} *}
 | 
| 
 | 
  1029  | 
  | 
| 
 | 
  1030  | 
lemma Pi'_UN:
  | 
| 
 | 
  1031  | 
  fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
  | 
| 
 | 
  1032  | 
  assumes "finite I"
  | 
| 
 | 
  1033  | 
  assumes mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
  | 
| 
 | 
  1034  | 
  shows "(\<Union>n. Pi' I (A n)) = Pi' I (\<lambda>i. \<Union>n. A n i)"
  | 
| 
 | 
  1035  | 
proof (intro set_eqI iffI)
  | 
| 
 | 
  1036  | 
  fix f assume "f \<in> Pi' I (\<lambda>i. \<Union>n. A n i)"
  | 
| 
 | 
  1037  | 
  then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" "domain f = I" by (auto simp: `finite I` Pi'_def)
  | 
| 
 | 
  1038  | 
  from bchoice[OF this(1)] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
  | 
| 
 | 
  1039  | 
  obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
  | 
| 
 | 
  1040  | 
    using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
  | 
| 
 | 
  1041  | 
  have "f \<in> Pi' I (\<lambda>i. A k i)"
  | 
| 
 | 
  1042  | 
  proof
  | 
| 
 | 
  1043  | 
    fix i assume "i \<in> I"
  | 
| 
 | 
  1044  | 
    from mono[OF this, of "n i" k] k[OF this] n[OF this] `domain f = I` `i \<in> I`
  | 
| 
 | 
  1045  | 
    show "f i \<in> A k i " by (auto simp: `finite I`)
  | 
| 
 | 
  1046  | 
  qed (simp add: `domain f = I` `finite I`)
  | 
| 
 | 
  1047  | 
  then show "f \<in> (\<Union>n. Pi' I (A n))" by auto
  | 
| 
 | 
  1048  | 
qed (auto simp: Pi'_def `finite I`)
  | 
| 
 | 
  1049  | 
  | 
| 
 | 
  1050  | 
text {* adapted from @{thm sigma_prod_algebra_sigma_eq} *}
 | 
| 
 | 
  1051  | 
  | 
| 
 | 
  1052  | 
lemma sigma_fprod_algebra_sigma_eq:
  | 
| 
 | 
  1053  | 
  fixes E :: "'i \<Rightarrow> 'a set set"
  | 
| 
 | 
  1054  | 
  assumes [simp]: "finite I" "I \<noteq> {}"
 | 
| 
 | 
  1055  | 
  assumes S_mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
  | 
| 
 | 
  1056  | 
    and S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
  | 
| 
 | 
  1057  | 
    and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
  | 
| 
 | 
  1058  | 
  assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
  | 
| 
 | 
  1059  | 
    and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
  | 
| 
 | 
  1060  | 
  defines "P == { Pi' I F | F. \<forall>i\<in>I. F i \<in> E i }"
 | 
| 
 | 
  1061  | 
  shows "sets (PiF {I} M) = sigma_sets (space (PiF {I} M)) P"
 | 
| 
 | 
  1062  | 
proof
  | 
| 
 | 
  1063  | 
  let ?P = "sigma (space (Pi\<^isub>F {I} M)) P"
 | 
| 
 | 
  1064  | 
  have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>F {I} M))"
 | 
| 
 | 
  1065  | 
    using E_closed by (auto simp: space_PiF P_def Pi'_iff subset_eq)
  | 
| 
 | 
  1066  | 
  then have space_P: "space ?P = (\<Pi>' i\<in>I. space (M i))"
  | 
| 
 | 
  1067  | 
    by (simp add: space_PiF)
  | 
| 
 | 
  1068  | 
  have "sets (PiF {I} M) =
 | 
| 
 | 
  1069  | 
      sigma_sets (space ?P) {{f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
 | 
| 
 | 
  1070  | 
    using sets_PiF_single[of I M] by (simp add: space_P)
  | 
| 
 | 
  1071  | 
  also have "\<dots> \<subseteq> sets (sigma (space (PiF {I} M)) P)"
 | 
| 
 | 
  1072  | 
  proof (safe intro!: sigma_sets_subset)
  | 
| 
 | 
  1073  | 
    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
  | 
| 
 | 
  1074  | 
    have "(\<lambda>x. (x)\<^isub>F i) \<in> measurable ?P (sigma (space (M i)) (E i))"
  | 
| 
 | 
  1075  | 
    proof (subst measurable_iff_measure_of)
  | 
| 
 | 
  1076  | 
      show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact
  | 
| 
 | 
  1077  | 
      from space_P `i \<in> I` show "(\<lambda>x. (x)\<^isub>F i) \<in> space ?P \<rightarrow> space (M i)"
  | 
| 
 | 
  1078  | 
        by auto
  | 
| 
 | 
  1079  | 
      show "\<forall>A\<in>E i. (\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
  | 
| 
 | 
  1080  | 
      proof
  | 
| 
 | 
  1081  | 
        fix A assume A: "A \<in> E i"
  | 
| 
 | 
  1082  | 
        then have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P = (\<Pi>' j\<in>I. if i = j then A else space (M j))"
  | 
| 
 | 
  1083  | 
          using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
  | 
| 
 | 
  1084  | 
        also have "\<dots> = (\<Pi>' j\<in>I. \<Union>n. if i = j then A else S j n)"
  | 
| 
 | 
  1085  | 
          by (intro Pi'_cong) (simp_all add: S_union)
  | 
| 
 | 
  1086  | 
        also have "\<dots> = (\<Union>n. \<Pi>' j\<in>I. if i = j then A else S j n)"
  | 
| 
 | 
  1087  | 
          using S_mono
  | 
| 
 | 
  1088  | 
          by (subst Pi'_UN[symmetric, OF `finite I`]) (auto simp: incseq_def)
  | 
| 
 | 
  1089  | 
        also have "\<dots> \<in> sets ?P"
  | 
| 
 | 
  1090  | 
        proof (safe intro!: countable_UN)
  | 
| 
 | 
  1091  | 
          fix n show "(\<Pi>' j\<in>I. if i = j then A else S j n) \<in> sets ?P"
  | 
| 
 | 
  1092  | 
            using A S_in_E
  | 
| 
 | 
  1093  | 
            by (simp add: P_closed)
  | 
| 
 | 
  1094  | 
               (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j n"])
  | 
| 
 | 
  1095  | 
        qed
  | 
| 
 | 
  1096  | 
        finally show "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
  | 
| 
 | 
  1097  | 
          using P_closed by simp
  | 
| 
 | 
  1098  | 
      qed
  | 
| 
 | 
  1099  | 
    qed
  | 
| 
 | 
  1100  | 
    from measurable_sets[OF this, of A] A `i \<in> I` E_closed
  | 
| 
 | 
  1101  | 
    have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P \<in> sets ?P"
  | 
| 
 | 
  1102  | 
      by (simp add: E_generates)
  | 
| 
 | 
  1103  | 
    also have "(\<lambda>x. (x)\<^isub>F i) -` A \<inter> space ?P = {f \<in> \<Pi>' i\<in>I. space (M i). f i \<in> A}"
 | 
| 
 | 
  1104  | 
      using P_closed by (auto simp: space_PiF)
  | 
| 
 | 
  1105  | 
    finally show "\<dots> \<in> sets ?P" .
  | 
| 
 | 
  1106  | 
  qed
  | 
| 
 | 
  1107  | 
  finally show "sets (PiF {I} M) \<subseteq> sigma_sets (space (PiF {I} M)) P"
 | 
| 
 | 
  1108  | 
    by (simp add: P_closed)
  | 
| 
 | 
  1109  | 
  show "sigma_sets (space (PiF {I} M)) P \<subseteq> sets (PiF {I} M)"
 | 
| 
 | 
  1110  | 
    using `finite I` `I \<noteq> {}`
 | 
| 
 | 
  1111  | 
    by (auto intro!: sigma_sets_subset product_in_sets_PiFI simp: E_generates P_def)
  | 
| 
 | 
  1112  | 
qed
  | 
| 
 | 
  1113  | 
  | 
| 
 | 
  1114  | 
lemma enumerable_sigma_fprod_algebra_sigma_eq:
  | 
| 
 | 
  1115  | 
  assumes "I \<noteq> {}"
 | 
| 
 | 
  1116  | 
  assumes [simp]: "finite I"
  | 
| 
 | 
  1117  | 
  shows "sets (PiF {I} (\<lambda>_. borel)) = sigma_sets (space (PiF {I} (\<lambda>_. borel)))
 | 
| 
 | 
  1118  | 
    {Pi' I F |F. (\<forall>i\<in>I. F i \<in> range enum_basis)}"
 | 
| 
 | 
  1119  | 
proof -
  | 
| 
 | 
  1120  | 
  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  | 
| 
 | 
  1121  | 
  show ?thesis
  | 
| 
 | 
  1122  | 
  proof (rule sigma_fprod_algebra_sigma_eq)
  | 
| 
 | 
  1123  | 
    show "finite I" by simp
  | 
| 
 | 
  1124  | 
    show "I \<noteq> {}" by fact
 | 
| 
 | 
  1125  | 
    show "incseq S" "(\<Union>j. S j) = space borel" "range S \<subseteq> range enum_basis"
  | 
| 
 | 
  1126  | 
      using S by simp_all
  | 
| 
 | 
  1127  | 
    show "range enum_basis \<subseteq> Pow (space borel)" by simp
  | 
| 
 | 
  1128  | 
    show "sets borel = sigma_sets (space borel) (range enum_basis)"
  | 
| 
 | 
  1129  | 
      by (simp add: borel_eq_enum_basis)
  | 
| 
 | 
  1130  | 
  qed
  | 
| 
 | 
  1131  | 
qed
  | 
| 
 | 
  1132  | 
  | 
| 
 | 
  1133  | 
text {* adapted from @{thm enumerable_sigma_fprod_algebra_sigma_eq} *}
 | 
| 
 | 
  1134  | 
  | 
| 
 | 
  1135  | 
lemma enumerable_sigma_prod_algebra_sigma_eq:
  | 
| 
 | 
  1136  | 
  assumes "I \<noteq> {}"
 | 
| 
 | 
  1137  | 
  assumes [simp]: "finite I"
  | 
| 
 | 
  1138  | 
  shows "sets (PiM I (\<lambda>_. borel)) = sigma_sets (space (PiM I (\<lambda>_. borel)))
  | 
| 
 | 
  1139  | 
    {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range enum_basis}"
 | 
| 
 | 
  1140  | 
proof -
  | 
| 
 | 
  1141  | 
  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  | 
| 
 | 
  1142  | 
  show ?thesis
  | 
| 
 | 
  1143  | 
  proof (rule sigma_prod_algebra_sigma_eq)
  | 
| 
 | 
  1144  | 
    show "finite I" by simp note[[show_types]]
  | 
| 
 | 
  1145  | 
    fix i show "(\<Union>j. S j) = space borel" "range S \<subseteq> range enum_basis"
  | 
| 
 | 
  1146  | 
      using S by simp_all
  | 
| 
 | 
  1147  | 
    show "range enum_basis \<subseteq> Pow (space borel)" by simp
  | 
| 
 | 
  1148  | 
    show "sets borel = sigma_sets (space borel) (range enum_basis)"
  | 
| 
 | 
  1149  | 
      by (simp add: borel_eq_enum_basis)
  | 
| 
 | 
  1150  | 
  qed
  | 
| 
 | 
  1151  | 
qed
  | 
| 
 | 
  1152  | 
  | 
| 
 | 
  1153  | 
lemma product_open_generates_sets_PiF_single:
  | 
| 
 | 
  1154  | 
  assumes "I \<noteq> {}"
 | 
| 
 | 
  1155  | 
  assumes [simp]: "finite I"
  | 
| 
 | 
  1156  | 
  shows "sets (PiF {I} (\<lambda>_. borel::'b::enumerable_basis measure)) =
 | 
| 
 | 
  1157  | 
    sigma_sets (space (PiF {I} (\<lambda>_. borel))) {Pi' I F |F. (\<forall>i\<in>I. F i \<in> Collect open)}"
 | 
| 
 | 
  1158  | 
proof -
  | 
| 
 | 
  1159  | 
  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  | 
| 
 | 
  1160  | 
  show ?thesis
  | 
| 
 | 
  1161  | 
  proof (rule sigma_fprod_algebra_sigma_eq)
  | 
| 
 | 
  1162  | 
    show "finite I" by simp
  | 
| 
 | 
  1163  | 
    show "I \<noteq> {}" by fact
 | 
| 
 | 
  1164  | 
    show "incseq S" "(\<Union>j. S j) = space borel" "range S \<subseteq> Collect open"
  | 
| 
 | 
  1165  | 
      using S by (auto simp: open_enum_basis)
  | 
| 
 | 
  1166  | 
    show "Collect open \<subseteq> Pow (space borel)" by simp
  | 
| 
 | 
  1167  | 
    show "sets borel = sigma_sets (space borel) (Collect open)"
  | 
| 
 | 
  1168  | 
      by (simp add: borel_def)
  | 
| 
 | 
  1169  | 
  qed
  | 
| 
 | 
  1170  | 
qed
  | 
| 
 | 
  1171  | 
  | 
| 
 | 
  1172  | 
lemma product_open_generates_sets_PiM:
  | 
| 
 | 
  1173  | 
  assumes "I \<noteq> {}"
 | 
| 
 | 
  1174  | 
  assumes [simp]: "finite I"
  | 
| 
 | 
  1175  | 
  shows "sets (PiM I (\<lambda>_. borel::'b::enumerable_basis measure)) =
  | 
| 
 | 
  1176  | 
    sigma_sets (space (PiM I (\<lambda>_. borel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> Collect open}"
 | 
| 
 | 
  1177  | 
proof -
  | 
| 
 | 
  1178  | 
  from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'b set" . note S = this
  | 
| 
 | 
  1179  | 
  show ?thesis
  | 
| 
 | 
  1180  | 
  proof (rule sigma_prod_algebra_sigma_eq)
  | 
| 
 | 
  1181  | 
    show "finite I" by simp note[[show_types]]
  | 
| 
 | 
  1182  | 
    fix i show "(\<Union>j. S j) = space borel" "range S \<subseteq> Collect open"
  | 
| 
 | 
  1183  | 
      using S by (auto simp: open_enum_basis)
  | 
| 
 | 
  1184  | 
    show "Collect open \<subseteq> Pow (space borel)" by simp
  | 
| 
 | 
  1185  | 
    show "sets borel = sigma_sets (space borel) (Collect open)"
  | 
| 
 | 
  1186  | 
      by (simp add: borel_def)
  | 
| 
 | 
  1187  | 
  qed
  | 
| 
 | 
  1188  | 
qed
  | 
| 
 | 
  1189  | 
  | 
| 
 | 
  1190  | 
lemma finmap_UNIV[simp]: "(\<Union>J\<in>Collect finite. J \<leadsto> UNIV) = UNIV" by auto
  | 
| 
 | 
  1191  | 
  | 
| 
 | 
  1192  | 
lemma borel_eq_PiF_borel:
  | 
| 
 | 
  1193  | 
  shows "(borel :: ('i::countable \<Rightarrow>\<^isub>F 'a::polish_space) measure) =
 | 
| 
 | 
  1194  | 
  PiF (Collect finite) (\<lambda>_. borel :: 'a measure)"
  | 
| 
 | 
  1195  | 
proof (rule measure_eqI)
  | 
| 
 | 
  1196  | 
  have C: "Collect finite \<noteq> {}" by auto
 | 
| 
 | 
  1197  | 
  show "sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure) = sets (PiF (Collect finite) (\<lambda>_. borel))"
 | 
| 
 | 
  1198  | 
  proof
  | 
| 
 | 
  1199  | 
    show "sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure) \<subseteq> sets (PiF (Collect finite) (\<lambda>_. borel))"
 | 
| 
 | 
  1200  | 
      apply (simp add: borel_def sets_PiF)
  | 
| 
 | 
  1201  | 
    proof (rule sigma_sets_mono, safe, cases)
  | 
| 
 | 
  1202  | 
      fix X::"('i \<Rightarrow>\<^isub>F 'a) set" assume "open X" "X \<noteq> {}"
 | 
| 
 | 
  1203  | 
      from open_basisE[OF this] guess NA NB . note N = this
  | 
| 
 | 
  1204  | 
      hence "X = (\<Union>i. Pi' (NA i) (NB i))" by simp
  | 
| 
 | 
  1205  | 
      also have "\<dots> \<in>
  | 
| 
 | 
  1206  | 
        sigma_sets UNIV {Pi' J S |S J. finite J \<and> S \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}"
 | 
| 
 | 
  1207  | 
        using N by (intro Union sigma_sets.Basic) blast
  | 
| 
 | 
  1208  | 
      finally show "X \<in> sigma_sets UNIV
  | 
| 
 | 
  1209  | 
        {Pi' J X |X J. finite J \<and> X \<in> J \<rightarrow> sigma_sets UNIV (Collect open)}" .
 | 
| 
 | 
  1210  | 
    qed (auto simp: Empty)
  | 
| 
 | 
  1211  | 
  next
  | 
| 
 | 
  1212  | 
    show "sets (PiF (Collect finite) (\<lambda>_. borel)) \<subseteq> sets (borel::('i \<Rightarrow>\<^isub>F 'a) measure)"
 | 
| 
 | 
  1213  | 
    proof
  | 
| 
 | 
  1214  | 
      fix x assume x: "x \<in> sets (PiF (Collect finite::'i set set) (\<lambda>_. borel::'a measure))"
  | 
| 
 | 
  1215  | 
      hence x_sp: "x \<subseteq> space (PiF (Collect finite) (\<lambda>_. borel))" by (rule sets_into_space)
  | 
| 
 | 
  1216  | 
      let ?x = "\<lambda>J. x \<inter> {x. domain x = J}"
 | 
| 
 | 
  1217  | 
      have "x = \<Union>{?x J |J. finite J}" by auto
 | 
| 
 | 
  1218  | 
      also have "\<dots> \<in> sets borel"
  | 
| 
 | 
  1219  | 
      proof (rule countable_finite_comprehension, assumption)
  | 
| 
 | 
  1220  | 
        fix J::"'i set" assume "finite J"
  | 
| 
 | 
  1221  | 
        { assume ef: "J = {}"
 | 
| 
 | 
  1222  | 
          { assume e: "?x J = {}"
 | 
| 
 | 
  1223  | 
            hence "?x J \<in> sets borel" by simp
  | 
| 
 | 
  1224  | 
          } moreover {
 | 
| 
 | 
  1225  | 
            assume "?x J \<noteq> {}"
 | 
| 
 | 
  1226  | 
            then obtain f where "f \<in> x" "domain f = {}" using ef by auto
 | 
| 
 | 
  1227  | 
            hence "?x J = {f}" using `J = {}`
 | 
| 
 | 
  1228  | 
              by (auto simp: finmap_eq_iff)
  | 
| 
 | 
  1229  | 
            also have "{f} \<in> sets borel" by simp
 | 
| 
 | 
  1230  | 
            finally have "?x J \<in> sets borel" .
  | 
| 
 | 
  1231  | 
          } ultimately have "?x J \<in> sets borel" by blast
  | 
| 
 | 
  1232  | 
        } moreover {
 | 
| 
 | 
  1233  | 
          assume "J \<noteq> ({}::'i set)"
 | 
| 
 | 
  1234  | 
          from open_incseqE[OF open_UNIV] guess S::"nat \<Rightarrow> 'a set" . note S = this
  | 
| 
 | 
  1235  | 
          have "(?x J) = x \<inter> {m. domain m \<in> {J}}" by auto
 | 
| 
 | 
  1236  | 
          also have "\<dots> \<in> sets (PiF {J} (\<lambda>_. borel))"
 | 
| 
 | 
  1237  | 
            using x by (rule restrict_sets_measurable) (auto simp: `finite J`)
  | 
| 
 | 
  1238  | 
          also have "\<dots> = sigma_sets (space (PiF {J} (\<lambda>_. borel)))
 | 
| 
 | 
  1239  | 
            {Pi' (J) F |F. (\<forall>j\<in>J. F j \<in> range enum_basis)}"
 | 
| 
 | 
  1240  | 
            (is "_ = sigma_sets _ ?P")
  | 
| 
 | 
  1241  | 
            by (rule enumerable_sigma_fprod_algebra_sigma_eq[OF `J \<noteq> {}` `finite J`])
 | 
| 
 | 
  1242  | 
          also have "\<dots> \<subseteq> sets borel"
  | 
| 
 | 
  1243  | 
          proof
  | 
| 
 | 
  1244  | 
            fix x
  | 
| 
 | 
  1245  | 
            assume "x \<in> sigma_sets (space (PiF {J} (\<lambda>_. borel))) ?P"
 | 
| 
 | 
  1246  | 
            thus "x \<in> sets borel"
  | 
| 
 | 
  1247  | 
            proof (rule sigma_sets.induct, safe)
  | 
| 
 | 
  1248  | 
              fix F::"'i \<Rightarrow> 'a set"
  | 
| 
 | 
  1249  | 
              assume "\<forall>j\<in>J. F j \<in> range enum_basis"
  | 
| 
 | 
  1250  | 
              hence "Pi' J F \<in> range enum_basis_finmap"
  | 
| 
 | 
  1251  | 
                unfolding range_enum_basis_eq
  | 
| 
 | 
  1252  | 
                by (auto simp: `finite J` intro!: exI[where x=J] exI[where x=F])
  | 
| 
 | 
  1253  | 
              hence "open (Pi' (J) F)" by (rule range_enum_basis_finmap_imp_open)
  | 
| 
 | 
  1254  | 
              thus "Pi' (J) F \<in> sets borel" by simp
  | 
| 
 | 
  1255  | 
            next
  | 
| 
 | 
  1256  | 
              fix a::"('i \<Rightarrow>\<^isub>F 'a) set"
 | 
| 
 | 
  1257  | 
              have "space (PiF {J::'i set} (\<lambda>_. borel::'a measure)) =
 | 
| 
 | 
  1258  | 
                Pi' (J) (\<lambda>_. UNIV)"
  | 
| 
 | 
  1259  | 
                by (auto simp: space_PiF product_def)
  | 
| 
 | 
  1260  | 
              moreover have "open (Pi' (J::'i set) (\<lambda>_. UNIV::'a set))"
  | 
| 
 | 
  1261  | 
                by (intro open_Pi'I) auto
  | 
| 
 | 
  1262  | 
              ultimately
  | 
| 
 | 
  1263  | 
              have "space (PiF {J::'i set} (\<lambda>_. borel::'a measure)) \<in> sets borel"
 | 
| 
 | 
  1264  | 
                by simp
  | 
| 
 | 
  1265  | 
              moreover
  | 
| 
 | 
  1266  | 
              assume "a \<in> sets borel"
  | 
| 
 | 
  1267  | 
              ultimately show "space (PiF {J} (\<lambda>_. borel)) - a \<in> sets borel" ..
 | 
| 
 | 
  1268  | 
            qed auto
  | 
| 
 | 
  1269  | 
          qed
  | 
| 
 | 
  1270  | 
          finally have "(?x J) \<in> sets borel" .
  | 
| 
 | 
  1271  | 
        } ultimately show "(?x J) \<in> sets borel" by blast
  | 
| 
 | 
  1272  | 
      qed
  | 
| 
 | 
  1273  | 
      finally show "x \<in> sets (borel)" .
  | 
| 
 | 
  1274  | 
    qed
  | 
| 
 | 
  1275  | 
  qed
  | 
| 
 | 
  1276  | 
qed (simp add: emeasure_sigma borel_def PiF_def)
  | 
| 
 | 
  1277  | 
  | 
| 
 | 
  1278  | 
subsection {* Isomorphism between Functions and Finite Maps *}
 | 
| 
 | 
  1279  | 
  | 
| 
 | 
  1280  | 
lemma
  | 
| 
 | 
  1281  | 
  measurable_compose:
  | 
| 
 | 
  1282  | 
  fixes f::"'a \<Rightarrow> 'b"
  | 
| 
 | 
  1283  | 
  assumes inj: "\<And>j. j \<in> J \<Longrightarrow> f' (f j) = j"
  | 
| 
 | 
  1284  | 
  assumes "finite J"
  | 
| 
 | 
  1285  | 
  shows "(\<lambda>m. compose J m f) \<in> measurable (PiM (f ` J) (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
  | 
| 
 | 
  1286  | 
proof (rule measurable_PiM)
  | 
| 
 | 
  1287  | 
  show "(\<lambda>m. compose J m f)
  | 
| 
 | 
  1288  | 
    \<in> space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) \<rightarrow>
  | 
| 
 | 
  1289  | 
      (J \<rightarrow> space M) \<inter> extensional J"
  | 
| 
 | 
  1290  | 
  proof safe
  | 
| 
 | 
  1291  | 
    fix x and i
  | 
| 
 | 
  1292  | 
    assume x: "x \<in> space (PiM (f ` J) (\<lambda>_. M))" "i \<in> J"
  | 
| 
 | 
  1293  | 
    with inj show  "compose J x f i \<in> space M"
  | 
| 
 | 
  1294  | 
      by (auto simp: space_PiM compose_def)
  | 
| 
 | 
  1295  | 
  next
  | 
| 
 | 
  1296  | 
    fix x assume "x \<in> space (PiM (f ` J) (\<lambda>_. M))"
  | 
| 
 | 
  1297  | 
    show "(compose J x f) \<in> extensional J" by (rule compose_extensional)
  | 
| 
 | 
  1298  | 
  qed
  | 
| 
 | 
  1299  | 
next
  | 
| 
 | 
  1300  | 
  fix S X
  | 
| 
 | 
  1301  | 
  have inv: "\<And>j. j \<in> f ` J \<Longrightarrow> f (f' j) = j" using assms by auto
  | 
| 
 | 
  1302  | 
  assume S: "S \<noteq> {} \<or> J = {}" "finite S" "S \<subseteq> J" and P: "\<And>i. i \<in> S \<Longrightarrow> X i \<in> sets M"
 | 
| 
 | 
  1303  | 
  have "(\<lambda>m. compose J m f) -` prod_emb J (\<lambda>_. M) S (Pi\<^isub>E S X) \<inter>
  | 
| 
 | 
  1304  | 
    space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) = prod_emb (f ` J) (\<lambda>_. M) (f ` S) (Pi\<^isub>E (f ` S) (\<lambda>b. X (f' b)))"
  | 
| 
 | 
  1305  | 
    using assms inv S sets_into_space[OF P]
  | 
| 
 | 
  1306  | 
    by (force simp: prod_emb_iff compose_def space_PiM extensional_def Pi_def intro: imageI)
  | 
| 
 | 
  1307  | 
  also have "\<dots> \<in> sets (Pi\<^isub>M (f ` J) (\<lambda>_. M))"
  | 
| 
 | 
  1308  | 
  proof
  | 
| 
 | 
  1309  | 
    from S show "f ` S \<subseteq> f `  J" by auto
  | 
| 
 | 
  1310  | 
    show "(\<Pi>\<^isub>E b\<in>f ` S. X (f' b)) \<in> sets (Pi\<^isub>M (f ` S) (\<lambda>_. M))"
  | 
| 
 | 
  1311  | 
    proof (rule sets_PiM_I_finite)
  | 
| 
 | 
  1312  | 
      show "finite (f ` S)" using S by simp
  | 
| 
 | 
  1313  | 
      fix i assume "i \<in> f ` S" hence "f' i \<in> S" using S assms by auto
  | 
| 
 | 
  1314  | 
      thus "X (f' i) \<in> sets M" by (rule P)
  | 
| 
 | 
  1315  | 
    qed
  | 
| 
 | 
  1316  | 
  qed
  | 
| 
 | 
  1317  | 
  finally show "(\<lambda>m. compose J m f) -` prod_emb J (\<lambda>_. M) S (Pi\<^isub>E S X) \<inter>
  | 
| 
 | 
  1318  | 
    space (Pi\<^isub>M (f ` J) (\<lambda>_. M)) \<in> sets (Pi\<^isub>M (f ` J) (\<lambda>_. M))" .
  | 
| 
 | 
  1319  | 
qed
  | 
| 
 | 
  1320  | 
  | 
| 
 | 
  1321  | 
lemma
  | 
| 
 | 
  1322  | 
  measurable_compose_inv:
  | 
| 
 | 
  1323  | 
  fixes f::"'a \<Rightarrow> 'b"
  | 
| 
 | 
  1324  | 
  assumes inj: "\<And>j. j \<in> J \<Longrightarrow> f' (f j) = j"
  | 
| 
 | 
  1325  | 
  assumes "finite J"
  | 
| 
 | 
  1326  | 
  shows "(\<lambda>m. compose (f ` J) m f') \<in> measurable (PiM J (\<lambda>_. M)) (PiM (f ` J) (\<lambda>_. M))"
  | 
| 
 | 
  1327  | 
proof -
  | 
| 
 | 
  1328  | 
  have "(\<lambda>m. compose (f ` J) m f') \<in> measurable (Pi\<^isub>M (f' ` f ` J) (\<lambda>_. M)) (Pi\<^isub>M (f ` J) (\<lambda>_. M))"
  | 
| 
 | 
  1329  | 
    using assms by (auto intro: measurable_compose)
  | 
| 
 | 
  1330  | 
  moreover
  | 
| 
 | 
  1331  | 
  from inj have "f' ` f ` J = J" by (metis (hide_lams, mono_tags) image_iff set_eqI)
  | 
| 
 | 
  1332  | 
  ultimately show ?thesis by simp
  | 
| 
 | 
  1333  | 
qed
  | 
| 
 | 
  1334  | 
  | 
| 
 | 
  1335  | 
locale function_to_finmap =
  | 
| 
 | 
  1336  | 
  fixes J::"'a set" and f :: "'a \<Rightarrow> 'b::countable" and f'
  | 
| 
 | 
  1337  | 
  assumes [simp]: "finite J"
  | 
| 
 | 
  1338  | 
  assumes inv: "i \<in> J \<Longrightarrow> f' (f i) = i"
  | 
| 
 | 
  1339  | 
begin
  | 
| 
 | 
  1340  | 
  | 
| 
 | 
  1341  | 
text {* to measure finmaps *}
 | 
| 
 | 
  1342  | 
  | 
| 
 | 
  1343  | 
definition "fm = (finmap_of (f ` J)) o (\<lambda>g. compose (f ` J) g f')"
  | 
| 
 | 
  1344  | 
  | 
| 
 | 
  1345  | 
lemma domain_fm[simp]: "domain (fm x) = f ` J"
  | 
| 
 | 
  1346  | 
  unfolding fm_def by simp
  | 
| 
 | 
  1347  | 
  | 
| 
 | 
  1348  | 
lemma fm_restrict[simp]: "fm (restrict y J) = fm y"
  | 
| 
 | 
  1349  | 
  unfolding fm_def by (auto simp: compose_def inv intro: restrict_ext)
  | 
| 
 | 
  1350  | 
  | 
| 
 | 
  1351  | 
lemma fm_product:
  | 
| 
 | 
  1352  | 
  assumes "\<And>i. space (M i) = UNIV"
  | 
| 
 | 
  1353  | 
  shows "fm -` Pi' (f ` J) S \<inter> space (Pi\<^isub>M J M) = (\<Pi>\<^isub>E j \<in> J. S (f j))"
  | 
| 
 | 
  1354  | 
  using assms
  | 
| 
 | 
  1355  | 
  by (auto simp: inv fm_def compose_def space_PiM Pi'_def)
  | 
| 
 | 
  1356  | 
  | 
| 
 | 
  1357  | 
lemma fm_measurable:
  | 
| 
 | 
  1358  | 
  assumes "f ` J \<in> N"
  | 
| 
 | 
  1359  | 
  shows "fm \<in> measurable (Pi\<^isub>M J (\<lambda>_. M)) (Pi\<^isub>F N (\<lambda>_. M))"
  | 
| 
 | 
  1360  | 
  unfolding fm_def
  | 
| 
 | 
  1361  | 
proof (rule measurable_comp, rule measurable_compose_inv)
  | 
| 
 | 
  1362  | 
  show "finmap_of (f ` J) \<in> measurable (Pi\<^isub>M (f ` J) (\<lambda>_. M)) (PiF N (\<lambda>_. M)) "
  | 
| 
 | 
  1363  | 
    using assms by (intro measurable_finmap_of measurable_component_singleton) auto
  | 
| 
 | 
  1364  | 
qed (simp_all add: inv)
  | 
| 
 | 
  1365  | 
  | 
| 
 | 
  1366  | 
lemma proj_fm:
  | 
| 
 | 
  1367  | 
  assumes "x \<in> J"
  | 
| 
 | 
  1368  | 
  shows "fm m (f x) = m x"
  | 
| 
 | 
  1369  | 
  using assms by (auto simp: fm_def compose_def o_def inv)
  | 
| 
 | 
  1370  | 
  | 
| 
 | 
  1371  | 
lemma inj_on_compose_f': "inj_on (\<lambda>g. compose (f ` J) g f') (extensional J)"
  | 
| 
 | 
  1372  | 
proof (rule inj_on_inverseI)
  | 
| 
 | 
  1373  | 
  fix x::"'a \<Rightarrow> 'c" assume "x \<in> extensional J"
  | 
| 
 | 
  1374  | 
  thus "(\<lambda>x. compose J x f) (compose (f ` J) x f') = x"
  | 
| 
 | 
  1375  | 
    by (auto simp: compose_def inv extensional_def)
  | 
| 
 | 
  1376  | 
qed
  | 
| 
 | 
  1377  | 
  | 
| 
 | 
  1378  | 
lemma inj_on_fm:
  | 
| 
 | 
  1379  | 
  assumes "\<And>i. space (M i) = UNIV"
  | 
| 
 | 
  1380  | 
  shows "inj_on fm (space (Pi\<^isub>M J M))"
  | 
| 
 | 
  1381  | 
  using assms
  | 
| 
 | 
  1382  | 
  apply (auto simp: fm_def space_PiM)
  | 
| 
 | 
  1383  | 
  apply (rule comp_inj_on)
  | 
| 
 | 
  1384  | 
  apply (rule inj_on_compose_f')
  | 
| 
 | 
  1385  | 
  apply (rule finmap_of_inj_on_extensional_finite)
  | 
| 
 | 
  1386  | 
  apply simp
  | 
| 
 | 
  1387  | 
  apply (auto)
  | 
| 
 | 
  1388  | 
  done
  | 
| 
 | 
  1389  | 
  | 
| 
 | 
  1390  | 
text {* to measure functions *}
 | 
| 
 | 
  1391  | 
  | 
| 
 | 
  1392  | 
definition "mf = (\<lambda>g. compose J g f) o proj"
  | 
| 
 | 
  1393  | 
  | 
| 
 | 
  1394  | 
lemma
  | 
| 
 | 
  1395  | 
  assumes "x \<in> space (Pi\<^isub>M J (\<lambda>_. M))" "finite J"
  | 
| 
 | 
  1396  | 
  shows "proj (finmap_of J x) = x"
  | 
| 
 | 
  1397  | 
  using assms by (auto simp: space_PiM extensional_def)
  | 
| 
 | 
  1398  | 
  | 
| 
 | 
  1399  | 
lemma
  | 
| 
 | 
  1400  | 
  assumes "x \<in> space (Pi\<^isub>F {J} (\<lambda>_. M))"
 | 
| 
 | 
  1401  | 
  shows "finmap_of J (proj x) = x"
  | 
| 
 | 
  1402  | 
  using assms by (auto simp: space_PiF Pi'_def finmap_eq_iff)
  | 
| 
 | 
  1403  | 
  | 
| 
 | 
  1404  | 
lemma mf_fm:
  | 
| 
 | 
  1405  | 
  assumes "x \<in> space (Pi\<^isub>M J (\<lambda>_. M))"
  | 
| 
 | 
  1406  | 
  shows "mf (fm x) = x"
  | 
| 
 | 
  1407  | 
proof -
  | 
| 
 | 
  1408  | 
  have "mf (fm x) \<in> extensional J"
  | 
| 
 | 
  1409  | 
    by (auto simp: mf_def extensional_def compose_def)
  | 
| 
 | 
  1410  | 
  moreover
  | 
| 
 | 
  1411  | 
  have "x \<in> extensional J" using assms sets_into_space
  | 
| 
 | 
  1412  | 
    by (force simp: space_PiM)
  | 
| 
 | 
  1413  | 
  moreover
  | 
| 
 | 
  1414  | 
  { fix i assume "i \<in> J"
 | 
| 
 | 
  1415  | 
    hence "mf (fm x) i = x i"
  | 
| 
 | 
  1416  | 
      by (auto simp: inv mf_def compose_def fm_def)
  | 
| 
 | 
  1417  | 
  }
  | 
| 
 | 
  1418  | 
  ultimately
  | 
| 
 | 
  1419  | 
  show ?thesis by (rule extensionalityI)
  | 
| 
 | 
  1420  | 
qed
  | 
| 
 | 
  1421  | 
  | 
| 
 | 
  1422  | 
lemma mf_measurable:
  | 
| 
 | 
  1423  | 
  assumes "space M = UNIV"
  | 
| 
 | 
  1424  | 
  shows "mf \<in> measurable (PiF {f ` J} (\<lambda>_. M)) (PiM J (\<lambda>_. M))"
 | 
| 
 | 
  1425  | 
  unfolding mf_def
  | 
| 
 | 
  1426  | 
proof (rule measurable_comp, rule measurable_proj_PiM)
  | 
| 
 | 
  1427  | 
  show "(\<lambda>g. compose J g f) \<in>
  | 
| 
 | 
  1428  | 
    measurable (Pi\<^isub>M (f ` J) (\<lambda>x. M)) (Pi\<^isub>M J (\<lambda>_. M))"
  | 
| 
 | 
  1429  | 
    by (rule measurable_compose, rule inv) auto
  | 
| 
 | 
  1430  | 
qed (auto simp add: space_PiM extensional_def assms)
  | 
| 
 | 
  1431  | 
  | 
| 
 | 
  1432  | 
lemma fm_image_measurable:
  | 
| 
 | 
  1433  | 
  assumes "space M = UNIV"
  | 
| 
 | 
  1434  | 
  assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M))"
  | 
| 
 | 
  1435  | 
  shows "fm ` X \<in> sets (PiF {f ` J} (\<lambda>_. M))"
 | 
| 
 | 
  1436  | 
proof -
  | 
| 
 | 
  1437  | 
  have "fm ` X = (mf) -` X \<inter> space (PiF {f ` J} (\<lambda>_. M))"
 | 
| 
 | 
  1438  | 
  proof safe
  | 
| 
 | 
  1439  | 
    fix x assume "x \<in> X"
  | 
| 
 | 
  1440  | 
    with mf_fm[of x] sets_into_space[OF assms(2)] show "fm x \<in> mf -` X" by auto
  | 
| 
 | 
  1441  | 
    show "fm x \<in> space (PiF {f ` J} (\<lambda>_. M))" by (simp add: space_PiF assms)
 | 
| 
 | 
  1442  | 
  next
  | 
| 
 | 
  1443  | 
    fix y x
  | 
| 
 | 
  1444  | 
    assume x: "mf y \<in> X"
  | 
| 
 | 
  1445  | 
    assume y: "y \<in> space (PiF {f ` J} (\<lambda>_. M))"
 | 
| 
 | 
  1446  | 
    thus "y \<in> fm ` X"
  | 
| 
 | 
  1447  | 
      by (intro image_eqI[OF _ x], unfold finmap_eq_iff)
  | 
| 
 | 
  1448  | 
         (auto simp: space_PiF fm_def mf_def compose_def inv Pi'_def)
  | 
| 
 | 
  1449  | 
  qed
  | 
| 
 | 
  1450  | 
  also have "\<dots> \<in> sets (PiF {f ` J} (\<lambda>_. M))"
 | 
| 
 | 
  1451  | 
    using assms
  | 
| 
 | 
  1452  | 
    by (intro measurable_sets[OF mf_measurable]) auto
  | 
| 
 | 
  1453  | 
  finally show ?thesis .
  | 
| 
 | 
  1454  | 
qed
  | 
| 
 | 
  1455  | 
  | 
| 
 | 
  1456  | 
lemma fm_image_measurable_finite:
  | 
| 
 | 
  1457  | 
  assumes "space M = UNIV"
  | 
| 
 | 
  1458  | 
  assumes "X \<in> sets (Pi\<^isub>M J (\<lambda>_. M::'c measure))"
  | 
| 
 | 
  1459  | 
  shows "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. M::'c measure))"
  | 
| 
 | 
  1460  | 
  using fm_image_measurable[OF assms]
  | 
| 
 | 
  1461  | 
  by (rule subspace_set_in_sets) (auto simp: finite_subset)
  | 
| 
 | 
  1462  | 
  | 
| 
 | 
  1463  | 
text {* measure on finmaps *}
 | 
| 
 | 
  1464  | 
  | 
| 
 | 
  1465  | 
definition "mapmeasure M N = distr M (PiF (Collect finite) N) (fm)"
  | 
| 
 | 
  1466  | 
  | 
| 
 | 
  1467  | 
lemma sets_mapmeasure[simp]: "sets (mapmeasure M N) = sets (PiF (Collect finite) N)"
  | 
| 
 | 
  1468  | 
  unfolding mapmeasure_def by simp
  | 
| 
 | 
  1469  | 
  | 
| 
 | 
  1470  | 
lemma space_mapmeasure[simp]: "space (mapmeasure M N) = space (PiF (Collect finite) N)"
  | 
| 
 | 
  1471  | 
  unfolding mapmeasure_def by simp
  | 
| 
 | 
  1472  | 
  | 
| 
 | 
  1473  | 
lemma mapmeasure_PiF:
  | 
| 
 | 
  1474  | 
  assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
  | 
| 
 | 
  1475  | 
  assumes s2: "sets M = (Pi\<^isub>M J (\<lambda>_. N))"
  | 
| 
 | 
  1476  | 
  assumes "space N = UNIV"
  | 
| 
 | 
  1477  | 
  assumes "X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
  | 
| 
 | 
  1478  | 
  shows "emeasure (mapmeasure M (\<lambda>_. N)) X = emeasure M ((fm -` X \<inter> extensional J))"
  | 
| 
 | 
  1479  | 
  using assms
  | 
| 
 | 
  1480  | 
  by (auto simp: measurable_eqI[OF s1 refl s2 refl] mapmeasure_def emeasure_distr
  | 
| 
 | 
  1481  | 
    fm_measurable space_PiM)
  | 
| 
 | 
  1482  | 
  | 
| 
 | 
  1483  | 
lemma mapmeasure_PiM:
  | 
| 
 | 
  1484  | 
  fixes N::"'c measure"
  | 
| 
 | 
  1485  | 
  assumes s1: "space M = space (Pi\<^isub>M J (\<lambda>_. N))"
  | 
| 
 | 
  1486  | 
  assumes s2: "sets M = (Pi\<^isub>M J (\<lambda>_. N))"
  | 
| 
 | 
  1487  | 
  assumes N: "space N = UNIV"
  | 
| 
 | 
  1488  | 
  assumes X: "X \<in> sets M"
  | 
| 
 | 
  1489  | 
  shows "emeasure M X = emeasure (mapmeasure M (\<lambda>_. N)) (fm ` X)"
  | 
| 
 | 
  1490  | 
  unfolding mapmeasure_def
  | 
| 
 | 
  1491  | 
proof (subst emeasure_distr, subst measurable_eqI[OF s1 refl s2 refl], rule fm_measurable)
  | 
| 
 | 
  1492  | 
  have "X \<subseteq> space (Pi\<^isub>M J (\<lambda>_. N))" using assms by (simp add: sets_into_space)
  | 
| 
 | 
  1493  | 
  from assms inj_on_fm[of "\<lambda>_. N"] set_mp[OF this] have "fm -` fm ` X \<inter> space (Pi\<^isub>M J (\<lambda>_. N)) = X"
  | 
| 
 | 
  1494  | 
    by (auto simp: vimage_image_eq inj_on_def)
  | 
| 
 | 
  1495  | 
  thus "emeasure M X = emeasure M (fm -` fm ` X \<inter> space M)" using s1
  | 
| 
 | 
  1496  | 
    by simp
  | 
| 
 | 
  1497  | 
  show "fm ` X \<in> sets (PiF (Collect finite) (\<lambda>_. N))"
  | 
| 
 | 
  1498  | 
    by (rule fm_image_measurable_finite[OF N X[simplified s2]])
  | 
| 
 | 
  1499  | 
qed simp
  | 
| 
 | 
  1500  | 
  | 
| 
 | 
  1501  | 
end
  | 
| 
 | 
  1502  | 
  | 
| 
 | 
  1503  | 
end
  |