src/HOL/Probability/ex/Measure_Not_CCC.thy
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81179 cf2c03967178
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     1
(* Author: Johannes Hölzl <hoelzl@in.tum.de> *)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     2
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     3
section \<open>The Category of Measurable Spaces is not Cartesian Closed\<close>
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     4
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     5
theory Measure_Not_CCC
66992
69673025292e less global theories -- avoid confusion about special cases;
wenzelm
parents: 66453
diff changeset
     6
  imports "HOL-Probability.Probability"
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     7
begin
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     8
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
     9
text \<open>
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    10
  We show that the category of measurable spaces with measurable functions as morphisms is not a
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    11
  Cartesian closed category. While the category has products and terminal objects, the exponential
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    12
  does not exist for each pair of measurable spaces.
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    13
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    14
  We show that the exponential $\mathbb{B}^\mathbb{C}$ does not exist, where $\mathbb{B}$ is the
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    15
  discrete measurable space on boolean values, and $\mathbb{C}$ is the $\sigma$-algebra consisting
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    16
  of all countable and co-countable real sets. We also define $\mathbb{R}$ to be the discrete
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    17
  measurable space on the reals.
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    18
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69313
diff changeset
    19
  Now, the diagonal predicate \<^term>\<open>\<lambda>x y. x = y\<close> is $\mathbb{R}$-$\mathbb{B}^\mathbb{C}$-measurable,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69313
diff changeset
    20
  but \<^term>\<open>\<lambda>(x, y). x = y\<close> is not $(\mathbb{R} \times \mathbb{C})$-$\mathbb{B}$-measurable.
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    21
\<close>
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    22
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    23
definition COCOUNT :: "real measure" where
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    24
  "COCOUNT = sigma UNIV {{x} | x. True}"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    25
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    26
abbreviation POW :: "real measure" where
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    27
  "POW \<equiv> count_space UNIV"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    28
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    29
abbreviation BOOL :: "bool measure" where
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    30
  "BOOL \<equiv> count_space UNIV"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    31
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    32
lemma measurable_const_iff: "(\<lambda>x. c) \<in> measurable A B \<longleftrightarrow> (space A = {} \<or> c \<in> space B)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    33
  by (auto simp: measurable_def)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    34
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66992
diff changeset
    35
lemma measurable_eq[measurable]: "((=) x) \<in> measurable COCOUNT BOOL"
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    36
  unfolding pred_def by (auto simp: COCOUNT_def)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    37
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    38
lemma COCOUNT_eq: "A \<in> COCOUNT \<longleftrightarrow> countable A \<or> countable (UNIV - A)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    39
proof
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    40
  fix A assume "A \<in> COCOUNT"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    41
  then have "A \<in> sigma_sets UNIV {{x} | x. True}"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    42
    by (auto simp: COCOUNT_def)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    43
  then show "countable A \<or> countable (UNIV - A)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    44
  proof induction
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    45
    case (Union F)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    46
    moreover
81179
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    47
    have "countable (UNIV - (\<Union>i. F i))" if "countable (UNIV - F i)" for i
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    48
      using that by (rule countable_subset[rotated]) auto
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    49
    ultimately show "countable (\<Union>i. F i) \<or> countable (UNIV - (\<Union>i. F i))"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    50
      by blast
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    51
  qed (auto simp: Diff_Diff_Int)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    52
next
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    53
  assume "countable A \<or> countable (UNIV - A)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    54
  moreover
81179
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    55
  have A: "A \<in> COCOUNT" if "countable A" for A :: "real set"
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    56
  proof -
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    57
    have "A = (\<Union>a\<in>A. {a})"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    58
      by auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    59
    also have "\<dots> \<in> COCOUNT"
