src/HOL/Analysis/Measurable.thy
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 78801 42ae6e0ecfd4
child 82513 8281047b896d
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63627
6ddb43c6b711 rename HOL-Multivariate_Analysis to HOL-Analysis.
hoelzl
parents: 63626
diff changeset
     1
(*  Title:      HOL/Analysis/Measurable.thy
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     2
    Author:     Johannes Hölzl <hoelzl@in.tum.de>
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     3
*)
69517
dc20f278e8f3 tuned style and headers
nipkow
parents: 69260
diff changeset
     4
section \<open>Measurability Prover\<close>
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     5
theory Measurable
56021
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
     6
  imports
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
     7
    Sigma_Algebra
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 64320
diff changeset
     8
    "HOL-Library.Order_Continuity"
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
     9
begin
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    10
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    11
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    12
lemma (in algebra) sets_Collect_finite_All:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    13
  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    14
  shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    15
proof -
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    16
  have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (if S = {} then \<Omega> else \<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    17
    by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    18
  with assms show ?thesis by (auto intro!: sets_Collect_finite_All')
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    19
qed
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    20
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    21
abbreviation "pred M P \<equiv> P \<in> measurable M (count_space (UNIV::bool set))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    22
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    23
lemma pred_def: "pred M P \<longleftrightarrow> {x\<in>space M. P x} \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    24
proof
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    25
  assume "pred M P"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    26
  then have "P -` {True} \<inter> space M \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    27
    by (auto simp: measurable_count_space_eq2)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    28
  also have "P -` {True} \<inter> space M = {x\<in>space M. P x}" by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    29
  finally show "{x\<in>space M. P x} \<in> sets M" .
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    30
next
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    31
  assume P: "{x\<in>space M. P x} \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    32
  moreover
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    33
  { fix X
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    34
    have "X \<in> Pow (UNIV :: bool set)" by simp
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    35
    then have "P -` X \<inter> space M = {x\<in>space M. ((X = {True} \<longrightarrow> P x) \<and> (X = {False} \<longrightarrow> \<not> P x) \<and> X \<noteq> {})}"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    36
      unfolding UNIV_bool Pow_insert Pow_empty by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    37
    then have "P -` X \<inter> space M \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    38
      by (auto intro!: sets.sets_Collect_neg sets.sets_Collect_imp sets.sets_Collect_conj sets.sets_Collect_const P) }
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    39
  then show "pred M P"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    40
    by (auto simp: measurable_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    41
qed
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    42
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    43
lemma pred_sets1: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> f \<in> measurable N M \<Longrightarrow> pred N (\<lambda>x. P (f x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    44
  by (rule measurable_compose[where f=f and N=M]) (auto simp: pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    45
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    46
lemma pred_sets2: "A \<in> sets N \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    47
  by (rule measurable_compose[where f=f and N=N]) (auto simp: pred_def Int_def[symmetric])
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    48
69605
a96320074298 isabelle update -u path_cartouches;
wenzelm
parents: 69597
diff changeset
    49
ML_file \<open>measurable.ML\<close>
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    50
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
    51
attribute_setup measurable = \<open>
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    52
  Scan.lift (
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    53
    (Args.add >> K true || Args.del >> K false || Scan.succeed true) --
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    54
    Scan.optional (Args.parens (
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    55
      Scan.optional (Args.$$$ "raw" >> K true) false --
58965
a62cdcc5344b add del option to measurable;
Andreas Lochbihler
parents: 57025
diff changeset
    56
      Scan.optional (Args.$$$ "generic" >> K Measurable.Generic) Measurable.Concrete))
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    57
    (false, Measurable.Concrete) >>
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    58
    Measurable.measurable_thm_attr)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
    59
\<close> "declaration of measurability theorems"
53043
8cbfbeb566a4 more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
wenzelm
parents: 50530
diff changeset
    60
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    61
attribute_setup measurable_dest = Measurable.dest_thm_attr
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    62
  "add dest rule to measurability prover"
53043
8cbfbeb566a4 more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
wenzelm
parents: 50530
diff changeset
    63
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    64
attribute_setup measurable_cong = Measurable.cong_thm_attr
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    65
  "add congurence rules to measurability prover"
53043
8cbfbeb566a4 more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
wenzelm
parents: 50530
diff changeset
    66
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69861
diff changeset
    67
method_setup\<^marker>\<open>tag important\<close> measurable = \<open> Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \<close>
59047
8d7cec9b861d cleanup measurability prover
hoelzl
parents: 59000
diff changeset
    68
  "measurability prover"
53043
8cbfbeb566a4 more standard attribute_setup / method_setup -- export key ML operations instead of parsers;
wenzelm
parents: 50530
diff changeset
    69
78801
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 71834
diff changeset
    70
simproc_setup\<^marker>\<open>tag important\<close> measurable ("A \<in> sets M" | "f \<in> measurable M N") =
42ae6e0ecfd4 tuned signature;
wenzelm
parents: 71834
diff changeset
    71
  \<open>K Measurable.proc\<close>
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    72
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
    73
setup \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69517
diff changeset
    74
  Global_Theory.add_thms_dynamic (\<^binding>\<open>measurable\<close>, Measurable.get_all)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
    75
\<close>
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    76
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    77
declare
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    78
  pred_sets1[measurable_dest]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    79
  pred_sets2[measurable_dest]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    80
  sets.sets_into_space[measurable_dest]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    81
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    82
declare
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    83
  sets.top[measurable]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    84
  sets.empty_sets[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    85
  sets.Un[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    86
  sets.Diff[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    87
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    88
declare
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    89
  measurable_count_space[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    90
  measurable_ident[measurable (raw)]
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    91
  measurable_id[measurable (raw)]
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    92
  measurable_const[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    93
  measurable_If[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    94
  measurable_comp[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    95
  measurable_sets[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
    96
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    97
declare measurable_cong_sets[measurable_cong]
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
    98
declare sets_restrict_space_cong[measurable_cong]
59361
fd5da2434be4 piecewise measurability using restrict_space; cleanup Borel_Space
hoelzl
parents: 59353
diff changeset
    99
declare sets_restrict_UNIV[measurable_cong]
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   100
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   101
lemma predE[measurable (raw)]:
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   102
  "pred M P \<Longrightarrow> {x\<in>space M. P x} \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   103
  unfolding pred_def .
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   104
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   105
lemma pred_intros_imp'[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   106
  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<longrightarrow> P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   107
  by (cases K) auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   108
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   109
lemma pred_intros_conj1'[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   110
  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<and> P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   111
  by (cases K) auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   112
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   113
lemma pred_intros_conj2'[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   114
  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<and> K)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   115
  by (cases K) auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   116
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   117
lemma pred_intros_disj1'[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   118
  "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<or> P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   119
  by (cases K) auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   120
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   121
lemma pred_intros_disj2'[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   122
  "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<or> K)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   123
  by (cases K) auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   124
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   125
lemma pred_intros_logic[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   126
  "pred M (\<lambda>x. x \<in> space M)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   127
  "pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. \<not> P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   128
  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<and> P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   129
  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<longrightarrow> P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   130
  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<or> P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   131
  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x = P x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   132
  "pred M (\<lambda>x. f x \<in> UNIV)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   133
  "pred M (\<lambda>x. f x \<in> {})"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   134
  "pred M (\<lambda>x. P' (f x) x) \<Longrightarrow> pred M (\<lambda>x. f x \<in> {y. P' y x})"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   135
  "pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> - (B x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   136
  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) - (B x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   137
  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<inter> (B x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   138
  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<union> (B x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   139
  "pred M (\<lambda>x. g x (f x) \<in> (X x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (g x) -` (X x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   140
  by (auto simp: iff_conv_conj_imp pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   141
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   142
lemma pred_intros_countable[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   143
  fixes P :: "'a \<Rightarrow> 'i :: countable \<Rightarrow> bool"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   144
  shows
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   145
    "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i. P x i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   146
    "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i. P x i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   147
  by (auto intro!: sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex simp: pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   148
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   149
lemma pred_intros_countable_bounded[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   150
  fixes X :: "'i :: countable set"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   151
  shows
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   152
    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>X. N x i))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   153
    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>X. N x i))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   154
    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   155
    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>X. P x i)"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 61808
diff changeset
   156
  by simp_all (auto simp: Bex_def Ball_def)
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   157
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   158
lemma pred_intros_finite[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   159
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>I. N x i))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   160
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>I. N x i))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   161
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>I. P x i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   162
  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>I. P x i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   163
  by (auto intro!: sets.sets_Collect_finite_Ex sets.sets_Collect_finite_All simp: iff_conv_conj_imp pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   164
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   165
lemma countable_Un_Int[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   166
  "(\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Union>i\<in>I. N i) \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   167
  "I \<noteq> {} \<Longrightarrow> (\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Inter>i\<in>I. N i) \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   168
  by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   169
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   170
declare
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   171
  finite_UN[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   172
  finite_INT[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   173
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   174
lemma sets_Int_pred[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   175
  assumes space: "A \<inter> B \<subseteq> space M" and [measurable]: "pred M (\<lambda>x. x \<in> A)" "pred M (\<lambda>x. x \<in> B)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   176
  shows "A \<inter> B \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   177
proof -
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   178
  have "{x\<in>space M. x \<in> A \<inter> B} \<in> sets M" by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   179
  also have "{x\<in>space M. x \<in> A \<inter> B} = A \<inter> B"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   180
    using space by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   181
  finally show ?thesis .
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   182
qed
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   183
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   184
lemma [measurable (raw generic)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   185
  assumes f: "f \<in> measurable M N" and c: "c \<in> space N \<Longrightarrow> {c} \<in> sets N"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   186
  shows pred_eq_const1: "pred M (\<lambda>x. f x = c)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   187
    and pred_eq_const2: "pred M (\<lambda>x. c = f x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   188
proof -
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   189
  show "pred M (\<lambda>x. f x = c)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   190
  proof cases
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   191
    assume "c \<in> space N"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   192
    with measurable_sets[OF f c] show ?thesis
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   193
      by (auto simp: Int_def conj_commute pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   194
  next
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   195
    assume "c \<notin> space N"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   196
    with f[THEN measurable_space] have "{x \<in> space M. f x = c} = {}" by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   197
    then show ?thesis by (auto simp: pred_def cong: conj_cong)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   198
  qed
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   199
  then show "pred M (\<lambda>x. c = f x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   200
    by (simp add: eq_commute)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   201
qed
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   202
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   203
lemma pred_count_space_const1[measurable (raw)]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   204
  "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. f x = c)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   205
  by (intro pred_eq_const1[where N="count_space UNIV"]) (auto )
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   206
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   207
lemma pred_count_space_const2[measurable (raw)]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   208
  "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. c = f x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   209
  by (intro pred_eq_const2[where N="count_space UNIV"]) (auto )
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   210
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   211
lemma pred_le_const[measurable (raw generic)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   212
  assumes f: "f \<in> measurable M N" and c: "{.. c} \<in> sets N" shows "pred M (\<lambda>x. f x \<le> c)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   213
  using measurable_sets[OF f c]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   214
  by (auto simp: Int_def conj_commute eq_commute pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   215
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   216
lemma pred_const_le[measurable (raw generic)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   217
  assumes f: "f \<in> measurable M N" and c: "{c ..} \<in> sets N" shows "pred M (\<lambda>x. c \<le> f x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   218
  using measurable_sets[OF f c]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   219
  by (auto simp: Int_def conj_commute eq_commute pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   220
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   221
lemma pred_less_const[measurable (raw generic)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   222
  assumes f: "f \<in> measurable M N" and c: "{..< c} \<in> sets N" shows "pred M (\<lambda>x. f x < c)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   223
  using measurable_sets[OF f c]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   224
  by (auto simp: Int_def conj_commute eq_commute pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   225
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   226
lemma pred_const_less[measurable (raw generic)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   227
  assumes f: "f \<in> measurable M N" and c: "{c <..} \<in> sets N" shows "pred M (\<lambda>x. c < f x)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   228
  using measurable_sets[OF f c]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   229
  by (auto simp: Int_def conj_commute eq_commute pred_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   230
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   231
declare
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   232
  sets.Int[measurable (raw)]
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   233
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   234
lemma pred_in_If[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   235
  "(P \<Longrightarrow> pred M (\<lambda>x. x \<in> A x)) \<Longrightarrow> (\<not> P \<Longrightarrow> pred M (\<lambda>x. x \<in> B x)) \<Longrightarrow>
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   236
    pred M (\<lambda>x. x \<in> (if P then A x else B x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   237
  by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   238
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   239
lemma sets_range[measurable_dest]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   240
  "A ` I \<subseteq> sets M \<Longrightarrow> i \<in> I \<Longrightarrow> A i \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   241
  by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   242
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   243
lemma pred_sets_range[measurable_dest]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   244
  "A ` I \<subseteq> sets N \<Longrightarrow> i \<in> I \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   245
  using pred_sets2[OF sets_range] by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   246
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   247
lemma sets_All[measurable_dest]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   248
  "\<forall>i. A i \<in> sets (M i) \<Longrightarrow> A i \<in> sets (M i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   249
  by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   250
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   251
lemma pred_sets_All[measurable_dest]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   252
  "\<forall>i. A i \<in> sets (N i) \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   253
  using pred_sets2[OF sets_All, of A N f] by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   254
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   255
lemma sets_Ball[measurable_dest]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   256
  "\<forall>i\<in>I. A i \<in> sets (M i) \<Longrightarrow> i\<in>I \<Longrightarrow> A i \<in> sets (M i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   257
  by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   258
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   259
lemma pred_sets_Ball[measurable_dest]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   260
  "\<forall>i\<in>I. A i \<in> sets (N i) \<Longrightarrow> i\<in>I \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   261
  using pred_sets2[OF sets_Ball, of _ _ _ f] by auto
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   262
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   263
lemma measurable_finite[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   264
  fixes S :: "'a \<Rightarrow> nat set"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   265
  assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   266
  shows "pred M (\<lambda>x. finite (S x))"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   267
  unfolding finite_nat_set_iff_bounded by (simp add: Ball_def)
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   268
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   269
lemma measurable_Least[measurable]:
63333
158ab2239496 Probability: show that measures form a complete lattice
hoelzl
parents: 63040
diff changeset
   270
  assumes [measurable]: "(\<And>i::nat. (\<lambda>x. P i x) \<in> measurable M (count_space UNIV))"
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   271
  shows "(\<lambda>x. LEAST i. P i x) \<in> measurable M (count_space UNIV)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   272
  unfolding measurable_def by (safe intro!: sets_Least) simp_all
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   273
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   274
lemma measurable_Max_nat[measurable (raw)]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   275
  fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   276
  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   277
  shows "(\<lambda>x. Max {i. P i x}) \<in> measurable M (count_space UNIV)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   278
  unfolding measurable_count_space_eq2_countable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   279
proof safe
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   280
  fix n
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   281
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   282
  { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   283
    then have "infinite {i. P i x}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   284
      unfolding infinite_nat_iff_unbounded_le by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   285
    then have "Max {i. P i x} = the None"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   286
      by (rule Max.infinite) }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   287
  note 1 = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   288
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   289
  { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   290
    then have "finite {i. P i x}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   291
      by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   292
    with \<open>P i x\<close> have "P (Max {i. P i x}) x" "i \<le> Max {i. P i x}" "finite {i. P i x}"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   293
      using Max_in[of "{i. P i x}"] by auto }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   294
  note 2 = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   295
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   296
  have "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Max {i. P i x} = n}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   297
    by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   298
  also have "\<dots> =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   299
    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   300
      if (\<exists>i. P i x) then P n x \<and> (\<forall>i>n. \<not> P i x)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   301
      else Max {} = n}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   302
    by (intro arg_cong[where f=Collect] ext conj_cong)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   303
       (auto simp add: 1 2 not_le[symmetric] intro!: Max_eqI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   304
  also have "\<dots> \<in> sets M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   305
    by measurable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   306
  finally show "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   307
qed simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   308
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   309
lemma measurable_Min_nat[measurable (raw)]:
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   310
  fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   311
  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   312
  shows "(\<lambda>x. Min {i. P i x}) \<in> measurable M (count_space UNIV)"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   313
  unfolding measurable_count_space_eq2_countable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   314
proof safe
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   315
  fix n
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   316
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   317
  { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   318
    then have "infinite {i. P i x}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   319
      unfolding infinite_nat_iff_unbounded_le by auto
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   320
    then have "Min {i. P i x} = the None"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   321
      by (rule Min.infinite) }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   322
  note 1 = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   323
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   324
  { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   325
    then have "finite {i. P i x}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   326
      by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   327
    with \<open>P i x\<close> have "P (Min {i. P i x}) x" "Min {i. P i x} \<le> i" "finite {i. P i x}"
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   328
      using Min_in[of "{i. P i x}"] by auto }
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   329
  note 2 = this
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   330
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   331
  have "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Min {i. P i x} = n}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   332
    by auto
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   333
  also have "\<dots> =
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   334
    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
56993
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   335
      if (\<exists>i. P i x) then P n x \<and> (\<forall>i<n. \<not> P i x)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   336
      else Min {} = n}"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   337
    by (intro arg_cong[where f=Collect] ext conj_cong)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   338
       (auto simp add: 1 2 not_le[symmetric] intro!: Min_eqI)
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   339
  also have "\<dots> \<in> sets M"
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   340
    by measurable
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   341
  finally show "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   342
qed simp
e5366291d6aa introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents: 56045
diff changeset
   343
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   344
lemma measurable_count_space_insert[measurable (raw)]:
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   345
  "s \<in> S \<Longrightarrow> A \<in> sets (count_space S) \<Longrightarrow> insert s A \<in> sets (count_space S)"
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   346
  by simp
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   347
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   348
lemma sets_UNIV [measurable (raw)]: "A \<in> sets (count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   349
  by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   350
57025
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   351
lemma measurable_card[measurable]:
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   352
  fixes S :: "'a \<Rightarrow> nat set"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   353
  assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   354
  shows "(\<lambda>x. card (S x)) \<in> measurable M (count_space UNIV)"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   355
  unfolding measurable_count_space_eq2_countable
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   356
proof safe
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   357
  fix n show "(\<lambda>x. card (S x)) -` {n} \<inter> space M \<in> sets M"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   358
  proof (cases n)
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   359
    case 0
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   360
    then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M = {x\<in>space M. infinite (S x) \<or> (\<forall>i. i \<notin> S x)}"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   361
      by auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   362
    also have "\<dots> \<in> sets M"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   363
      by measurable
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   364
    finally show ?thesis .
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   365
  next
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   366
    case (Suc i)
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   367
    then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M =
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   368
      (\<Union>F\<in>{A\<in>{A. finite A}. card A = n}. {x\<in>space M. (\<forall>i. i \<in> S x \<longleftrightarrow> i \<in> F)})"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   369
      unfolding set_eq_iff[symmetric] Collect_bex_eq[symmetric] by (auto intro: card_ge_0_finite)
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   370
    also have "\<dots> \<in> sets M"
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   371
      by (intro sets.countable_UN' countable_Collect countable_Collect_finite) auto
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   372
    finally show ?thesis .
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   373
  qed
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   374
qed rule
e7fd64f82876 add various lemmas
hoelzl
parents: 56993
diff changeset
   375
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   376
lemma measurable_pred_countable[measurable (raw)]:
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   377
  assumes "countable X"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   378
  shows
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   379
    "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   380
    "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<exists>i\<in>X. P x i)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   381
  unfolding pred_def
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   382
  by (auto intro!: sets.sets_Collect_countable_All' sets.sets_Collect_countable_Ex' assms)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   383
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69861
diff changeset
   384
subsection\<^marker>\<open>tag unimportant\<close> \<open>Measurability for (co)inductive predicates\<close>
56021
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   385
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   386
lemma measurable_bot[measurable]: "bot \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   387
  by (simp add: bot_fun_def)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   388
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   389
lemma measurable_top[measurable]: "top \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   390
  by (simp add: top_fun_def)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   391
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   392
lemma measurable_SUP[measurable]:
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   393
  fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   394
  assumes [simp]: "countable I"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   395
  assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)"
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67962
diff changeset
   396
  shows "(\<lambda>x. SUP i\<in>I. F i x) \<in> measurable M (count_space UNIV)"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   397
  unfolding measurable_count_space_eq2_countable
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   398
proof (safe intro!: UNIV_I)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   399
  fix a
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67962
diff changeset
   400
  have "(\<lambda>x. SUP i\<in>I. F i x) -` {a} \<inter> space M =
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   401
    {x\<in>space M. (\<forall>i\<in>I. F i x \<le> a) \<and> (\<forall>b. (\<forall>i\<in>I. F i x \<le> b) \<longrightarrow> a \<le> b)}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   402
    unfolding SUP_le_iff[symmetric] by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   403
  also have "\<dots> \<in> sets M"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   404
    by measurable
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67962
diff changeset
   405
  finally show "(\<lambda>x. SUP i\<in>I. F i x) -` {a} \<inter> space M \<in> sets M" .
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   406
qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   407
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   408
lemma measurable_INF[measurable]:
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   409
  fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   410
  assumes [simp]: "countable I"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   411
  assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)"
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67962
diff changeset
   412
  shows "(\<lambda>x. INF i\<in>I. F i x) \<in> measurable M (count_space UNIV)"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   413
  unfolding measurable_count_space_eq2_countable
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   414
proof (safe intro!: UNIV_I)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   415
  fix a
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67962
diff changeset
   416
  have "(\<lambda>x. INF i\<in>I. F i x) -` {a} \<inter> space M =
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   417
    {x\<in>space M. (\<forall>i\<in>I. a \<le> F i x) \<and> (\<forall>b. (\<forall>i\<in>I. b \<le> F i x) \<longrightarrow> b \<le> a)}"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   418
    unfolding le_INF_iff[symmetric] by auto
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   419
  also have "\<dots> \<in> sets M"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   420
    by measurable
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 67962
diff changeset
   421
  finally show "(\<lambda>x. INF i\<in>I. F i x) -` {a} \<inter> space M \<in> sets M" .
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   422
qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   423
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   424
lemma measurable_lfp_coinduct[consumes 1, case_names continuity step]:
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   425
  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   426
  assumes "P M"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 59361
diff changeset
   427
  assumes F: "sup_continuous F"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   428
  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   429
  shows "lfp F \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   430
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   431
  { fix i from \<open>P M\<close> have "((F ^^ i) bot) \<in> measurable M (count_space UNIV)"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   432
      by (induct i arbitrary: M) (auto intro!: *) }
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   433
  then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   434
    by measurable
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   435
  also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = lfp F"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69605
diff changeset
   436
    by (subst sup_continuous_lfp) (auto intro: F simp: image_comp)
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   437
  finally show ?thesis .
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   438
qed
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   439
56021
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   440
lemma measurable_lfp:
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   441
  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 59361
diff changeset
   442
  assumes F: "sup_continuous F"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   443
  assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   444
  shows "lfp F \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   445
  by (coinduction rule: measurable_lfp_coinduct[OF _ F]) (blast intro: *)
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   446
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   447
lemma measurable_gfp_coinduct[consumes 1, case_names continuity step]:
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   448
  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   449
  assumes "P M"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 59361
diff changeset
   450
  assumes F: "inf_continuous F"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   451
  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   452
  shows "gfp F \<in> measurable M (count_space UNIV)"
56021
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   453
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   454
  { fix i from \<open>P M\<close> have "((F ^^ i) top) \<in> measurable M (count_space UNIV)"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   455
      by (induct i arbitrary: M) (auto intro!: *) }
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   456
  then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> measurable M (count_space UNIV)"
56021
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   457
    by measurable
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   458
  also have "(\<lambda>x. INF i. (F ^^ i) top x) = gfp F"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69605
diff changeset
   459
    by (subst inf_continuous_gfp) (auto intro: F simp: image_comp)
56021
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   460
  finally show ?thesis .
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   461
qed
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   462
e0c9d76c2a6d add measurability rule for (co)inductive predicates
hoelzl
parents: 53043
diff changeset
   463
lemma measurable_gfp:
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   464
  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 59361
diff changeset
   465
  assumes F: "inf_continuous F"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   466
  assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   467
  shows "gfp F \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   468
  by (coinduction rule: measurable_gfp_coinduct[OF _ F]) (blast intro: *)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   469
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   470
lemma measurable_lfp2_coinduct[consumes 1, case_names continuity step]:
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   471
  fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   472
  assumes "P M s"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 59361
diff changeset
   473
  assumes F: "sup_continuous F"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   474
  assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   475
  shows "lfp F s \<in> measurable M (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   476
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   477
  { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   478
      by (induct i arbitrary: M s) (auto intro!: *) }
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   479
  then have "(\<lambda>x. SUP i. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   480
    by measurable
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   481
  also have "(\<lambda>x. SUP i. (F ^^ i) bot s x) = lfp F s"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69605
diff changeset
   482
    by (subst sup_continuous_lfp) (auto simp: F simp: image_comp)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   483
  finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   484
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   485
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   486
lemma measurable_gfp2_coinduct[consumes 1, case_names continuity step]:
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   487
  fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   488
  assumes "P M s"
60172
423273355b55 rename continuous and down_continuous in Order_Continuity to sup_/inf_continuous; relate them with topological continuity
hoelzl
parents: 59361
diff changeset
   489
  assumes F: "inf_continuous F"
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   490
  assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   491
  shows "gfp F s \<in> measurable M (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   492
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   493
  { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   494
      by (induct i arbitrary: M s) (auto intro!: *) }
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   495
  then have "(\<lambda>x. INF i. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   496
    by measurable
59088
ff2bd4a14ddb generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
hoelzl
parents: 59048
diff changeset
   497
  also have "(\<lambda>x. INF i. (F ^^ i) top s x) = gfp F s"
69861
62e47f06d22c avoid context-sensitive simp rules whose context-free form (image_comp) is not simp by default
haftmann
parents: 69605
diff changeset
   498
    by (subst inf_continuous_gfp) (auto simp: F simp: image_comp)
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   499
  finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   500
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   501
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   502
lemma measurable_enat_coinduct:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   503
  fixes f :: "'a \<Rightarrow> enat"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   504
  assumes "R f"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   505
  assumes *: "\<And>f. R f \<Longrightarrow> \<exists>g h i P. R g \<and> f = (\<lambda>x. if P x then h x else eSuc (g (i x))) \<and>
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   506
    Measurable.pred M P \<and>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   507
    i \<in> measurable M M \<and>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   508
    h \<in> measurable M (count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   509
  shows "f \<in> measurable M (count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   510
proof (simp add: measurable_count_space_eq2_countable, rule )
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   511
  fix a :: enat
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   512
  have "f -` {a} \<inter> space M = {x\<in>space M. f x = a}"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   513
    by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   514
  { fix i :: nat
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   515
    from \<open>R f\<close> have "Measurable.pred M (\<lambda>x. f x = enat i)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   516
    proof (induction i arbitrary: f)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   517
      case 0
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   518
      from *[OF this] obtain g h i P
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   519
        where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   520
          [measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   521
        by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   522
      have "Measurable.pred M (\<lambda>x. P x \<and> h x = 0)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   523
        by measurable
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   524
      also have "(\<lambda>x. P x \<and> h x = 0) = (\<lambda>x. f x = enat 0)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   525
        by (auto simp: f zero_enat_def[symmetric])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   526
      finally show ?case .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   527
    next
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   528
      case (Suc n)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   529
      from *[OF Suc.prems] obtain g h i P
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   530
        where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and "R g" and
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   531
          M[measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   532
        by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   533
      have "(\<lambda>x. f x = enat (Suc n)) =
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   534
        (\<lambda>x. (P x \<longrightarrow> h x = enat (Suc n)) \<and> (\<not> P x \<longrightarrow> g (i x) = enat n))"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   535
        by (auto simp: f zero_enat_def[symmetric] eSuc_enat[symmetric])
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   536
      also have "Measurable.pred M \<dots>"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60172
diff changeset
   537
        by (intro pred_intros_logic measurable_compose[OF M(2)] Suc \<open>R g\<close>) measurable
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   538
      finally show ?case .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   539
    qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   540
    then have "f -` {enat i} \<inter> space M \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   541
      by (simp add: pred_def Int_def conj_commute) }
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   542
  note fin = this
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   543
  show "f -` {a} \<inter> space M \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   544
  proof (cases a)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   545
    case infinity
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   546
    then have "f -` {a} \<inter> space M = space M - (\<Union>n. f -` {enat n} \<inter> space M)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   547
      by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   548
    also have "\<dots> \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   549
      by (intro sets.Diff sets.top sets.Un sets.countable_UN) (auto intro!: fin)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   550
    finally show ?thesis .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   551
  qed (simp add: fin)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   552
qed
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   553
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   554
lemma measurable_THE:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   555
  fixes P :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   556
  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   557
  assumes I[simp]: "countable I" "\<And>i x. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> i \<in> I"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   558
  assumes unique: "\<And>x i j. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> P j x \<Longrightarrow> i = j"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   559
  shows "(\<lambda>x. THE i. P i x) \<in> measurable M (count_space UNIV)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   560
  unfolding measurable_def
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   561
proof safe
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   562
  fix X
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   563
  define f where "f x = (THE i. P i x)" for x
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62975
diff changeset
   564
  define undef where "undef = (THE i::'a. False)"
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   565
  { fix i x assume "x \<in> space M" "P i x" then have "f x = i"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   566
      unfolding f_def using unique by auto }
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   567
  note f_eq = this
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   568
  { fix x assume "x \<in> space M" "\<forall>i\<in>I. \<not> P i x"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   569
    then have "\<And>i. \<not> P i x"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   570
      using I(2)[of x] by auto
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   571
    then have "f x = undef"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   572
      by (auto simp: undef_def f_def) }
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   573
  then have "f -` X \<inter> space M = (\<Union>i\<in>I \<inter> X. {x\<in>space M. P i x}) \<union>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   574
     (if undef \<in> X then space M - (\<Union>i\<in>I. {x\<in>space M. P i x}) else {})"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   575
    by (auto dest: f_eq)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   576
  also have "\<dots> \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   577
    by (auto intro!: sets.Diff sets.countable_UN')
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   578
  finally show "f -` X \<inter> space M \<in> sets M" .
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   579
qed simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   580
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   581
lemma measurable_Ex1[measurable (raw)]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   582
  assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> Measurable.pred M (P i)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   583
  shows "Measurable.pred M (\<lambda>x. \<exists>!i\<in>I. P i x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   584
  unfolding bex1_def by measurable
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   585
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   586
lemma measurable_Sup_nat[measurable (raw)]:
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   587
  fixes F :: "'a \<Rightarrow> nat set"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   588
  assumes [measurable]: "\<And>i. Measurable.pred M (\<lambda>x. i \<in> F x)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   589
  shows "(\<lambda>x. Sup (F x)) \<in> M \<rightarrow>\<^sub>M count_space UNIV"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   590
proof (clarsimp simp add: measurable_count_space_eq2_countable)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   591
  fix a
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   592
  have F_empty_iff: "F x = {} \<longleftrightarrow> (\<forall>i. i \<notin> F x)" for x
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   593
    by auto
71834
919a55257e62 Fixes for Sup{} = (0::nat)
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   594
  have "Measurable.pred M (\<lambda>x. if finite (F x) then if F x = {} then a = 0
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   595
    else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None)"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   596
    unfolding finite_nat_set_iff_bounded Ball_def F_empty_iff by measurable
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   597
  moreover have "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M =
71834
919a55257e62 Fixes for Sup{} = (0::nat)
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   598
    {x\<in>space M. if finite (F x) then if F x = {} then a = 0
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   599
      else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None}"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   600
    by (intro set_eqI)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   601
       (auto simp: Sup_nat_def Max.infinite intro!: Max_in Max_eqI)
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   602
  ultimately show "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M \<in> sets M"
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   603
    by auto
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   604
qed
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   605
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
   606
lemma measurable_if_split[measurable (raw)]:
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   607
  "(c \<Longrightarrow> Measurable.pred M f) \<Longrightarrow> (\<not> c \<Longrightarrow> Measurable.pred M g) \<Longrightarrow>
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   608
   Measurable.pred M (if c then f else g)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   609
  by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   610
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   611
lemma pred_restrict_space:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   612
  assumes "S \<in> sets M"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   613
  shows "Measurable.pred (restrict_space M S) P \<longleftrightarrow> Measurable.pred M (\<lambda>x. x \<in> S \<and> P x)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   614
  unfolding pred_def sets_Collect_restrict_space_iff[OF assms] ..
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   615
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   616
lemma measurable_predpow[measurable]:
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   617
  assumes "Measurable.pred M T"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   618
  assumes "\<And>Q. Measurable.pred M Q \<Longrightarrow> Measurable.pred M (R Q)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   619
  shows "Measurable.pred M ((R ^^ n) T)"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   620
  by (induct n) (auto intro: assms)
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58965
diff changeset
   621
64008
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   622
lemma measurable_compose_countable_restrict:
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   623
  assumes P: "countable {i. P i}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   624
    and f: "f \<in> M \<rightarrow>\<^sub>M count_space UNIV"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   625
    and Q: "\<And>i. P i \<Longrightarrow> pred M (Q i)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   626
  shows "pred M (\<lambda>x. P (f x) \<and> Q (f x) x)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   627
proof -
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   628
  have P_f: "{x \<in> space M. P (f x)} \<in> sets M"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   629
    unfolding pred_def[symmetric] by (rule measurable_compose[OF f]) simp
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   630
  have "pred (restrict_space M {x\<in>space M. P (f x)}) (\<lambda>x. Q (f x) x)"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   631
  proof (rule measurable_compose_countable'[where g=f, OF _ _ P])
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   632
    show "f \<in> restrict_space M {x\<in>space M. P (f x)} \<rightarrow>\<^sub>M count_space {i. P i}"
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   633
      by (rule measurable_count_space_extend[OF subset_UNIV])
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   634
         (auto simp: space_restrict_space intro!: measurable_restrict_space1 f)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   635
  qed (auto intro!: measurable_restrict_space1 Q)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   636
  then show ?thesis
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   637
    unfolding pred_restrict_space[OF P_f] by (simp cong: measurable_cong)
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   638
qed
17a20ca86d62 HOL-Probability: more about probability, prepare for Markov processes in the AFP
hoelzl
parents: 63627
diff changeset
   639
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   640
lemma measurable_limsup [measurable (raw)]:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   641
  assumes [measurable]: "\<And>n. A n \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   642
  shows "limsup A \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   643
by (subst limsup_INF_SUP, auto)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   644
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   645
lemma measurable_liminf [measurable (raw)]:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   646
  assumes [measurable]: "\<And>n. A n \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   647
  shows "liminf A \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   648
by (subst liminf_SUP_INF, auto)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 64008
diff changeset
   649
64320
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   650
lemma measurable_case_enat[measurable (raw)]:
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   651
  assumes f: "f \<in> M \<rightarrow>\<^sub>M count_space UNIV" and g: "\<And>i. g i \<in> M \<rightarrow>\<^sub>M N" and h: "h \<in> M \<rightarrow>\<^sub>M N"
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   652
  shows "(\<lambda>x. case f x of enat i \<Rightarrow> g i x | \<infinity> \<Rightarrow> h x) \<in> M \<rightarrow>\<^sub>M N"
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   653
  apply (rule measurable_compose_countable[OF _ f])
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   654
  subgoal for i
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   655
    by (cases i) (auto intro: g h)
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   656
  done
ba194424b895 HOL-Probability: move stopping time from AFP/Markov_Models
hoelzl
parents: 64283
diff changeset
   657
50387
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   658
hide_const (open) pred
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   659
3d8863c41fe8 Move the measurability prover to its own file
hoelzl
parents:
diff changeset
   660
end
59048
7dc8ac6f0895 add congruence solver to measurability prover
hoelzl
parents: 59047
diff changeset
   661