src/HOL/Probability/Complete_Measure.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Wed, 13 Jul 2011 04:00:32 +0900
changeset 43799 a72661ba7239
parent 42866 b0746bd57a41
child 43920 cedb5cb948fd
permissions -rw-r--r--
merge
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41983
2dc6e382a58b standardized headers;
wenzelm
parents: 41981
diff changeset
     1
(*  Title:      HOL/Probability/Complete_Measure.thy
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
     2
    Author:     Robert Himmelmann, Johannes Hoelzl, TU Muenchen
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
     3
*)
41983
2dc6e382a58b standardized headers;
wenzelm
parents: 41981
diff changeset
     4
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
     5
theory Complete_Measure
42146
5b52c6a9c627 split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents: 41983
diff changeset
     6
imports Lebesgue_Integration
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
     7
begin
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
     8
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
     9
locale completeable_measure_space = measure_space
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    10
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    11
definition (in completeable_measure_space)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    12
  "split_completion A p = (\<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and>
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    13
    fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    14
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    15
definition (in completeable_measure_space)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    16
  "main_part A = fst (Eps (split_completion A))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    17
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    18
definition (in completeable_measure_space)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    19
  "null_part A = snd (Eps (split_completion A))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    20
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    21
abbreviation (in completeable_measure_space) "\<mu>' A \<equiv> \<mu> (main_part A)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    22
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    23
definition (in completeable_measure_space) completion :: "('a, 'b) measure_space_scheme" where
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    24
  "completion = \<lparr> space = space M,
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    25
                  sets = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N' },
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    26
                  measure = \<mu>',
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    27
                  \<dots> = more M \<rparr>"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
    28
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    29
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    30
lemma (in completeable_measure_space) space_completion[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    31
  "space completion = space M" unfolding completion_def by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    32
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    33
lemma (in completeable_measure_space) sets_completionE:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    34
  assumes "A \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    35
  obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets" "S \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    36
  using assms unfolding completion_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    37
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    38
lemma (in completeable_measure_space) sets_completionI:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    39
  assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets" "S \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    40
  shows "A \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    41
  using assms unfolding completion_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    42
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    43
lemma (in completeable_measure_space) sets_completionI_sets[intro]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    44
  "A \<in> sets M \<Longrightarrow> A \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    45
  unfolding completion_def by force
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    46
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    47
lemma (in completeable_measure_space) null_sets_completion:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    48
  assumes "N' \<in> null_sets" "N \<subseteq> N'" shows "N \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    49
  apply(rule sets_completionI[of N "{}" N N'])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    50
  using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    51
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    52
sublocale completeable_measure_space \<subseteq> completion!: sigma_algebra completion
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    53
proof (unfold sigma_algebra_iff2, safe)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    54
  fix A x assume "A \<in> sets completion" "x \<in> A"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    55
  with sets_into_space show "x \<in> space completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    56
    by (auto elim!: sets_completionE)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    57
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    58
  fix A assume "A \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    59
  from this[THEN sets_completionE] guess S N N' . note A = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    60
  let ?C = "space completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    61
  show "?C - A \<in> sets completion" using A
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    62
    by (intro sets_completionI[of _ "(?C - S) \<inter> (?C - N')" "(?C - S) \<inter> N' \<inter> (?C - N)"])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    63
       auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    64
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    65
  fix A ::"nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    66
  then have "\<forall>n. \<exists>S N N'. A n = S \<union> N \<and> S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N'"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    67
    unfolding completion_def by (auto simp: image_subset_iff)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    68
  from choice[OF this] guess S ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    69
  from choice[OF this] guess N ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    70
  from choice[OF this] guess N' ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    71
  then show "UNION UNIV A \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    72
    using null_sets_UN[of N']
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    73
    by (intro sets_completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    74
       auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    75
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    76
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    77
lemma (in completeable_measure_space) split_completion:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    78
  assumes "A \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    79
  shows "split_completion A (main_part A, null_part A)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    80
  unfolding main_part_def null_part_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    81
proof (rule someI2_ex)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    82
  from assms[THEN sets_completionE] guess S N N' . note A = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    83
  let ?P = "(S, N - S)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    84
  show "\<exists>p. split_completion A p"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    85
    unfolding split_completion_def using A
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    86
  proof (intro exI conjI)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    87
    show "A = fst ?P \<union> snd ?P" using A by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    88
    show "snd ?P \<subseteq> N'" using A by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    89
  qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    90
qed auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    91
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    92
lemma (in completeable_measure_space)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    93
  assumes "S \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    94
  shows main_part_sets[intro, simp]: "main_part S \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    95
    and main_part_null_part_Un[simp]: "main_part S \<union> null_part S = S"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    96
    and main_part_null_part_Int[simp]: "main_part S \<inter> null_part S = {}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    97
  using split_completion[OF assms] by (auto simp: split_completion_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    98
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
    99
lemma (in completeable_measure_space) null_part:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   100
  assumes "S \<in> sets completion" shows "\<exists>N. N\<in>null_sets \<and> null_part S \<subseteq> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   101
  using split_completion[OF assms] by (auto simp: split_completion_def)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   102
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   103
lemma (in completeable_measure_space) null_part_sets[intro, simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   104
  assumes "S \<in> sets M" shows "null_part S \<in> sets M" "\<mu> (null_part S) = 0"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   105
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   106
  have S: "S \<in> sets completion" using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   107
  have "S - main_part S \<in> sets M" using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   108
  moreover
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   109
  from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   110
  have "S - main_part S = null_part S" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   111
  ultimately show sets: "null_part S \<in> sets M" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   112
  from null_part[OF S] guess N ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   113
  with measure_eq_0[of N "null_part S"] sets
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   114
  show "\<mu> (null_part S) = 0" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   115
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   116
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   117
lemma (in completeable_measure_space) \<mu>'_set[simp]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   118
  assumes "S \<in> sets M" shows "\<mu>' S = \<mu> S"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   119
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   120
  have S: "S \<in> sets completion" using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   121
  then have "\<mu> S = \<mu> (main_part S \<union> null_part S)" by simp
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   122
  also have "\<dots> = \<mu>' S"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   123
    using S assms measure_additive[of "main_part S" "null_part S"]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   124
    by (auto simp: measure_additive)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   125
  finally show ?thesis by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   126
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   127
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   128
lemma (in completeable_measure_space) sets_completionI_sub:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   129
  assumes N: "N' \<in> null_sets" "N \<subseteq> N'"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   130
  shows "N \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   131
  using assms by (intro sets_completionI[of _ "{}" N N']) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   132
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   133
lemma (in completeable_measure_space) \<mu>_main_part_UN:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   134
  fixes S :: "nat \<Rightarrow> 'a set"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   135
  assumes "range S \<subseteq> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   136
  shows "\<mu>' (\<Union>i. (S i)) = \<mu> (\<Union>i. main_part (S i))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   137
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   138
  have S: "\<And>i. S i \<in> sets completion" using assms by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   139
  then have UN: "(\<Union>i. S i) \<in> sets completion" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   140
  have "\<forall>i. \<exists>N. N \<in> null_sets \<and> null_part (S i) \<subseteq> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   141
    using null_part[OF S] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   142
  from choice[OF this] guess N .. note N = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   143
  then have UN_N: "(\<Union>i. N i) \<in> null_sets" by (intro null_sets_UN) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   144
  have "(\<Union>i. S i) \<in> sets completion" using S by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   145
  from null_part[OF this] guess N' .. note N' = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   146
  let ?N = "(\<Union>i. N i) \<union> N'"
42866
b0746bd57a41 the measurable sets with null measure form a ring
hoelzl
parents: 42146
diff changeset
   147
  have null_set: "?N \<in> null_sets" using N' UN_N by (intro nullsets.Un) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   148
  have "main_part (\<Union>i. S i) \<union> ?N = (main_part (\<Union>i. S i) \<union> null_part (\<Union>i. S i)) \<union> ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   149
    using N' by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   150
  also have "\<dots> = (\<Union>i. main_part (S i) \<union> null_part (S i)) \<union> ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   151
    unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   152
  also have "\<dots> = (\<Union>i. main_part (S i)) \<union> ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   153
    using N by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   154
  finally have *: "main_part (\<Union>i. S i) \<union> ?N = (\<Union>i. main_part (S i)) \<union> ?N" .
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   155
  have "\<mu> (main_part (\<Union>i. S i)) = \<mu> (main_part (\<Union>i. S i) \<union> ?N)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   156
    using null_set UN by (intro measure_Un_null_set[symmetric]) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   157
  also have "\<dots> = \<mu> ((\<Union>i. main_part (S i)) \<union> ?N)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   158
    unfolding * ..
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   159
  also have "\<dots> = \<mu> (\<Union>i. main_part (S i))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   160
    using null_set S by (intro measure_Un_null_set) auto
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   161
  finally show ?thesis .
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   162
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   163
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   164
lemma (in completeable_measure_space) \<mu>_main_part_Un:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   165
  assumes S: "S \<in> sets completion" and T: "T \<in> sets completion"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   166
  shows "\<mu>' (S \<union> T) = \<mu> (main_part S \<union> main_part T)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   167
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   168
  have UN: "(\<Union>i. binary (main_part S) (main_part T) i) = (\<Union>i. main_part (binary S T i))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   169
    unfolding binary_def by (auto split: split_if_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   170
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   171
    using \<mu>_main_part_UN[of "binary S T"] assms
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   172
    unfolding range_binary_eq Un_range_binary UN by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   173
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   174
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   175
sublocale completeable_measure_space \<subseteq> completion!: measure_space completion
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   176
  where "measure completion = \<mu>'"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   177
proof -
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   178
  show "measure_space completion"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   179
  proof
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   180
    show "positive completion (measure completion)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   181
      by (auto simp: completion_def positive_def)
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   182
  next
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   183
    show "countably_additive completion (measure completion)"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   184
    proof (intro countably_additiveI)
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   185
      fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion" "disjoint_family A"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   186
      have "disjoint_family (\<lambda>i. main_part (A i))"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   187
      proof (intro disjoint_family_on_bisimulation[OF A(2)])
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   188
        fix n m assume "A n \<inter> A m = {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   189
        then have "(main_part (A n) \<union> null_part (A n)) \<inter> (main_part (A m) \<union> null_part (A m)) = {}"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   190
          using A by (subst (1 2) main_part_null_part_Un) auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   191
        then show "main_part (A n) \<inter> main_part (A m) = {}" by auto
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   192
      qed
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   193
      then have "(\<Sum>n. measure completion (A n)) = \<mu> (\<Union>i. main_part (A i))"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   194
        unfolding completion_def using A by (auto intro!: measure_countably_additive)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   195
      then show "(\<Sum>n. measure completion (A n)) = measure completion (UNION UNIV A)"
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   196
        by (simp add: completion_def \<mu>_main_part_UN[OF A(1)])
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   197
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   198
  qed
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   199
  show "measure completion = \<mu>'" unfolding completion_def by simp
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   200
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   201
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   202
lemma (in completeable_measure_space) completion_ex_simple_function:
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   203
  assumes f: "simple_function completion f"
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   204
  shows "\<exists>f'. simple_function M f' \<and> (AE x. f x = f' x)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   205
proof -
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   206
  let "?F x" = "f -` {x} \<inter> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   207
  have F: "\<And>x. ?F x \<in> sets completion" and fin: "finite (f`space M)"
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   208
    using completion.simple_functionD[OF f]
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   209
      completion.simple_functionD[OF f] by simp_all
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   210
  have "\<forall>x. \<exists>N. N \<in> null_sets \<and> null_part (?F x) \<subseteq> N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   211
    using F null_part by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   212
  from choice[OF this] obtain N where
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   213
    N: "\<And>x. null_part (?F x) \<subseteq> N x" "\<And>x. N x \<in> null_sets" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   214
  let ?N = "\<Union>x\<in>f`space M. N x" let "?f' x" = "if x \<in> ?N then undefined else f x"
42866
b0746bd57a41 the measurable sets with null measure form a ring
hoelzl
parents: 42146
diff changeset
   215
  have sets: "?N \<in> null_sets" using N fin by (intro nullsets.finite_UN) auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   216
  show ?thesis unfolding simple_function_def
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   217
  proof (safe intro!: exI[of _ ?f'])
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   218
    have "?f' ` space M \<subseteq> f`space M \<union> {undefined}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   219
    from finite_subset[OF this] completion.simple_functionD(1)[OF f]
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   220
    show "finite (?f' ` space M)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   221
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   222
    fix x assume "x \<in> space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   223
    have "?f' -` {?f' x} \<inter> space M =
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   224
      (if x \<in> ?N then ?F undefined \<union> ?N
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   225
       else if f x = undefined then ?F (f x) \<union> ?N
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   226
       else ?F (f x) - ?N)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   227
      using N(2) sets_into_space by (auto split: split_if_asm)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   228
    moreover { fix y have "?F y \<union> ?N \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   229
      proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   230
        assume y: "y \<in> f`space M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   231
        have "?F y \<union> ?N = (main_part (?F y) \<union> null_part (?F y)) \<union> ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   232
          using main_part_null_part_Un[OF F] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   233
        also have "\<dots> = main_part (?F y) \<union> ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   234
          using y N by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   235
        finally show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   236
          using F sets by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   237
      next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   238
        assume "y \<notin> f`space M" then have "?F y = {}" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   239
        then show ?thesis using sets by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   240
      qed }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   241
    moreover {
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   242
      have "?F (f x) - ?N = main_part (?F (f x)) \<union> null_part (?F (f x)) - ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   243
        using main_part_null_part_Un[OF F] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   244
      also have "\<dots> = main_part (?F (f x)) - ?N"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   245
        using N `x \<in> space M` by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   246
      finally have "?F (f x) - ?N \<in> sets M"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   247
        using F sets by auto }
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   248
    ultimately show "?f' -` {?f' x} \<inter> space M \<in> sets M" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   249
  next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   250
    show "AE x. f x = ?f' x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   251
      by (rule AE_I', rule sets) auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   252
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   253
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   254
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   255
lemma (in completeable_measure_space) completion_ex_borel_measurable_pos:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   256
  fixes g :: "'a \<Rightarrow> extreal"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   257
  assumes g: "g \<in> borel_measurable completion" and "\<And>x. 0 \<le> g x"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   258
  shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   259
proof -
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   260
  from g[THEN completion.borel_measurable_implies_simple_function_sequence'] guess f . note f = this
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   261
  from this(1)[THEN completion_ex_simple_function]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   262
  have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x. f i x = f' x)" ..
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   263
  from this[THEN choice] obtain f' where
41689
3e39b0e730d6 the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents: 41097
diff changeset
   264
    sf: "\<And>i. simple_function M (f' i)" and
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   265
    AE: "\<forall>i. AE x. f i x = f' i x" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   266
  show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   267
  proof (intro bexI)
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   268
    from AE[unfolded AE_all_countable[symmetric]]
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41023
diff changeset
   269
    show "AE x. g x = (SUP i. f' i x)" (is "AE x. g x = ?f x")
41705
1100512e16d8 add auto support for AE_mp
hoelzl
parents: 41689
diff changeset
   270
    proof (elim AE_mp, safe intro!: AE_I2)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   271
      fix x assume eq: "\<forall>i. f i x = f' i x"
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   272
      moreover have "g x = (SUP i. f i x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   273
        unfolding f using `0 \<le> g x` by (auto split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   274
      ultimately show "g x = ?f x" by auto
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   275
    qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   276
    show "?f \<in> borel_measurable M"
41097
a1abfa4e2b44 use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents: 41023
diff changeset
   277
      using sf by (auto intro: borel_measurable_simple_function)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   278
  qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   279
qed
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   280
41981
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   281
lemma (in completeable_measure_space) completion_ex_borel_measurable:
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   282
  fixes g :: "'a \<Rightarrow> extreal"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   283
  assumes g: "g \<in> borel_measurable completion"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   284
  shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   285
proof -
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   286
  have "(\<lambda>x. max 0 (g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (g x)" using g by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   287
  from completion_ex_borel_measurable_pos[OF this] guess g_pos ..
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   288
  moreover
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   289
  have "(\<lambda>x. max 0 (- g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (- g x)" using g by auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   290
  from completion_ex_borel_measurable_pos[OF this] guess g_neg ..
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   291
  ultimately
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   292
  show ?thesis
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   293
  proof (safe intro!: bexI[of _ "\<lambda>x. g_pos x - g_neg x"])
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   294
    show "AE x. max 0 (- g x) = g_neg x \<longrightarrow> max 0 (g x) = g_pos x \<longrightarrow> g x = g_pos x - g_neg x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   295
    proof (intro AE_I2 impI)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   296
      fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x"
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   297
      show "g x = g_pos x - g_neg x" unfolding g[symmetric]
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   298
        by (cases "g x") (auto split: split_max)
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   299
    qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   300
  qed auto
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   301
qed
cdf7693bbe08 reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents: 41959
diff changeset
   302
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents:
diff changeset
   303
end