81179
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    60
      by (intro sets.countable_UN' that) (auto simp: COCOUNT_def) 
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    61
    finally show ?thesis .
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    62
  qed
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    63
  note A[of A]
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    64
  moreover
81179
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    65
  have "A \<in> COCOUNT" if "countable (UNIV - A)"
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    66
  proof -
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    67
    from A that have "space COCOUNT - (UNIV - A) \<in> COCOUNT" by simp
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    68
    then show ?thesis
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    69
      by (auto simp: COCOUNT_def Diff_Diff_Int)
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
    70
  qed
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    71
  ultimately show "A \<in> COCOUNT" 
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    72
    by blast
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    73
qed
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    74
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    75
lemma pair_COCOUNT:
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    76
  assumes A: "A \<in> sets (COCOUNT \<Otimes>\<^sub>M M)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    77
  shows "\<exists>J F X. X \<in> sets M \<and> F \<in> J \<rightarrow> sets M \<and> countable J \<and> A = (UNIV - J) \<times> X \<union> (SIGMA j:J. F j)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    78
  using A unfolding sets_pair_measure
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    79
proof induction
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    80
  case (Basic X)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    81
  then obtain a b where X: "X = a \<times> b" and b: "b \<in> sets M" and a: "countable a \<or> countable (UNIV - a)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    82
    by (auto simp: COCOUNT_eq)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    83
  from a show ?case
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    84
  proof 
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    85
    assume "countable a" with X b show ?thesis
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    86
      by (intro exI[of _ a] exI[of _ "\<lambda>_. b"] exI[of _ "{}"]) auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    87
  next
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    88
    assume "countable (UNIV - a)" with X b show ?thesis
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    89
      by (intro exI[of _ "UNIV - a"] exI[of _ "\<lambda>_. {}"] exI[of _ "b"]) auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    90
  qed
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    91
next
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    92
  case Empty then show ?case
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    93
    by (intro exI[of _ "{}"] exI[of _ "\<lambda>_. {}"] exI[of _ "{}"]) auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    94
next
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    95
  case (Compl A)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    96
  then obtain J F X where XFJ: "X \<in> sets M" "F \<in> J \<rightarrow> sets M" "countable J"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    97
    and A: "A = (UNIV - J) \<times> X \<union> Sigma J F"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    98
    by auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
    99
  have *: "space COCOUNT \<times> space M - A = (UNIV - J) \<times> (space M - X) \<union> (SIGMA j:J. space M - F j)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   100
    unfolding A by (auto simp: COCOUNT_def)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   101
  show ?case
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   102
    using XFJ unfolding *
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   103
    by (intro exI[of _ J] exI[of _ "space M - X"] exI[of _ "\<lambda>j. space M - F j"]) auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   104
next
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   105
  case (Union A)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   106
  obtain J F X where XFJ: "\<And>i. X i \<in> sets M" "\<And>i. F i \<in> J i \<rightarrow> sets M" "\<And>i. countable (J i)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   107
    and A_eq: "A = (\<lambda>i. (UNIV - J i) \<times> X i \<union> Sigma (J i) (F i))"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   108
    unfolding fun_eq_iff using Union.IH by metis
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   109
  show ?case
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   110
  proof (intro exI conjI)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
   111
    define G where "G j = (\<Union>i. if j \<in> J i then F i j else X i)" for j
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   112
    show "(\<Union>i. X i) \<in> sets M" "countable (\<Union>i. J i)" "G \<in> (\<Union>i. J i) \<rightarrow> sets M"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   113
      using XFJ by (auto simp: G_def Pi_iff)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 67399
diff changeset
   114
    show "\<Union>(A ` UNIV) = (UNIV - (\<Union>i. J i)) \<times> (\<Union>i. X i) \<union> (SIGMA j:\<Union>i. J i. \<Union>i. if j \<in> J i then F i j else X i)"
62390
842917225d56 more canonical names
nipkow
parents: 61808
diff changeset
   115
      unfolding A_eq by (auto split: if_split_asm)
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   116
  qed
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   117
qed
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   118
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   119
context
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   120
  fixes EXP :: "(real \<Rightarrow> bool) measure"
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 59144
diff changeset
   121
  assumes eq: "\<And>P. case_prod P \<in> measurable (POW \<Otimes>\<^sub>M COCOUNT) BOOL \<longleftrightarrow> P \<in> measurable POW EXP"
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   122
begin
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   123
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   124
lemma space_EXP: "space EXP = measurable COCOUNT BOOL"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   125
proof -
81179
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
   126
  have "f \<in> space EXP \<longleftrightarrow> f \<in> measurable COCOUNT BOOL" for f
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
   127
  proof -
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   128
    have "f \<in> space EXP \<longleftrightarrow> (\<lambda>(a, b). f b) \<in> measurable (POW \<Otimes>\<^sub>M COCOUNT) BOOL"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   129
      using eq[of "\<lambda>x. f"] by (simp add: measurable_const_iff)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   130
    also have "\<dots> \<longleftrightarrow> f \<in> measurable COCOUNT BOOL"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   131
      by auto
81179
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
   132
    finally show ?thesis .
cf2c03967178 tuned proofs;
wenzelm
parents: 69597
diff changeset
   133
  qed
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   134
  then show ?thesis by auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   135
qed
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   136
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   137
lemma measurable_eq_EXP: "(\<lambda>x y. x = y) \<in> measurable POW EXP"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   138
  unfolding measurable_def by (auto simp: space_EXP)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   139
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   140
lemma measurable_eq_pair: "(\<lambda>(y, x). x = y) \<in> measurable (COCOUNT \<Otimes>\<^sub>M POW) BOOL"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   141
  using measurable_eq_EXP unfolding eq[symmetric]
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   142
  by (subst measurable_pair_swap_iff) simp
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   143
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   144
lemma ce: False
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   145
proof -
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   146
  have "{(y, x) \<in> space (COCOUNT \<Otimes>\<^sub>M POW). x = y} \<in> sets (COCOUNT \<Otimes>\<^sub>M POW)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   147
    using measurable_eq_pair unfolding pred_def by (simp add: split_beta')
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   148
  also have "{(y, x) \<in> space (COCOUNT \<Otimes>\<^sub>M POW). x = y} = (SIGMA j:UNIV. {j})"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   149
    by (auto simp: space_pair_measure COCOUNT_def)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   150
  finally obtain X F J where "countable (J::real set)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   151
    and eq: "(SIGMA j:UNIV. {j}) = (UNIV - J) \<times> X \<union> (SIGMA j:J. F j)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   152
    using pair_COCOUNT[of "SIGMA j:UNIV. {j}" POW] by auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   153
  have X_single: "\<And>x. x \<notin> J \<Longrightarrow> X = {x}"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   154
    using eq[unfolded set_eq_iff] by force
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   155
  
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   156
  have "uncountable (UNIV - J)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61424
diff changeset
   157
    using \<open>countable J\<close> uncountable_UNIV_real uncountable_minus_countable by blast
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   158
  then have "infinite (UNIV - J)"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   159
    by (auto intro: countable_finite)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   160
  then have "\<exists>A. finite A \<and> card A = 2 \<and> A \<subseteq> UNIV - J"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   161
    by (rule infinite_arbitrarily_large)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   162
  then obtain i j where ij: "i \<in> UNIV - J" "j \<in> UNIV - J" "i \<noteq> j"
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   163
    by (auto simp add: card_Suc_eq numeral_2_eq_2)
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   164
  have "{(i, i), (j, j)} \<subseteq> (SIGMA j:UNIV. {j})" by auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   165
  with ij X_single[of i] X_single[of j] show False
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   166
    by auto
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   167
qed
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   168
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   169
end
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   170
61424
c3658c18b7bc prod_case as canonical name for product type eliminator
haftmann
parents: 59144
diff changeset
   171
corollary "\<not> (\<exists>EXP. \<forall>P. case_prod P \<in> measurable (POW \<Otimes>\<^sub>M COCOUNT) BOOL \<longleftrightarrow> P \<in> measurable POW EXP)"
59144
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   172
  using ce by blast
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   173
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   174
end
c9b75c03de3c unfortunately, there is no general function space in the measurable spaces
hoelzl
parents:
diff changeset
   175