src/HOL/Probability/Positive_Extended_Real.thy
author hoelzl
Fri, 03 Dec 2010 15:25:14 +0100
changeset 41023 9118eb4eb8dc
parent 40874 src/HOL/Probability/Positive_Infinite_Real.thy@f5a74b17a69e
child 41096 843c40bbc379
permissions -rw-r--r--
it is known as the extended reals, not the infinite reals
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     1
(* Author: Johannes Hoelzl, TU Muenchen *)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     2
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     3
header {* A type for positive real numbers with infinity *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     4
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
     5
theory Positive_Extended_Real
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     6
  imports Complex_Main Nat_Bijection Multivariate_Analysis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     7
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
     8
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
     9
lemma (in complete_lattice) Sup_start:
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    10
  assumes *: "\<And>x. f x \<le> f 0"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    11
  shows "(SUP n. f n) = f 0"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    12
proof (rule antisym)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    13
  show "f 0 \<le> (SUP n. f n)" by (rule le_SUPI) auto
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    14
  show "(SUP n. f n) \<le> f 0" by (rule SUP_leI[OF *])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    15
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    16
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    17
lemma (in complete_lattice) Inf_start:
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    18
  assumes *: "\<And>x. f 0 \<le> f x"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    19
  shows "(INF n. f n) = f 0"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    20
proof (rule antisym)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    21
  show "(INF n. f n) \<le> f 0" by (rule INF_leI) simp
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    22
  show "f 0 \<le> (INF n. f n)" by (rule le_INFI[OF *])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    23
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    24
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    25
lemma (in complete_lattice) Sup_mono_offset:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    26
  fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    27
  assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y" and "0 \<le> k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    28
  shows "(SUP n . f (k + n)) = (SUP n. f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    29
proof (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    30
  show "(SUP n. f (k + n)) \<le> (SUP n. f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    31
    by (auto intro!: Sup_mono simp: SUPR_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    32
  { fix n :: 'b
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    33
    have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    34
    with * have "f n \<le> f (k + n)" by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    35
  thus "(SUP n. f n) \<le> (SUP n. f (k + n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    36
    by (auto intro!: Sup_mono exI simp: SUPR_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    37
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    38
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    39
lemma (in complete_lattice) Sup_mono_offset_Suc:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    40
  assumes *: "\<And>x. f x \<le> f (Suc x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    41
  shows "(SUP n . f (Suc n)) = (SUP n. f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    42
  unfolding Suc_eq_plus1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    43
  apply (subst add_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    44
  apply (rule Sup_mono_offset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    45
  by (auto intro!: order.lift_Suc_mono_le[of "op \<le>" "op <" f, OF _ *]) default
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    46
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    47
lemma (in complete_lattice) Inf_mono_offset:
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    48
  fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \<Rightarrow> 'a"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    49
  assumes *: "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x" and "0 \<le> k"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    50
  shows "(INF n . f (k + n)) = (INF n. f n)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    51
proof (rule antisym)
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    52
  show "(INF n. f n) \<le> (INF n. f (k + n))"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    53
    by (auto intro!: Inf_mono simp: INFI_def)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    54
  { fix n :: 'b
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    55
    have "0 + n \<le> k + n" using `0 \<le> k` by (rule add_right_mono)
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    56
    with * have "f (k + n) \<le> f n" by simp }
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    57
  thus "(INF n. f (k + n)) \<le> (INF n. f n)"
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
    58
    by (auto intro!: Inf_mono exI simp: INFI_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    59
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    60
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    61
lemma (in complete_lattice) isotone_converge:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    62
  fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f x \<le> f y "
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    63
  shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    64
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    65
  have "\<And>n. (SUP m. f (n + m)) = (SUP n. f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    66
    apply (rule Sup_mono_offset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    67
    apply (rule assms)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    68
    by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    69
  moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    70
  { fix n have "(INF m. f (n + m)) = f n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    71
      using Inf_start[of "\<lambda>m. f (n + m)"] assms by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    72
  ultimately show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    73
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    74
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    75
lemma (in complete_lattice) antitone_converges:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    76
  fixes f :: "nat \<Rightarrow> 'a" assumes "\<And>x y. x \<le> y \<Longrightarrow> f y \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    77
  shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    78
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    79
  have "\<And>n. (INF m. f (n + m)) = (INF n. f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    80
    apply (rule Inf_mono_offset)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    81
    apply (rule assms)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    82
    by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    83
  moreover
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    84
  { fix n have "(SUP m. f (n + m)) = f n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    85
      using Sup_start[of "\<lambda>m. f (n + m)"] assms by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    86
  ultimately show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    87
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
    88
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    89
lemma (in complete_lattice) lim_INF_le_lim_SUP:
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    90
  fixes f :: "nat \<Rightarrow> 'a"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    91
  shows "(SUP n. INF m. f (n + m)) \<le> (INF n. SUP m. f (n + m))"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    92
proof (rule SUP_leI, rule le_INFI)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    93
  fix i j show "(INF m. f (i + m)) \<le> (SUP m. f (j + m))"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    94
  proof (cases rule: le_cases)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    95
    assume "i \<le> j"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    96
    have "(INF m. f (i + m)) \<le> f (i + (j - i))" by (rule INF_leI) simp
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    97
    also have "\<dots> = f (j + 0)" using `i \<le> j` by auto
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    98
    also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
    99
    finally show ?thesis .
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   100
  next
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   101
    assume "j \<le> i"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   102
    have "(INF m. f (i + m)) \<le> f (i + 0)" by (rule INF_leI) simp
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   103
    also have "\<dots> = f (j + (i - j))" using `j \<le> i` by auto
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   104
    also have "\<dots> \<le> (SUP m. f (j + m))" by (rule le_SUPI) simp
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   105
    finally show ?thesis .
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   106
  qed
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   107
qed
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
   108
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   109
text {*
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   110
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   111
We introduce the the positive real numbers as needed for measure theory.
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   112
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   113
*}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   114
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   115
typedef pextreal = "(Some ` {0::real..}) \<union> {None}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   116
  by (rule exI[of _ None]) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   117
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   118
subsection "Introduce @{typ pextreal} similar to a datatype"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   119
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   120
definition "Real x = Abs_pextreal (Some (sup 0 x))"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   121
definition "\<omega> = Abs_pextreal None"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   122
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   123
definition "pextreal_case f i x = (if x = \<omega> then i else f (THE r. 0 \<le> r \<and> x = Real r))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   124
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   125
definition "of_pextreal = pextreal_case (\<lambda>x. x) 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   126
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   127
defs (overloaded)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   128
  real_of_pextreal_def [code_unfold]: "real == of_pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   129
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   130
lemma pextreal_Some[simp]: "0 \<le> x \<Longrightarrow> Some x \<in> pextreal"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   131
  unfolding pextreal_def by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   132
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   133
lemma pextreal_Some_sup[simp]: "Some (sup 0 x) \<in> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   134
  by (simp add: sup_ge1)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   135
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   136
lemma pextreal_None[simp]: "None \<in> pextreal"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   137
  unfolding pextreal_def by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   138
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   139
lemma Real_inj[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   140
  assumes  "0 \<le> x" and "0 \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   141
  shows "Real x = Real y \<longleftrightarrow> x = y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   142
  unfolding Real_def assms[THEN sup_absorb2]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   143
  using assms by (simp add: Abs_pextreal_inject)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   144
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   145
lemma Real_neq_\<omega>[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   146
  "Real x = \<omega> \<longleftrightarrow> False"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   147
  "\<omega> = Real x \<longleftrightarrow> False"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   148
  by (simp_all add: Abs_pextreal_inject \<omega>_def Real_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   149
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   150
lemma Real_neg: "x < 0 \<Longrightarrow> Real x = Real 0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   151
  unfolding Real_def by (auto simp add: Abs_pextreal_inject intro!: sup_absorb1)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   152
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   153
lemma pextreal_cases[case_names preal infinite, cases type: pextreal]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   154
  assumes preal: "\<And>r. x = Real r \<Longrightarrow> 0 \<le> r \<Longrightarrow> P" and inf: "x = \<omega> \<Longrightarrow> P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   155
  shows P
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   156
proof (cases x rule: pextreal.Abs_pextreal_cases)
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   157
  case (Abs_pextreal y)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   158
  hence "y = None \<or> (\<exists>x \<ge> 0. y = Some x)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   159
    unfolding pextreal_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   160
  thus P
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   161
  proof (rule disjE)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   162
    assume "\<exists>x\<ge>0. y = Some x" then guess x ..
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   163
    thus P by (simp add: preal[of x] Real_def Abs_pextreal(1) sup_absorb2)
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   164
  qed (simp add: \<omega>_def Abs_pextreal(1) inf)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   165
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   166
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   167
lemma pextreal_case_\<omega>[simp]: "pextreal_case f i \<omega> = i"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   168
  unfolding pextreal_case_def by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   169
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   170
lemma pextreal_case_Real[simp]: "pextreal_case f i (Real x) = (if 0 \<le> x then f x else f 0)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   171
proof (cases "0 \<le> x")
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   172
  case True thus ?thesis unfolding pextreal_case_def by (auto intro: theI2)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   173
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   174
  case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   175
  moreover have "(THE r. 0 \<le> r \<and> Real 0 = Real r) = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   176
    by (auto intro!: the_equality)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   177
  ultimately show ?thesis unfolding pextreal_case_def by (simp add: Real_neg)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   178
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   179
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   180
lemma pextreal_case_cancel[simp]: "pextreal_case (\<lambda>c. i) i x = i"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   181
  by (cases x) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   182
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   183
lemma pextreal_case_split:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   184
  "P (pextreal_case f i x) = ((x = \<omega> \<longrightarrow> P i) \<and> (\<forall>r\<ge>0. x = Real r \<longrightarrow> P (f r)))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   185
  by (cases x) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   186
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   187
lemma pextreal_case_split_asm:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   188
  "P (pextreal_case f i x) = (\<not> (x = \<omega> \<and> \<not> P i \<or> (\<exists>r. r \<ge> 0 \<and> x = Real r \<and> \<not> P (f r))))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   189
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   190
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   191
lemma pextreal_case_cong[cong]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   192
  assumes eq: "x = x'" "i = i'" and cong: "\<And>r. 0 \<le> r \<Longrightarrow> f r = f' r"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   193
  shows "pextreal_case f i x = pextreal_case f' i' x'"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   194
  unfolding eq using cong by (cases x') simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   195
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   196
lemma real_Real[simp]: "real (Real x) = (if 0 \<le> x then x else 0)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   197
  unfolding real_of_pextreal_def of_pextreal_def by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   198
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   199
lemma Real_real_image:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   200
  assumes "\<omega> \<notin> A" shows "Real ` real ` A = A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   201
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   202
  fix x assume "x \<in> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   203
  hence *: "x = Real (real x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   204
    using `\<omega> \<notin> A` by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   205
  show "x \<in> Real ` real ` A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   206
    using `x \<in> A` by (subst *) (auto intro!: imageI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   207
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   208
  fix x assume "x \<in> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   209
  thus "Real (real x) \<in> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   210
    using `\<omega> \<notin> A` by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   211
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   212
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   213
lemma real_pextreal_nonneg[simp, intro]: "0 \<le> real (x :: pextreal)"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   214
  unfolding real_of_pextreal_def of_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   215
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   216
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   217
lemma real_\<omega>[simp]: "real \<omega> = 0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   218
  unfolding real_of_pextreal_def of_pextreal_def by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   219
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   220
lemma pextreal_noteq_omega_Ex: "X \<noteq> \<omega> \<longleftrightarrow> (\<exists>r\<ge>0. X = Real r)" by (cases X) auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   221
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   222
subsection "@{typ pextreal} is a monoid for addition"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   223
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   224
instantiation pextreal :: comm_monoid_add
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   225
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   226
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   227
definition "0 = Real 0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   228
definition "x + y = pextreal_case (\<lambda>r. pextreal_case (\<lambda>p. Real (r + p)) \<omega> y) \<omega> x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   229
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   230
lemma pextreal_plus[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   231
  "Real r + Real p = (if 0 \<le> r then if 0 \<le> p then Real (r + p) else Real r else Real p)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   232
  "x + 0 = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   233
  "0 + x = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   234
  "x + \<omega> = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   235
  "\<omega> + x = \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   236
  by (simp_all add: plus_pextreal_def Real_neg zero_pextreal_def split: pextreal_case_split)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   237
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   238
lemma \<omega>_neq_0[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   239
  "\<omega> = 0 \<longleftrightarrow> False"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   240
  "0 = \<omega> \<longleftrightarrow> False"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   241
  by (simp_all add: zero_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   242
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   243
lemma Real_eq_0[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   244
  "Real r = 0 \<longleftrightarrow> r \<le> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   245
  "0 = Real r \<longleftrightarrow> r \<le> 0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   246
  by (auto simp add: Abs_pextreal_inject zero_pextreal_def Real_def sup_real_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   247
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   248
lemma Real_0[simp]: "Real 0 = 0" by (simp add: zero_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   249
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   250
instance
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   251
proof
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   252
  fix a :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   253
  show "0 + a = a" by (cases a) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   254
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   255
  fix b show "a + b = b + a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   256
    by (cases a, cases b) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   257
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   258
  fix c show "a + b + c = a + (b + c)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   259
    by (cases a, cases b, cases c) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   260
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   261
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   262
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   263
lemma pextreal_plus_eq_\<omega>[simp]: "(a :: pextreal) + b = \<omega> \<longleftrightarrow> a = \<omega> \<or> b = \<omega>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   264
  by (cases a, cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   265
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   266
lemma pextreal_add_cancel_left:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   267
  "a + b = a + c \<longleftrightarrow> (a = \<omega> \<or> b = c)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   268
  by (cases a, cases b, cases c, simp_all, cases c, simp_all)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   269
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   270
lemma pextreal_add_cancel_right:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   271
  "b + a = c + a \<longleftrightarrow> (a = \<omega> \<or> b = c)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   272
  by (cases a, cases b, cases c, simp_all, cases c, simp_all)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   273
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   274
lemma Real_eq_Real:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   275
  "Real a = Real b \<longleftrightarrow> (a = b \<or> (a \<le> 0 \<and> b \<le> 0))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   276
proof (cases "a \<le> 0 \<or> b \<le> 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   277
  case False with Real_inj[of a b] show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   278
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   279
  case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   280
  thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   281
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   282
    assume "a \<le> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   283
    hence *: "Real a = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   284
    show ?thesis using `a \<le> 0` unfolding * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   285
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   286
    assume "b \<le> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   287
    hence *: "Real b = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   288
    show ?thesis using `b \<le> 0` unfolding * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   289
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   290
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   291
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   292
lemma real_pextreal_0[simp]: "real (0 :: pextreal) = 0"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   293
  unfolding zero_pextreal_def real_Real by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   294
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   295
lemma real_of_pextreal_eq_0: "real X = 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   296
  by (cases X) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   297
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   298
lemma real_of_pextreal_eq: "real X = real Y \<longleftrightarrow>
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   299
    (X = Y \<or> (X = 0 \<and> Y = \<omega>) \<or> (Y = 0 \<and> X = \<omega>))"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   300
  by (cases X, cases Y) (auto simp add: real_of_pextreal_eq_0)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   301
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   302
lemma real_of_pextreal_add: "real X + real Y =
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   303
    (if X = \<omega> then real Y else if Y = \<omega> then real X else real (X + Y))"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   304
  by (auto simp: pextreal_noteq_omega_Ex)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   305
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   306
subsection "@{typ pextreal} is a monoid for multiplication"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   307
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   308
instantiation pextreal :: comm_monoid_mult
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   309
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   310
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   311
definition "1 = Real 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   312
definition "x * y = (if x = 0 \<or> y = 0 then 0 else
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   313
  pextreal_case (\<lambda>r. pextreal_case (\<lambda>p. Real (r * p)) \<omega> y) \<omega> x)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   314
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   315
lemma pextreal_times[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   316
  "Real r * Real p = (if 0 \<le> r \<and> 0 \<le> p then Real (r * p) else 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   317
  "\<omega> * x = (if x = 0 then 0 else \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   318
  "x * \<omega> = (if x = 0 then 0 else \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   319
  "0 * x = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   320
  "x * 0 = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   321
  "1 = \<omega> \<longleftrightarrow> False"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   322
  "\<omega> = 1 \<longleftrightarrow> False"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   323
  by (auto simp add: times_pextreal_def one_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   324
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   325
lemma pextreal_one_mult[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   326
  "Real x + 1 = (if 0 \<le> x then Real (x + 1) else 1)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   327
  "1 + Real x = (if 0 \<le> x then Real (1 + x) else 1)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   328
  unfolding one_pextreal_def by simp_all
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   329
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   330
instance
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   331
proof
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   332
  fix a :: pextreal show "1 * a = a"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   333
    by (cases a) (simp_all add: one_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   334
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   335
  fix b show "a * b = b * a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   336
    by (cases a, cases b) (simp_all add: mult_nonneg_nonneg)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   337
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   338
  fix c show "a * b * c = a * (b * c)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   339
    apply (cases a, cases b, cases c)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   340
    apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   341
    apply (cases b, cases c)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   342
    apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   343
    done
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   344
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   345
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   346
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   347
lemma pextreal_mult_cancel_left:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   348
  "a * b = a * c \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   349
  by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   350
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   351
lemma pextreal_mult_cancel_right:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   352
  "b * a = c * a \<longleftrightarrow> (a = 0 \<or> b = c \<or> (a = \<omega> \<and> b \<noteq> 0 \<and> c \<noteq> 0))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   353
  by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   354
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   355
lemma Real_1[simp]: "Real 1 = 1" by (simp add: one_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   356
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   357
lemma real_pextreal_1[simp]: "real (1 :: pextreal) = 1"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   358
  unfolding one_pextreal_def real_Real by simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   359
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   360
lemma real_of_pextreal_mult: "real X * real Y = real (X * Y :: pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   361
  by (cases X, cases Y) (auto simp: zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   362
40874
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   363
lemma Real_mult_nonneg: assumes "x \<ge> 0" "y \<ge> 0"
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   364
  shows "Real (x * y) = Real x * Real y" using assms by auto
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   365
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   366
lemma Real_setprod: assumes "\<forall>x\<in>A. f x \<ge> 0" shows "Real (setprod f A) = setprod (\<lambda>x. Real (f x)) A"
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   367
proof(cases "finite A")
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   368
  case True thus ?thesis using assms
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   369
  proof(induct A) case (insert x A)
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   370
    have "0 \<le> setprod f A" apply(rule setprod_nonneg) using insert by auto
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   371
    thus ?case unfolding setprod_insert[OF insert(1-2)] apply-
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   372
      apply(subst Real_mult_nonneg) prefer 3 apply(subst insert(3)[THEN sym])
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   373
      using insert by auto
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   374
  qed auto
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   375
qed auto
f5a74b17a69e Moved theorems to appropriate place.
hoelzl
parents: 40872
diff changeset
   376
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   377
subsection "@{typ pextreal} is a linear order"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   378
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   379
instantiation pextreal :: linorder
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   380
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   381
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   382
definition "x < y \<longleftrightarrow> pextreal_case (\<lambda>i. pextreal_case (\<lambda>j. i < j) True y) False x"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   383
definition "x \<le> y \<longleftrightarrow> pextreal_case (\<lambda>j. pextreal_case (\<lambda>i. i \<le> j) False x) True y"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   384
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   385
lemma pextreal_less[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   386
  "Real r < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   387
  "Real r < Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r < p else 0 < p)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   388
  "\<omega> < x \<longleftrightarrow> False"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   389
  "0 < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   390
  "0 < Real r \<longleftrightarrow> 0 < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   391
  "x < 0 \<longleftrightarrow> False"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   392
  "0 < (1::pextreal)"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   393
  by (simp_all add: less_pextreal_def zero_pextreal_def one_pextreal_def del: Real_0 Real_1)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   394
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   395
lemma pextreal_less_eq[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   396
  "x \<le> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   397
  "Real r \<le> Real p \<longleftrightarrow> (if 0 \<le> r \<and> 0 \<le> p then r \<le> p else r \<le> 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   398
  "0 \<le> x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   399
  by (simp_all add: less_eq_pextreal_def zero_pextreal_def del: Real_0)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   400
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   401
lemma pextreal_\<omega>_less_eq[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   402
  "\<omega> \<le> x \<longleftrightarrow> x = \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   403
  by (cases x) (simp_all add: not_le less_eq_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   404
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   405
lemma pextreal_less_eq_zero[simp]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   406
  "(x::pextreal) \<le> 0 \<longleftrightarrow> x = 0"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   407
  by (cases x) (simp_all add: zero_pextreal_def del: Real_0)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   408
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   409
instance
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   410
proof
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   411
  fix x :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   412
  show "x \<le> x" by (cases x) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   413
  fix y
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   414
  show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   415
    by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   416
  show "x \<le> y \<or> y \<le> x "
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   417
    by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   418
  { assume "x \<le> y" "y \<le> x" thus "x = y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   419
      by (cases x, cases y) auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   420
  { fix z assume "x \<le> y" "y \<le> z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   421
    thus "x \<le> z" by (cases x, cases y, cases z) auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   422
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   423
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   424
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   425
lemma pextreal_zero_lessI[intro]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   426
  "(a :: pextreal) \<noteq> 0 \<Longrightarrow> 0 < a"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   427
  by (cases a) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   428
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   429
lemma pextreal_less_omegaI[intro, simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   430
  "a \<noteq> \<omega> \<Longrightarrow> a < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   431
  by (cases a) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   432
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   433
lemma pextreal_plus_eq_0[simp]: "(a :: pextreal) + b = 0 \<longleftrightarrow> a = 0 \<and> b = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   434
  by (cases a, cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   435
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   436
lemma pextreal_le_add1[simp, intro]: "n \<le> n + (m::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   437
  by (cases n, cases m) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   438
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   439
lemma pextreal_le_add2: "(n::pextreal) + m \<le> k \<Longrightarrow> m \<le> k"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   440
  by (cases n, cases m, cases k) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   441
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   442
lemma pextreal_le_add3: "(n::pextreal) + m \<le> k \<Longrightarrow> n \<le> k"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   443
  by (cases n, cases m, cases k) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   444
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   445
lemma pextreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   446
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   447
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   448
lemma pextreal_0_less_mult_iff[simp]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   449
  fixes x y :: pextreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
39092
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   450
  by (cases x, cases y) (auto simp: zero_less_mult_iff)
98de40859858 move lemmas to correct theory files
hoelzl
parents: 38705
diff changeset
   451
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   452
lemma pextreal_ord_one[simp]:
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   453
  "Real p < 1 \<longleftrightarrow> p < 1"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   454
  "Real p \<le> 1 \<longleftrightarrow> p \<le> 1"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   455
  "1 < Real p \<longleftrightarrow> 1 < p"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   456
  "1 \<le> Real p \<longleftrightarrow> 1 \<le> p"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   457
  by (simp_all add: one_pextreal_def del: Real_1)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   458
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   459
subsection {* @{text "x - y"} on @{typ pextreal} *}
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   460
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   461
instantiation pextreal :: minus
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   462
begin
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   463
definition "x - y = (if y < x then THE d. x = y + d else 0 :: pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   464
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   465
lemma minus_pextreal_eq:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   466
  "(x - y = (z :: pextreal)) \<longleftrightarrow> (if y < x then x = y + z else z = 0)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   467
  (is "?diff \<longleftrightarrow> ?if")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   468
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   469
  assume ?diff
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   470
  thus ?if
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   471
  proof (cases "y < x")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   472
    case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   473
    then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   474
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   475
    show ?thesis unfolding `?diff`[symmetric] if_P[OF True] minus_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   476
    proof (rule theI2[where Q="\<lambda>d. x = y + d"])
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   477
      show "x = y + pextreal_case (\<lambda>r. Real (r - real y)) \<omega> x" (is "x = y + ?d")
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   478
        using `y < x` p by (cases x) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   479
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   480
      fix d assume "x = y + d"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   481
      thus "d = ?d" using `y < x` p by (cases d, cases x) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   482
    qed simp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   483
  qed (simp add: minus_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   484
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   485
  assume ?if
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   486
  thus ?diff
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   487
  proof (cases "y < x")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   488
    case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   489
    then obtain p where p: "y = Real p" "0 \<le> p" by (cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   490
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   491
    from True `?if` have "x = y + z" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   492
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   493
    show ?thesis unfolding minus_pextreal_def if_P[OF True] unfolding `x = y + z`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   494
    proof (rule the_equality)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   495
      fix d :: pextreal assume "y + z = y + d"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   496
      thus "d = z" using `y < x` p
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   497
        by (cases d, cases z) simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   498
    qed simp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   499
  qed (simp add: minus_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   500
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   501
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   502
instance ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   503
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   504
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   505
lemma pextreal_minus[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   506
  "Real r - Real p = (if 0 \<le> r \<and> p < r then if 0 \<le> p then Real (r - p) else Real r else 0)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   507
  "(A::pextreal) - A = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   508
  "\<omega> - Real r = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   509
  "Real r - \<omega> = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   510
  "A - 0 = A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   511
  "0 - A = 0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   512
  by (auto simp: minus_pextreal_eq not_less)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   513
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   514
lemma pextreal_le_epsilon:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   515
  fixes x y :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   516
  assumes "\<And>e. 0 < e \<Longrightarrow> x \<le> y + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   517
  shows "x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   518
proof (cases y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   519
  case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   520
  then obtain p where x: "x = Real p" "0 \<le> p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   521
    using assms[of 1] by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   522
  { fix e have "0 < e \<Longrightarrow> p \<le> r + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   523
      using assms[of "Real e"] preal x by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   524
  hence "p \<le> r" by (rule field_le_epsilon)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   525
  thus ?thesis using preal x by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   526
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   527
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   528
instance pextreal :: "{ordered_comm_semiring, comm_semiring_1}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   529
proof
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   530
  show "0 \<noteq> (1::pextreal)" unfolding zero_pextreal_def one_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   531
    by (simp del: Real_1 Real_0)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   532
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   533
  fix a :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   534
  show "0 * a = 0" "a * 0 = 0" by simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   535
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   536
  fix b c :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   537
  show "(a + b) * c = a * c + b * c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   538
    by (cases c, cases a, cases b)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   539
       (auto intro!: arg_cong[where f=Real] simp: field_simps not_le mult_le_0_iff mult_less_0_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   540
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   541
  { assume "a \<le> b" thus "c + a \<le> c + b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   542
     by (cases c, cases a, cases b) auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   543
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   544
  assume "a \<le> b" "0 \<le> c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   545
  thus "c * a \<le> c * b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   546
    apply (cases c, cases a, cases b)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   547
    by (auto simp: mult_left_mono mult_le_0_iff mult_less_0_iff not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   548
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   549
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   550
lemma mult_\<omega>[simp]: "x * y = \<omega> \<longleftrightarrow> (x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   551
  by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   552
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   553
lemma \<omega>_mult[simp]: "(\<omega> = x * y) = ((x = \<omega> \<or> y = \<omega>) \<and> x \<noteq> 0 \<and> y \<noteq> 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   554
  by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   555
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   556
lemma pextreal_mult_0[simp]: "x * y = 0 \<longleftrightarrow> x = 0 \<or> (y::pextreal) = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   557
  by (cases x, cases y) (auto simp: mult_le_0_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   558
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   559
lemma pextreal_mult_cancel:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   560
  fixes x y z :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   561
  assumes "y \<le> z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   562
  shows "x * y \<le> x * z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   563
  using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   564
  by (cases x, cases y, cases z)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   565
     (auto simp: mult_le_cancel_left mult_le_0_iff mult_less_0_iff not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   566
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   567
lemma Real_power[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   568
  "Real x ^ n = (if x \<le> 0 then (if n = 0 then 1 else 0) else Real (x ^ n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   569
  by (induct n) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   570
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   571
lemma Real_power_\<omega>[simp]:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   572
  "\<omega> ^ n = (if n = 0 then 1 else \<omega>)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   573
  by (induct n) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   574
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   575
lemma pextreal_of_nat[simp]: "of_nat m = Real (real m)"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   576
  by (induct m) (auto simp: real_of_nat_Suc one_pextreal_def simp del: Real_1)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   577
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   578
lemma less_\<omega>_Ex_of_nat: "x < \<omega> \<longleftrightarrow> (\<exists>n. x < of_nat n)"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   579
proof safe
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   580
  assume "x < \<omega>"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   581
  then obtain r where "0 \<le> r" "x = Real r" by (cases x) auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   582
  moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   583
  ultimately show "\<exists>n. x < of_nat n" by (auto simp: real_eq_of_nat)
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   584
qed auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   585
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   586
lemma real_of_pextreal_mono:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   587
  fixes a b :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   588
  assumes "b \<noteq> \<omega>" "a \<le> b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   589
  shows "real a \<le> real b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   590
using assms by (cases b, cases a) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   591
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   592
lemma setprod_pextreal_0:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   593
  "(\<Prod>i\<in>I. f i) = (0::pextreal) \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = 0)"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   594
proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   595
  assume "finite I" then show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   596
  proof (induct I)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   597
    case (insert i I)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   598
    then show ?case by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   599
  qed simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   600
qed simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   601
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   602
lemma setprod_\<omega>:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   603
  "(\<Prod>i\<in>I. f i) = \<omega> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<omega>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   604
proof cases
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   605
  assume "finite I" then show ?thesis
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   606
  proof (induct I)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   607
    case (insert i I) then show ?case
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   608
      by (auto simp: setprod_pextreal_0)
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   609
  qed simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   610
qed simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
   611
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   612
instance pextreal :: "semiring_char_0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   613
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   614
  fix m n
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   615
  show "inj (of_nat::nat\<Rightarrow>pextreal)" by (auto intro!: inj_onI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   616
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   617
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   618
subsection "@{typ pextreal} is a complete lattice"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   619
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   620
instantiation pextreal :: lattice
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   621
begin
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   622
definition [simp]: "sup x y = (max x y :: pextreal)"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   623
definition [simp]: "inf x y = (min x y :: pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   624
instance proof qed simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   625
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   626
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   627
instantiation pextreal :: complete_lattice
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   628
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   629
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   630
definition "bot = Real 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   631
definition "top = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   632
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   633
definition "Sup S = (LEAST z. \<forall>x\<in>S. x \<le> z :: pextreal)"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   634
definition "Inf S = (GREATEST z. \<forall>x\<in>S. z \<le> x :: pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   635
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   636
lemma pextreal_complete_Sup:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   637
  fixes S :: "pextreal set" assumes "S \<noteq> {}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   638
  shows "\<exists>x. (\<forall>y\<in>S. y \<le> x) \<and> (\<forall>z. (\<forall>y\<in>S. y \<le> z) \<longrightarrow> x \<le> z)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   639
proof (cases "\<exists>x\<ge>0. \<forall>a\<in>S. a \<le> Real x")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   640
  case False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   641
  hence *: "\<And>x. x\<ge>0 \<Longrightarrow> \<exists>a\<in>S. \<not>a \<le> Real x" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   642
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   643
  proof (safe intro!: exI[of _ \<omega>])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   644
    fix y assume **: "\<forall>z\<in>S. z \<le> y"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   645
    show "\<omega> \<le> y" unfolding pextreal_\<omega>_less_eq
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   646
    proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   647
      assume "y \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   648
      then obtain x where [simp]: "y = Real x" and "0 \<le> x" by (cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   649
      from *[OF `0 \<le> x`] show False using ** by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   650
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   651
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   652
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   653
  case True then obtain y where y: "\<And>a. a\<in>S \<Longrightarrow> a \<le> Real y" and "0 \<le> y" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   654
  from y[of \<omega>] have "\<omega> \<notin> S" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   655
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   656
  with `S \<noteq> {}` obtain x where "x \<in> S" and "x \<noteq> \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   657
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   658
  have bound: "\<forall>x\<in>real ` S. x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   659
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   660
    fix z assume "z \<in> real ` S" then guess a ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   661
    with y[of a] `\<omega> \<notin> S` `0 \<le> y` show "z \<le> y" by (cases a) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   662
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   663
  with reals_complete2[of "real ` S"] `x \<in> S`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   664
  obtain s where s: "\<forall>y\<in>S. real y \<le> s" "\<forall>z. ((\<forall>y\<in>S. real y \<le> z) \<longrightarrow> s \<le> z)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   665
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   666
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   667
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   668
  proof (safe intro!: exI[of _ "Real s"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   669
    fix z assume "z \<in> S" thus "z \<le> Real s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   670
      using s `\<omega> \<notin> S` by (cases z) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   671
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   672
    fix z assume *: "\<forall>y\<in>S. y \<le> z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   673
    show "Real s \<le> z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   674
    proof (cases z)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   675
      case (preal u)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   676
      { fix v assume "v \<in> S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   677
        hence "v \<le> Real u" using * preal by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   678
        hence "real v \<le> u" using `\<omega> \<notin> S` `0 \<le> u` by (cases v) auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   679
      hence "s \<le> u" using s(2) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   680
      thus "Real s \<le> z" using preal by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   681
    qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   682
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   683
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   684
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   685
lemma pextreal_complete_Inf:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   686
  fixes S :: "pextreal set" assumes "S \<noteq> {}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   687
  shows "\<exists>x. (\<forall>y\<in>S. x \<le> y) \<and> (\<forall>z. (\<forall>y\<in>S. z \<le> y) \<longrightarrow> z \<le> x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   688
proof (cases "S = {\<omega>}")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   689
  case True thus ?thesis by (auto intro!: exI[of _ \<omega>])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   690
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   691
  case False with `S \<noteq> {}` have "S - {\<omega>} \<noteq> {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   692
  hence not_empty: "\<exists>x. x \<in> uminus ` real ` (S - {\<omega>})" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   693
  have bounds: "\<exists>x. \<forall>y\<in>uminus ` real ` (S - {\<omega>}). y \<le> x" by (auto intro!: exI[of _ 0])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   694
  from reals_complete2[OF not_empty bounds]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   695
  obtain s where s: "\<And>y. y\<in>S - {\<omega>} \<Longrightarrow> - real y \<le> s" "\<forall>z. ((\<forall>y\<in>S - {\<omega>}. - real y \<le> z) \<longrightarrow> s \<le> z)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   696
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   697
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   698
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   699
  proof (safe intro!: exI[of _ "Real (-s)"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   700
    fix z assume "z \<in> S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   701
    show "Real (-s) \<le> z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   702
    proof (cases z)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   703
      case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   704
      with s `z \<in> S` have "z \<in> S - {\<omega>}" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   705
      hence "- r \<le> s" using preal s(1)[of z] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   706
      hence "- s \<le> r" by (subst neg_le_iff_le[symmetric]) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   707
      thus ?thesis using preal by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   708
    qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   709
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   710
    fix z assume *: "\<forall>y\<in>S. z \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   711
    show "z \<le> Real (-s)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   712
    proof (cases z)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   713
      case (preal u)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   714
      { fix v assume "v \<in> S-{\<omega>}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   715
        hence "Real u \<le> v" using * preal by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   716
        hence "- real v \<le> - u" using `0 \<le> u` `v \<in> S - {\<omega>}` by (cases v) auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   717
      hence "u \<le> - s" using s(2) by (subst neg_le_iff_le[symmetric]) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   718
      thus "z \<le> Real (-s)" using preal by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   719
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   720
      case infinite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   721
      with * have "S = {\<omega>}" using `S \<noteq> {}` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   722
      with `S - {\<omega>} \<noteq> {}` show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   723
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   724
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   725
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   726
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   727
instance
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   728
proof
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   729
  fix x :: pextreal and A
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   730
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   731
  show "bot \<le> x" by (cases x) (simp_all add: bot_pextreal_def)
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   732
  show "x \<le> top" by (simp add: top_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   733
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   734
  { assume "x \<in> A"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   735
    with pextreal_complete_Sup[of A]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   736
    obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   737
    hence "x \<le> s" using `x \<in> A` by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   738
    also have "... = Sup A" using s unfolding Sup_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   739
      by (auto intro!: Least_equality[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   740
    finally show "x \<le> Sup A" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   741
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   742
  { assume "x \<in> A"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   743
    with pextreal_complete_Inf[of A]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   744
    obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   745
    hence "Inf A = i" unfolding Inf_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   746
      by (auto intro!: Greatest_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   747
    also have "i \<le> x" using i `x \<in> A` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   748
    finally show "Inf A \<le> x" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   749
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   750
  { assume *: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   751
    show "Sup A \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   752
    proof (cases "A = {}")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   753
      case True
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   754
      hence "Sup A = 0" unfolding Sup_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   755
        by (auto intro!: Least_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   756
      thus "Sup A \<le> x" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   757
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   758
      case False
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   759
      with pextreal_complete_Sup[of A]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   760
      obtain s where s: "\<forall>y\<in>A. y \<le> s" "\<forall>z. (\<forall>y\<in>A. y \<le> z) \<longrightarrow> s \<le> z" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   761
      hence "Sup A = s"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   762
        unfolding Sup_pextreal_def by (auto intro!: Least_equality)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   763
      also have "s \<le> x" using * s by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   764
      finally show "Sup A \<le> x" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   765
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   766
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   767
  { assume *: "\<And>z. z \<in> A \<Longrightarrow> x \<le> z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   768
    show "x \<le> Inf A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   769
    proof (cases "A = {}")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   770
      case True
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   771
      hence "Inf A = \<omega>" unfolding Inf_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   772
        by (auto intro!: Greatest_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   773
      thus "x \<le> Inf A" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   774
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   775
      case False
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   776
      with pextreal_complete_Inf[of A]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   777
      obtain i where i: "\<forall>y\<in>A. i \<le> y" "\<forall>z. (\<forall>y\<in>A. z \<le> y) \<longrightarrow> z \<le> i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   778
      have "x \<le> i" using * i by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   779
      also have "i = Inf A" using i
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   780
        unfolding Inf_pextreal_def by (auto intro!: Greatest_equality[symmetric])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   781
      finally show "x \<le> Inf A" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   782
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   783
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   784
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   785
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   786
lemma Inf_pextreal_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   787
  fixes z :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   788
  shows "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> (\<exists>x\<in>X. x<y) \<longleftrightarrow> Inf X < y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   789
  by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   790
            order_less_le_trans)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   791
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   792
lemma Inf_greater:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   793
  fixes z :: pextreal assumes "Inf X < z"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   794
  shows "\<exists>x \<in> X. x < z"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   795
proof -
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   796
  have "X \<noteq> {}" using assms by (auto simp: Inf_empty top_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   797
  with assms show ?thesis
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   798
    by (metis Inf_pextreal_iff mem_def not_leE)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   799
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   800
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   801
lemma Inf_close:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   802
  fixes e :: pextreal assumes "Inf X \<noteq> \<omega>" "0 < e"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   803
  shows "\<exists>x \<in> X. x < Inf X + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   804
proof (rule Inf_greater)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   805
  show "Inf X < Inf X + e" using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   806
    by (cases "Inf X", cases e) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   807
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   808
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   809
lemma pextreal_SUPI:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   810
  fixes x :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   811
  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   812
  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> f i \<le> y) \<Longrightarrow> x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   813
  shows "(SUP i:A. f i) = x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   814
  unfolding SUPR_def Sup_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   815
  using assms by (auto intro!: Least_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   816
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   817
lemma Sup_pextreal_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   818
  fixes z :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   819
  shows "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> (\<exists>x\<in>X. y<x) \<longleftrightarrow> y < Sup X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   820
  by (metis complete_lattice_class.Sup_least complete_lattice_class.Sup_upper less_le_not_le linear
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   821
            order_less_le_trans)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   822
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   823
lemma Sup_lesser:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   824
  fixes z :: pextreal assumes "z < Sup X"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   825
  shows "\<exists>x \<in> X. z < x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   826
proof -
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   827
  have "X \<noteq> {}" using assms by (auto simp: Sup_empty bot_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   828
  with assms show ?thesis
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   829
    by (metis Sup_pextreal_iff mem_def not_leE)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   830
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   831
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   832
lemma Sup_eq_\<omega>: "\<omega> \<in> S \<Longrightarrow> Sup S = \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   833
  unfolding Sup_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   834
  by (auto intro!: Least_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   835
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   836
lemma Sup_close:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   837
  assumes "0 < e" and S: "Sup S \<noteq> \<omega>" "S \<noteq> {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   838
  shows "\<exists>X\<in>S. Sup S < X + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   839
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   840
  assume "Sup S = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   841
  moreover obtain X where "X \<in> S" using `S \<noteq> {}` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   842
  ultimately show ?thesis using `0 < e` by (auto intro!: bexI[OF _ `X\<in>S`])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   843
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   844
  assume "Sup S \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   845
  have "\<exists>X\<in>S. Sup S - e < X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   846
  proof (rule Sup_lesser)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   847
    show "Sup S - e < Sup S" using `0 < e` `Sup S \<noteq> 0` `Sup S \<noteq> \<omega>`
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   848
      by (cases e) (auto simp: pextreal_noteq_omega_Ex)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   849
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   850
  then guess X .. note X = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   851
  with `Sup S \<noteq> \<omega>` Sup_eq_\<omega> have "X \<noteq> \<omega>" by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   852
  thus ?thesis using `Sup S \<noteq> \<omega>` X unfolding pextreal_noteq_omega_Ex
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   853
    by (cases e) (auto intro!: bexI[OF _ `X\<in>S`] simp: split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   854
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   855
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   856
lemma Sup_\<omega>: "(SUP i::nat. Real (real i)) = \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   857
proof (rule pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   858
  fix y assume *: "\<And>i::nat. i \<in> UNIV \<Longrightarrow> Real (real i) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   859
  thus "\<omega> \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   860
  proof (cases y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   861
    case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   862
    then obtain k :: nat where "r < real k"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   863
      using ex_less_of_nat by (auto simp: real_eq_of_nat)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   864
    with *[of k] preal show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   865
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   866
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   867
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   868
lemma SUP_\<omega>: "(SUP i:A. f i) = \<omega> \<longleftrightarrow> (\<forall>x<\<omega>. \<exists>i\<in>A. x < f i)"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   869
proof
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   870
  assume *: "(SUP i:A. f i) = \<omega>"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   871
  show "(\<forall>x<\<omega>. \<exists>i\<in>A. x < f i)" unfolding *[symmetric]
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   872
  proof (intro allI impI)
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   873
    fix x assume "x < SUPR A f" then show "\<exists>i\<in>A. x < f i"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   874
      unfolding less_SUP_iff by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   875
  qed
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   876
next
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   877
  assume *: "\<forall>x<\<omega>. \<exists>i\<in>A. x < f i"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   878
  show "(SUP i:A. f i) = \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   879
  proof (rule pextreal_SUPI)
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   880
    fix y assume **: "\<And>i. i \<in> A \<Longrightarrow> f i \<le> y"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   881
    show "\<omega> \<le> y"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   882
    proof cases
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   883
      assume "y < \<omega>"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   884
      from *[THEN spec, THEN mp, OF this]
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   885
      obtain i where "i \<in> A" "\<not> (f i \<le> y)" by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   886
      with ** show ?thesis by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   887
    qed auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   888
  qed auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   889
qed
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
   890
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   891
subsubsection {* Equivalence between @{text "f ----> x"} and @{text SUP} on @{typ pextreal} *}
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   892
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   893
lemma monoseq_monoI: "mono f \<Longrightarrow> monoseq f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   894
  unfolding mono_def monoseq_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   895
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   896
lemma incseq_mono: "mono f \<longleftrightarrow> incseq f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   897
  unfolding mono_def incseq_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   898
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   899
lemma SUP_eq_LIMSEQ:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   900
  assumes "mono f" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   901
  shows "(SUP n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   902
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   903
  assume x: "(SUP n. Real (f n)) = Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   904
  { fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   905
    have "Real (f n) \<le> Real x" using x[symmetric] by (auto intro: le_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   906
    hence "f n \<le> x" using assms by simp }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   907
  show "f ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   908
  proof (rule LIMSEQ_I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   909
    fix r :: real assume "0 < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   910
    show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   911
    proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   912
      assume *: "\<not> ?thesis"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   913
      { fix N
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   914
        from * obtain n where "N \<le> n" "r \<le> x - f n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   915
          using `\<And>n. f n \<le> x` by (auto simp: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   916
        hence "f N \<le> f n" using `mono f` by (auto dest: monoD)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   917
        hence "f N \<le> x - r" using `r \<le> x - f n` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   918
        hence "Real (f N) \<le> Real (x - r)" and "r \<le> x" using `0 \<le> f N` by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   919
      hence "(SUP n. Real (f n)) \<le> Real (x - r)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   920
        and "Real (x - r) < Real x" using `0 < r` by (auto intro: SUP_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   921
      hence "(SUP n. Real (f n)) < Real x" by (rule le_less_trans)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   922
      thus False using x by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   923
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   924
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   925
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   926
  assume "f ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   927
  show "(SUP n. Real (f n)) = Real x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   928
  proof (rule pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   929
    fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   930
    from incseq_le[of f x] `mono f` `f ----> x`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   931
    show "Real (f n) \<le> Real x" using assms incseq_mono by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   932
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   933
    fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> Real (f n) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   934
    show "Real x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   935
    proof (cases y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   936
      case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   937
      with * have "\<exists>N. \<forall>n\<ge>N. f n \<le> r" using assms by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   938
      from LIMSEQ_le_const2[OF `f ----> x` this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   939
      show "Real x \<le> y" using `0 \<le> x` preal by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   940
    qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   941
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   942
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   943
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   944
lemma SUPR_bound:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   945
  assumes "\<forall>N. f N \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   946
  shows "(SUP n. f n) \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   947
  using assms by (simp add: SUPR_def Sup_le_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   948
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   949
lemma pextreal_less_eq_diff_eq_sum:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   950
  fixes x y z :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   951
  assumes "y \<le> x" and "x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   952
  shows "z \<le> x - y \<longleftrightarrow> z + y \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   953
  using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   954
  apply (cases z, cases y, cases x)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   955
  by (simp_all add: field_simps minus_pextreal_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   956
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   957
lemma Real_diff_less_omega: "Real r - x < \<omega>" by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   958
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   959
subsubsection {* Numbers on @{typ pextreal} *}
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   960
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   961
instantiation pextreal :: number
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   962
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   963
definition [simp]: "number_of x = Real (number_of x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   964
instance proof qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   965
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   966
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   967
subsubsection {* Division on @{typ pextreal} *}
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   968
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   969
instantiation pextreal :: inverse
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   970
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   971
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   972
definition "inverse x = pextreal_case (\<lambda>x. if x = 0 then \<omega> else Real (inverse x)) 0 x"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   973
definition [simp]: "x / y = x * inverse (y :: pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   974
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   975
instance proof qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   976
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   977
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   978
lemma pextreal_inverse[simp]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   979
  "inverse 0 = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   980
  "inverse (Real x) = (if x \<le> 0 then \<omega> else Real (inverse x))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   981
  "inverse \<omega> = 0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   982
  "inverse (1::pextreal) = 1"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   983
  "inverse (inverse x) = x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   984
  by (simp_all add: inverse_pextreal_def one_pextreal_def split: pextreal_case_split del: Real_1)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   985
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   986
lemma pextreal_inverse_le_eq:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   987
  assumes "x \<noteq> 0" "x \<noteq> \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
   988
  shows "y \<le> z / x \<longleftrightarrow> x * y \<le> (z :: pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   989
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   990
  from assms obtain r where r: "x = Real r" "0 < r" by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   991
  { fix p q :: real assume "0 \<le> p" "0 \<le> q"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   992
    have "p \<le> q * inverse r \<longleftrightarrow> p \<le> q / r" by (simp add: divide_inverse)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   993
    also have "... \<longleftrightarrow> p * r \<le> q" using `0 < r` by (auto simp: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   994
    finally have "p \<le> q * inverse r \<longleftrightarrow> p * r \<le> q" . }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   995
  with r show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   996
    by (cases y, cases z, auto simp: zero_le_mult_iff field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   997
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   998
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
   999
lemma inverse_antimono_strict:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1000
  fixes x y :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1001
  assumes "x < y" shows "inverse y < inverse x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1002
  using assms by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1003
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1004
lemma inverse_antimono:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1005
  fixes x y :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1006
  assumes "x \<le> y" shows "inverse y \<le> inverse x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1007
  using assms by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1008
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1009
lemma pextreal_inverse_\<omega>_iff[simp]: "inverse x = \<omega> \<longleftrightarrow> x = 0"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1010
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1011
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1012
subsection "Infinite sum over @{typ pextreal}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1013
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1014
text {*
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1015
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1016
The infinite sum over @{typ pextreal} has the nice property that it is always
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1017
defined.
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1018
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1019
*}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1020
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1021
definition psuminf :: "(nat \<Rightarrow> pextreal) \<Rightarrow> pextreal" (binder "\<Sum>\<^isub>\<infinity>" 10) where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1022
  "(\<Sum>\<^isub>\<infinity> x. f x) = (SUP n. \<Sum>i<n. f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1023
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1024
subsubsection {* Equivalence between @{text "\<Sum> n. f n"} and @{text "\<Sum>\<^isub>\<infinity> n. f n"} *}
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1025
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1026
lemma setsum_Real:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1027
  assumes "\<forall>x\<in>A. 0 \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1028
  shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1029
proof (cases "finite A")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1030
  case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1031
  thus ?thesis using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1032
  proof induct case (insert x s)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1033
    hence "0 \<le> setsum f s" apply-apply(rule setsum_nonneg) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1034
    thus ?case using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1035
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1036
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1037
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1038
lemma setsum_Real':
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1039
  assumes "\<forall>x. 0 \<le> f x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1040
  shows "(\<Sum>x\<in>A. Real (f x)) = Real (\<Sum>x\<in>A. f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1041
  apply(rule setsum_Real) using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1042
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1043
lemma setsum_\<omega>:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1044
  "(\<Sum>x\<in>P. f x) = \<omega> \<longleftrightarrow> (finite P \<and> (\<exists>i\<in>P. f i = \<omega>))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1045
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1046
  assume *: "setsum f P = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1047
  show "finite P"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1048
  proof (rule ccontr) assume "infinite P" with * show False by auto qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1049
  show "\<exists>i\<in>P. f i = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1050
  proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1051
    assume "\<not> ?thesis" hence "\<And>i. i\<in>P \<Longrightarrow> f i \<noteq> \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1052
    from `finite P` this have "setsum f P \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1053
      by induct auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1054
    with * show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1055
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1056
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1057
  fix i assume "finite P" "i \<in> P" "f i = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1058
  thus "setsum f P = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1059
  proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1060
    case (insert x A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1061
    show ?case using insert by (cases "x = i") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1062
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1063
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1064
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1065
lemma real_of_pextreal_setsum:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1066
  assumes "\<And>x. x \<in> S \<Longrightarrow> f x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1067
  shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1068
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1069
  assume "finite S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1070
  from this assms show ?thesis
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1071
    by induct (simp_all add: real_of_pextreal_add setsum_\<omega>)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1072
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1073
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1074
lemma setsum_0:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1075
  fixes f :: "'a \<Rightarrow> pextreal" assumes "finite A"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1076
  shows "(\<Sum>x\<in>A. f x) = 0 \<longleftrightarrow> (\<forall>i\<in>A. f i = 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1077
  using assms by induct auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1078
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1079
lemma suminf_imp_psuminf:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1080
  assumes "f sums x" and "\<forall>n. 0 \<le> f n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1081
  shows "(\<Sum>\<^isub>\<infinity> x. Real (f x)) = Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1082
  unfolding psuminf_def setsum_Real'[OF assms(2)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1083
proof (rule SUP_eq_LIMSEQ[THEN iffD2])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1084
  show "mono (\<lambda>n. setsum f {..<n})" (is "mono ?S")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1085
    unfolding mono_iff_le_Suc using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1086
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1087
  { fix n show "0 \<le> ?S n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1088
      using setsum_nonneg[of "{..<n}" f] assms by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1089
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1090
  thus "0 \<le> x" "?S ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1091
    using `f sums x` LIMSEQ_le_const
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1092
    by (auto simp: atLeast0LessThan sums_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1093
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1094
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1095
lemma psuminf_equality:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1096
  assumes "\<And>n. setsum f {..<n} \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1097
  and "\<And>y. y \<noteq> \<omega> \<Longrightarrow> (\<And>n. setsum f {..<n} \<le> y) \<Longrightarrow> x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1098
  shows "psuminf f = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1099
  unfolding psuminf_def
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1100
proof (safe intro!: pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1101
  fix n show "setsum f {..<n} \<le> x" using assms(1) .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1102
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1103
  fix y assume *: "\<forall>n. n \<in> UNIV \<longrightarrow> setsum f {..<n} \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1104
  show "x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1105
  proof (cases "y = \<omega>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1106
    assume "y \<noteq> \<omega>" from assms(2)[OF this] *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1107
    show "x \<le> y" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1108
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1109
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1110
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1111
lemma psuminf_\<omega>:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1112
  assumes "f i = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1113
  shows "(\<Sum>\<^isub>\<infinity> x. f x) = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1114
proof (rule psuminf_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1115
  fix y assume *: "\<And>n. setsum f {..<n} \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1116
  have "setsum f {..<Suc i} = \<omega>" 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1117
    using assms by (simp add: setsum_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1118
  thus "\<omega> \<le> y" using *[of "Suc i"] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1119
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1120
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1121
lemma psuminf_imp_suminf:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1122
  assumes "(\<Sum>\<^isub>\<infinity> x. f x) \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1123
  shows "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity> x. f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1124
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1125
  have "\<forall>i. \<exists>r. f i = Real r \<and> 0 \<le> r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1126
  proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1127
    fix i show "\<exists>r. f i = Real r \<and> 0 \<le> r" using psuminf_\<omega> assms by (cases "f i") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1128
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1129
  from choice[OF this] obtain r where f: "f = (\<lambda>i. Real (r i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1130
    and pos: "\<forall>i. 0 \<le> r i"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1131
    by (auto simp: fun_eq_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1132
  hence [simp]: "\<And>i. real (f i) = r i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1133
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1134
  have "mono (\<lambda>n. setsum r {..<n})" (is "mono ?S")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1135
    unfolding mono_iff_le_Suc using pos by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1136
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1137
  { fix n have "0 \<le> ?S n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1138
      using setsum_nonneg[of "{..<n}" r] pos by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1139
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1140
  from assms obtain p where *: "(\<Sum>\<^isub>\<infinity> x. f x) = Real p" and "0 \<le> p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1141
    by (cases "(\<Sum>\<^isub>\<infinity> x. f x)") auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1142
  show ?thesis unfolding * using * pos `0 \<le> p` SUP_eq_LIMSEQ[OF `mono ?S` `\<And>n. 0 \<le> ?S n` `0 \<le> p`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1143
    by (simp add: f atLeast0LessThan sums_def psuminf_def setsum_Real'[OF pos] f)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1144
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1145
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1146
lemma psuminf_bound:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1147
  assumes "\<forall>N. (\<Sum>n<N. f n) \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1148
  shows "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1149
  using assms by (simp add: psuminf_def SUPR_def Sup_le_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1150
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1151
lemma psuminf_bound_add:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1152
  assumes "\<forall>N. (\<Sum>n<N. f n) + y \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1153
  shows "(\<Sum>\<^isub>\<infinity> n. f n) + y \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1154
proof (cases "x = \<omega>")
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1155
  have "y \<le> x" using assms by (auto intro: pextreal_le_add2)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1156
  assume "x \<noteq> \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1157
  note move_y = pextreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1158
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1159
  have "\<forall>N. (\<Sum>n<N. f n) \<le> x - y" using assms by (simp add: move_y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1160
  hence "(\<Sum>\<^isub>\<infinity> n. f n) \<le> x - y" by (rule psuminf_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1161
  thus ?thesis by (simp add: move_y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1162
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1163
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1164
lemma psuminf_finite:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1165
  assumes "\<forall>N\<ge>n. f N = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1166
  shows "(\<Sum>\<^isub>\<infinity> n. f n) = (\<Sum>N<n. f N)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1167
proof (rule psuminf_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1168
  fix N
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1169
  show "setsum f {..<N} \<le> setsum f {..<n}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1170
  proof (cases rule: linorder_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1171
    assume "N < n" thus ?thesis by (auto intro!: setsum_mono3)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1172
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1173
    assume "n < N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1174
    hence *: "{..<N} = {..<n} \<union> {n..<N}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1175
    moreover have "setsum f {n..<N} = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1176
      using assms by (auto intro!: setsum_0')
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1177
    ultimately show ?thesis unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1178
      by (subst setsum_Un_disjoint) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1179
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1180
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1181
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1182
lemma psuminf_upper:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1183
  shows "(\<Sum>n<N. f n) \<le> (\<Sum>\<^isub>\<infinity> n. f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1184
  unfolding psuminf_def SUPR_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1185
  by (auto intro: complete_lattice_class.Sup_upper image_eqI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1186
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1187
lemma psuminf_le:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1188
  assumes "\<And>N. f N \<le> g N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1189
  shows "psuminf f \<le> psuminf g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1190
proof (safe intro!: psuminf_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1191
  fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1192
  have "setsum f {..<n} \<le> setsum g {..<n}" using assms by (auto intro: setsum_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1193
  also have "... \<le> psuminf g" by (rule psuminf_upper)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1194
  finally show "setsum f {..<n} \<le> psuminf g" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1195
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1196
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1197
lemma psuminf_const[simp]: "psuminf (\<lambda>n. c) = (if c = 0 then 0 else \<omega>)" (is "_ = ?if")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1198
proof (rule psuminf_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1199
  fix y assume *: "\<And>n :: nat. (\<Sum>n<n. c) \<le> y" and "y \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1200
  then obtain r p where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1201
    y: "y = Real r" "0 \<le> r" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1202
    c: "c = Real p" "0 \<le> p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1203
      using *[of 1] by (cases c, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1204
  show "(if c = 0 then 0 else \<omega>) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1205
  proof (cases "p = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1206
    assume "p = 0" with c show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1207
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1208
    assume "p \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1209
    with * c y have **: "\<And>n :: nat. real n \<le> r / p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1210
      by (auto simp: zero_le_mult_iff field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1211
    from ex_less_of_nat[of "r / p"] guess n ..
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1212
    with **[of n] show ?thesis by (simp add: real_eq_of_nat)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1213
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1214
qed (cases "c = 0", simp_all)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1215
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1216
lemma psuminf_add[simp]: "psuminf (\<lambda>n. f n + g n) = psuminf f + psuminf g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1217
proof (rule psuminf_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1218
  fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1219
  from psuminf_upper[of f n] psuminf_upper[of g n]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1220
  show "(\<Sum>n<n. f n + g n) \<le> psuminf f + psuminf g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1221
    by (auto simp add: setsum_addf intro!: add_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1222
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1223
  fix y assume *: "\<And>n. (\<Sum>n<n. f n + g n) \<le> y" and "y \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1224
  { fix n m
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1225
    have **: "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1226
    proof (cases rule: linorder_le_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1227
      assume "n \<le> m"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1228
      hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<m. f n) + (\<Sum>n<m. g n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1229
        by (auto intro!: add_right_mono setsum_mono3)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1230
      also have "... \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1231
        using * by (simp add: setsum_addf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1232
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1233
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1234
      assume "m \<le> n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1235
      hence "(\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> (\<Sum>n<n. f n) + (\<Sum>n<n. g n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1236
        by (auto intro!: add_left_mono setsum_mono3)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1237
      also have "... \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1238
        using * by (simp add: setsum_addf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1239
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1240
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1241
  hence "\<And>m. \<forall>n. (\<Sum>n<n. f n) + (\<Sum>n<m. g n) \<le> y" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1242
  from psuminf_bound_add[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1243
  have "\<forall>m. (\<Sum>n<m. g n) + psuminf f \<le> y" by (simp add: ac_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1244
  from psuminf_bound_add[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1245
  show "psuminf f + psuminf g \<le> y" by (simp add: ac_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1246
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1247
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1248
lemma psuminf_0: "psuminf f = 0 \<longleftrightarrow> (\<forall>i. f i = 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1249
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1250
  assume "\<forall>i. f i = 0"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  1251
  hence "f = (\<lambda>i. 0)" by (simp add: fun_eq_iff)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1252
  thus "psuminf f = 0" using psuminf_const by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1253
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1254
  fix i assume "psuminf f = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1255
  hence "(\<Sum>n<Suc i. f n) = 0" using psuminf_upper[of f "Suc i"] by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1256
  thus "f i = 0" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1257
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1258
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1259
lemma psuminf_cmult_right[simp]: "psuminf (\<lambda>n. c * f n) = c * psuminf f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1260
proof (rule psuminf_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1261
  fix n show "(\<Sum>n<n. c * f n) \<le> c * psuminf f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1262
   by (auto simp: setsum_right_distrib[symmetric] intro: mult_left_mono psuminf_upper)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1263
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1264
  fix y
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1265
  assume "\<And>n. (\<Sum>n<n. c * f n) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1266
  hence *: "\<And>n. c * (\<Sum>n<n. f n) \<le> y" by (auto simp add: setsum_right_distrib)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1267
  thus "c * psuminf f \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1268
  proof (cases "c = \<omega> \<or> c = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1269
    assume "c = \<omega> \<or> c = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1270
    thus ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1271
      using * by (fastsimp simp add: psuminf_0 setsum_0 split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1272
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1273
    assume "\<not> (c = \<omega> \<or> c = 0)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1274
    hence "c \<noteq> 0" "c \<noteq> \<omega>" by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1275
    note rewrite_div = pextreal_inverse_le_eq[OF this, of _ y]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1276
    hence "\<forall>n. (\<Sum>n<n. f n) \<le> y / c" using * by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1277
    hence "psuminf f \<le> y / c" by (rule psuminf_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1278
    thus ?thesis using rewrite_div by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1279
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1280
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1281
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1282
lemma psuminf_cmult_left[simp]: "psuminf (\<lambda>n. f n * c) = psuminf f * c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1283
  using psuminf_cmult_right[of c f] by (simp add: ac_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1284
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1285
lemma psuminf_half_series: "psuminf (\<lambda>n. (1/2)^Suc n) = 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1286
  using suminf_imp_psuminf[OF power_half_series] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1287
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1288
lemma setsum_pinfsum: "(\<Sum>\<^isub>\<infinity> n. \<Sum>m\<in>A. f n m) = (\<Sum>m\<in>A. (\<Sum>\<^isub>\<infinity> n. f n m))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1289
proof (cases "finite A")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1290
  assume "finite A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1291
  thus ?thesis by induct simp_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1292
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1293
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1294
lemma psuminf_reindex:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1295
  fixes f:: "nat \<Rightarrow> nat" assumes "bij f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1296
  shows "psuminf (g \<circ> f) = psuminf g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1297
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1298
  have [intro, simp]: "\<And>A. inj_on f A" using `bij f` unfolding bij_def by (auto intro: subset_inj_on)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1299
  have f[intro, simp]: "\<And>x. f (inv f x) = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1300
    using `bij f` unfolding bij_def by (auto intro: surj_f_inv_f)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1301
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1302
  proof (rule psuminf_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1303
    fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1304
    have "setsum (g \<circ> f) {..<n} = setsum g (f ` {..<n})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1305
      by (simp add: setsum_reindex)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1306
    also have "\<dots> \<le> setsum g {..Max (f ` {..<n})}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1307
      by (rule setsum_mono3) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1308
    also have "\<dots> \<le> psuminf g" unfolding lessThan_Suc_atMost[symmetric] by (rule psuminf_upper)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1309
    finally show "setsum (g \<circ> f) {..<n} \<le> psuminf g" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1310
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1311
    fix y assume *: "\<And>n. setsum (g \<circ> f) {..<n} \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1312
    show "psuminf g \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1313
    proof (safe intro!: psuminf_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1314
      fix N
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1315
      have "setsum g {..<N} \<le> setsum g (f ` {..Max (inv f ` {..<N})})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1316
        by (rule setsum_mono3) (auto intro!: image_eqI[where f="f", OF f[symmetric]])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1317
      also have "\<dots> = setsum (g \<circ> f) {..Max (inv f ` {..<N})}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1318
        by (simp add: setsum_reindex)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1319
      also have "\<dots> \<le> y" unfolding lessThan_Suc_atMost[symmetric] by (rule *)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1320
      finally show "setsum g {..<N} \<le> y" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1321
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1322
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1323
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1324
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1325
lemma pextreal_mult_less_right:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1326
  assumes "b * a < c * a" "0 < a" "a < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1327
  shows "b < c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1328
  using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1329
  by (cases a, cases b, cases c) (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1330
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1331
lemma pextreal_\<omega>_eq_plus[simp]: "\<omega> = a + b \<longleftrightarrow> (a = \<omega> \<or> b = \<omega>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1332
  by (cases a, cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1333
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1334
lemma pextreal_of_nat_le_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1335
  "(of_nat k :: pextreal) \<le> of_nat m \<longleftrightarrow> k \<le> m" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1336
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1337
lemma pextreal_of_nat_less_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1338
  "(of_nat k :: pextreal) < of_nat m \<longleftrightarrow> k < m" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1339
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1340
lemma pextreal_bound_add:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1341
  assumes "\<forall>N. f N + y \<le> (x::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1342
  shows "(SUP n. f n) + y \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1343
proof (cases "x = \<omega>")
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1344
  have "y \<le> x" using assms by (auto intro: pextreal_le_add2)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1345
  assume "x \<noteq> \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1346
  note move_y = pextreal_less_eq_diff_eq_sum[OF `y \<le> x` this]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1347
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1348
  have "\<forall>N. f N \<le> x - y" using assms by (simp add: move_y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1349
  hence "(SUP n. f n) \<le> x - y" by (rule SUPR_bound)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1350
  thus ?thesis by (simp add: move_y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1351
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1352
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1353
lemma SUPR_pextreal_add:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1354
  fixes f g :: "nat \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1355
  assumes f: "\<forall>n. f n \<le> f (Suc n)" and g: "\<forall>n. g n \<le> g (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1356
  shows "(SUP n. f n + g n) = (SUP n. f n) + (SUP n. g n)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1357
proof (rule pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1358
  fix n :: nat from le_SUPI[of n UNIV f] le_SUPI[of n UNIV g]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1359
  show "f n + g n \<le> (SUP n. f n) + (SUP n. g n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1360
    by (auto intro!: add_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1361
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1362
  fix y assume *: "\<And>n. n \<in> UNIV \<Longrightarrow> f n + g n \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1363
  { fix n m
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1364
    have "f n + g m \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1365
    proof (cases rule: linorder_le_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1366
      assume "n \<le> m"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1367
      hence "f n + g m \<le> f m + g m"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1368
        using f lift_Suc_mono_le by (auto intro!: add_right_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1369
      also have "\<dots> \<le> y" using * by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1370
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1371
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1372
      assume "m \<le> n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1373
      hence "f n + g m \<le> f n + g n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1374
        using g lift_Suc_mono_le by (auto intro!: add_left_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1375
      also have "\<dots> \<le> y" using * by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1376
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1377
    qed }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1378
  hence "\<And>m. \<forall>n. f n + g m \<le> y" by simp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1379
  from pextreal_bound_add[OF this]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1380
  have "\<forall>m. (g m) + (SUP n. f n) \<le> y" by (simp add: ac_simps)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1381
  from pextreal_bound_add[OF this]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1382
  show "SUPR UNIV f + SUPR UNIV g \<le> y" by (simp add: ac_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1383
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1384
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1385
lemma SUPR_pextreal_setsum:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1386
  fixes f :: "'x \<Rightarrow> nat \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1387
  assumes "\<And>i. i \<in> P \<Longrightarrow> \<forall>n. f i n \<le> f i (Suc n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1388
  shows "(SUP n. \<Sum>i\<in>P. f i n) = (\<Sum>i\<in>P. SUP n. f i n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1389
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1390
  assume "finite P" from this assms show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1391
  proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1392
    case (insert i P)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1393
    thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1394
      apply simp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1395
      apply (subst SUPR_pextreal_add)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1396
      by (auto intro!: setsum_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1397
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1398
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1399
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1400
lemma psuminf_SUP_eq:
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1401
  assumes "\<And>n i. f n i \<le> f (Suc n) i"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1402
  shows "(\<Sum>\<^isub>\<infinity> i. SUP n::nat. f n i) = (SUP n::nat. \<Sum>\<^isub>\<infinity> i. f n i)"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1403
proof -
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1404
  { fix n :: nat
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1405
    have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1406
      using assms by (auto intro!: SUPR_pextreal_setsum[symmetric]) }
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1407
  note * = this
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1408
  show ?thesis
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1409
    unfolding psuminf_def
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1410
    unfolding *
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1411
    apply (subst SUP_commute) ..
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1412
qed
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1413
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1414
lemma psuminf_commute:
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1415
  shows "(\<Sum>\<^isub>\<infinity> i j. f i j) = (\<Sum>\<^isub>\<infinity> j i. f i j)"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1416
proof -
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1417
  have "(SUP n. \<Sum> i < n. SUP m. \<Sum> j < m. f i j) = (SUP n. SUP m. \<Sum> i < n. \<Sum> j < m. f i j)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1418
    apply (subst SUPR_pextreal_setsum)
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1419
    by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1420
  also have "\<dots> = (SUP m n. \<Sum> j < m. \<Sum> i < n. f i j)"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1421
    apply (subst SUP_commute)
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1422
    apply (subst setsum_commute)
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1423
    by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1424
  also have "\<dots> = (SUP m. \<Sum> j < m. SUP n. \<Sum> i < n. f i j)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1425
    apply (subst SUPR_pextreal_setsum)
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1426
    by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1427
  finally show ?thesis
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1428
    unfolding psuminf_def by auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1429
qed
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  1430
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1431
lemma psuminf_2dimen:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1432
  fixes f:: "nat * nat \<Rightarrow> pextreal"
40872
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1433
  assumes fsums: "\<And>m. g m = (\<Sum>\<^isub>\<infinity> n. f (m,n))"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1434
  shows "psuminf (f \<circ> prod_decode) = psuminf g"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1435
proof (rule psuminf_equality)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1436
  fix n :: nat
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1437
  let ?P = "prod_decode ` {..<n}"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1438
  have "setsum (f \<circ> prod_decode) {..<n} = setsum f ?P"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1439
    by (auto simp: setsum_reindex inj_prod_decode)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1440
  also have "\<dots> \<le> setsum f ({..Max (fst ` ?P)} \<times> {..Max (snd ` ?P)})"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1441
  proof (safe intro!: setsum_mono3 Max_ge image_eqI)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1442
    fix a b x assume "(a, b) = prod_decode x"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1443
    from this[symmetric] show "a = fst (prod_decode x)" "b = snd (prod_decode x)"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1444
      by simp_all
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1445
  qed simp_all
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1446
  also have "\<dots> = (\<Sum>m\<le>Max (fst ` ?P). (\<Sum>n\<le>Max (snd ` ?P). f (m,n)))"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1447
    unfolding setsum_cartesian_product by simp
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1448
  also have "\<dots> \<le> (\<Sum>m\<le>Max (fst ` ?P). g m)"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1449
    by (auto intro!: setsum_mono psuminf_upper simp del: setsum_lessThan_Suc
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1450
        simp: fsums lessThan_Suc_atMost[symmetric])
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1451
  also have "\<dots> \<le> psuminf g"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1452
    by (auto intro!: psuminf_upper simp del: setsum_lessThan_Suc
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1453
        simp: lessThan_Suc_atMost[symmetric])
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1454
  finally show "setsum (f \<circ> prod_decode) {..<n} \<le> psuminf g" .
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1455
next
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1456
  fix y assume *: "\<And>n. setsum (f \<circ> prod_decode) {..<n} \<le> y"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1457
  have g: "g = (\<lambda>m. \<Sum>\<^isub>\<infinity> n. f (m,n))" unfolding fsums[symmetric] ..
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1458
  show "psuminf g \<le> y" unfolding g
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1459
  proof (rule psuminf_bound, unfold setsum_pinfsum[symmetric], safe intro!: psuminf_bound)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1460
    fix N M :: nat
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1461
    let ?P = "{..<N} \<times> {..<M}"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1462
    let ?M = "Max (prod_encode ` ?P)"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1463
    have "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> (\<Sum>(m, n)\<in>?P. f (m, n))"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1464
      unfolding setsum_commute[of _ _ "{..<M}"] unfolding setsum_cartesian_product ..
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1465
    also have "\<dots> \<le> (\<Sum>(m,n)\<in>(prod_decode ` {..?M}). f (m, n))"
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1466
      by (auto intro!: setsum_mono3 image_eqI[where f=prod_decode, OF prod_encode_inverse[symmetric]])
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1467
    also have "\<dots> \<le> y" using *[of "Suc ?M"]
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1468
      by (simp add: lessThan_Suc_atMost[symmetric] setsum_reindex
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1469
               inj_prod_decode del: setsum_lessThan_Suc)
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1470
    finally show "(\<Sum>n<M. \<Sum>m<N. f (m, n)) \<le> y" .
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1471
  qed
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1472
qed
7c556a9240de Move SUP_commute, SUP_less_iff to HOL image;
hoelzl
parents: 40871
diff changeset
  1473
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1474
lemma Real_max:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1475
  assumes "x \<ge> 0" "y \<ge> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1476
  shows "Real (max x y) = max (Real x) (Real y)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1477
  using assms unfolding max_def by (auto simp add:not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1478
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1479
lemma Real_real: "Real (real x) = (if x = \<omega> then 0 else x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1480
  using assms by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1481
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1482
lemma inj_on_real: "inj_on real (UNIV - {\<omega>})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1483
proof (rule inj_onI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1484
  fix x y assume mem: "x \<in> UNIV - {\<omega>}" "y \<in> UNIV - {\<omega>}" and "real x = real y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1485
  thus "x = y" by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1486
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1487
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1488
lemma inj_on_Real: "inj_on Real {0..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1489
  by (auto intro!: inj_onI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1490
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1491
lemma range_Real[simp]: "range Real = UNIV - {\<omega>}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1492
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1493
  fix x assume "x \<notin> range Real"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1494
  thus "x = \<omega>" by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1495
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1496
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1497
lemma image_Real[simp]: "Real ` {0..} = UNIV - {\<omega>}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1498
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1499
  fix x assume "x \<notin> Real ` {0..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1500
  thus "x = \<omega>" by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1501
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1502
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1503
lemma pextreal_SUP_cmult:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1504
  fixes f :: "'a \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1505
  shows "(SUP i : R. z * f i) = z * (SUP i : R. f i)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1506
proof (rule pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1507
  fix i assume "i \<in> R"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1508
  from le_SUPI[OF this]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1509
  show "z * f i \<le> z * (SUP i:R. f i)" by (rule pextreal_mult_cancel)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1510
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1511
  fix y assume "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1512
  hence *: "\<And>i. i\<in>R \<Longrightarrow> z * f i \<le> y" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1513
  show "z * (SUP i:R. f i) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1514
  proof (cases "\<forall>i\<in>R. f i = 0")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1515
    case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1516
    show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1517
    proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1518
      assume "R \<noteq> {}" hence "f ` R = {0}" using True by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1519
      thus ?thesis by (simp add: SUPR_def)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1520
    qed (simp add: SUPR_def Sup_empty bot_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1521
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1522
    case False then obtain i where i: "i \<in> R" and f0: "f i \<noteq> 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1523
    show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1524
    proof (cases "z = 0 \<or> z = \<omega>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1525
      case True with f0 *[OF i] show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1526
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1527
      case False hence z: "z \<noteq> 0" "z \<noteq> \<omega>" by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1528
      note div = pextreal_inverse_le_eq[OF this, symmetric]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1529
      hence "\<And>i. i\<in>R \<Longrightarrow> f i \<le> y / z" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1530
      thus ?thesis unfolding div SUP_le_iff by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1531
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1532
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1533
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1534
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1535
instantiation pextreal :: topological_space
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1536
begin
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1537
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1538
definition "open A \<longleftrightarrow>
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1539
  (\<exists>T. open T \<and> (Real ` (T\<inter>{0..}) = A - {\<omega>})) \<and> (\<omega> \<in> A \<longrightarrow> (\<exists>x\<ge>0. {Real x <..} \<subseteq> A))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1540
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1541
lemma open_omega: "open A \<Longrightarrow> \<omega> \<in> A \<Longrightarrow> (\<exists>x\<ge>0. {Real x<..} \<subseteq> A)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1542
  unfolding open_pextreal_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1543
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1544
lemma open_omegaD: assumes "open A" "\<omega> \<in> A" obtains x where "x\<ge>0" "{Real x<..} \<subseteq> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1545
  using open_omega[OF assms] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1546
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1547
lemma pextreal_openE: assumes "open A" obtains A' x where
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1548
  "open A'" "Real ` (A' \<inter> {0..}) = A - {\<omega>}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1549
  "x \<ge> 0" "\<omega> \<in> A \<Longrightarrow> {Real x<..} \<subseteq> A"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1550
  using assms open_pextreal_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1551
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1552
instance
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1553
proof
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1554
  let ?U = "UNIV::pextreal set"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1555
  show "open ?U" unfolding open_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1556
    by (auto intro!: exI[of _ "UNIV"] exI[of _ 0])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1557
next
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1558
  fix S T::"pextreal set" assume "open S" and "open T"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1559
  from `open S`[THEN pextreal_openE] guess S' xS . note S' = this
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1560
  from `open T`[THEN pextreal_openE] guess T' xT . note T' = this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1561
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1562
  from S'(1-3) T'(1-3)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1563
  show "open (S \<inter> T)" unfolding open_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1564
  proof (safe intro!: exI[of _ "S' \<inter> T'"] exI[of _ "max xS xT"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1565
    fix x assume *: "Real (max xS xT) < x" and "\<omega> \<in> S" "\<omega> \<in> T"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1566
    from `\<omega> \<in> S`[THEN S'(4)] * show "x \<in> S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1567
      by (cases x, auto simp: max_def split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1568
    from `\<omega> \<in> T`[THEN T'(4)] * show "x \<in> T"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1569
      by (cases x, auto simp: max_def split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1570
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1571
    fix x assume x: "x \<notin> Real ` (S' \<inter> T' \<inter> {0..})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1572
    have *: "S' \<inter> T' \<inter> {0..} = (S' \<inter> {0..}) \<inter> (T' \<inter> {0..})" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1573
    assume "x \<in> T" "x \<in> S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1574
    with S'(2) T'(2) show "x = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1575
      using x[unfolded *] inj_on_image_Int[OF inj_on_Real] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1576
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1577
next
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1578
  fix K assume openK: "\<forall>S \<in> K. open (S:: pextreal set)"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1579
  hence "\<forall>S\<in>K. \<exists>T. open T \<and> Real ` (T \<inter> {0..}) = S - {\<omega>}" by (auto simp: open_pextreal_def)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1580
  from bchoice[OF this] guess T .. note T = this[rule_format]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1581
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1582
  show "open (\<Union>K)" unfolding open_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1583
  proof (safe intro!: exI[of _ "\<Union>(T ` K)"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1584
    fix x S assume "0 \<le> x" "x \<in> T S" "S \<in> K"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1585
    with T[OF `S \<in> K`] show "Real x \<in> \<Union>K" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1586
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1587
    fix x S assume x: "x \<notin> Real ` (\<Union>T ` K \<inter> {0..})" "S \<in> K" "x \<in> S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1588
    hence "x \<notin> Real ` (T S \<inter> {0..})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1589
      by (auto simp: image_UN UN_simps[symmetric] simp del: UN_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1590
    thus "x = \<omega>" using T[OF `S \<in> K`] `x \<in> S` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1591
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1592
    fix S assume "\<omega> \<in> S" "S \<in> K"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1593
    from openK[rule_format, OF `S \<in> K`, THEN pextreal_openE] guess S' x .
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1594
    from this(3, 4) `\<omega> \<in> S`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1595
    show "\<exists>x\<ge>0. {Real x<..} \<subseteq> \<Union>K"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1596
      by (auto intro!: exI[of _ x] bexI[OF _ `S \<in> K`])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1597
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1598
    from T[THEN conjunct1] show "open (\<Union>T`K)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1599
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1600
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1601
end
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1602
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1603
lemma open_pextreal_lessThan[simp]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1604
  "open {..< a :: pextreal}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1605
proof (cases a)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1606
  case (preal x) thus ?thesis unfolding open_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1607
  proof (safe intro!: exI[of _ "{..< x}"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1608
    fix y assume "y < Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1609
    moreover assume "y \<notin> Real ` ({..<x} \<inter> {0..})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1610
    ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1611
    thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1612
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1613
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1614
  case infinite thus ?thesis
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1615
    unfolding open_pextreal_def by (auto intro!: exI[of _ UNIV])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1616
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1617
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1618
lemma open_pextreal_greaterThan[simp]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1619
  "open {a :: pextreal <..}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1620
proof (cases a)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1621
  case (preal x) thus ?thesis unfolding open_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1622
  proof (safe intro!: exI[of _ "{x <..}"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1623
    fix y assume "Real x < y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1624
    moreover assume "y \<notin> Real ` ({x<..} \<inter> {0..})"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1625
    ultimately have "y \<noteq> Real (real y)" using preal by (cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1626
    thus "y = \<omega>" by (auto simp: Real_real split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1627
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1628
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1629
  case infinite thus ?thesis
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1630
    unfolding open_pextreal_def by (auto intro!: exI[of _ "{}"])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1631
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1632
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1633
lemma pextreal_open_greaterThanLessThan[simp]: "open {a::pextreal <..< b}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1634
  unfolding greaterThanLessThan_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1635
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1636
lemma closed_pextreal_atLeast[simp, intro]: "closed {a :: pextreal ..}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1637
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1638
  have "- {a ..} = {..< a}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1639
  then show "closed {a ..}"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1640
    unfolding closed_def using open_pextreal_lessThan by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1641
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1642
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1643
lemma closed_pextreal_atMost[simp, intro]: "closed {.. b :: pextreal}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1644
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1645
  have "- {.. b} = {b <..}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1646
  then show "closed {.. b}" 
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1647
    unfolding closed_def using open_pextreal_greaterThan by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1648
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1649
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1650
lemma closed_pextreal_atLeastAtMost[simp, intro]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1651
  shows "closed {a :: pextreal .. b}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1652
  unfolding atLeastAtMost_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1653
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1654
lemma pextreal_dense:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1655
  fixes x y :: pextreal assumes "x < y"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1656
  shows "\<exists>z. x < z \<and> z < y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1657
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1658
  from `x < y` obtain p where p: "x = Real p" "0 \<le> p" by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1659
  show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1660
  proof (cases y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1661
    case (preal r) with p `x < y` have "p < r" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1662
    with dense obtain z where "p < z" "z < r" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1663
    thus ?thesis using preal p by (auto intro!: exI[of _ "Real z"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1664
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1665
    case infinite thus ?thesis using `x < y` p
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1666
      by (auto intro!: exI[of _ "Real p + 1"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1667
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1668
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1669
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1670
instance pextreal :: t2_space
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1671
proof
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1672
  fix x y :: pextreal assume "x \<noteq> y"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1673
  let "?P x (y::pextreal)" = "\<exists> U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1674
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1675
  { fix x y :: pextreal assume "x < y"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1676
    from pextreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1677
    have "?P x y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1678
      apply (rule exI[of _ "{..<z}"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1679
      apply (rule exI[of _ "{z<..}"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1680
      using z by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1681
  note * = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1682
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1683
  from `x \<noteq> y`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1684
  show "\<exists>U V. open U \<and> open V \<and> x \<in> U \<and> y \<in> V \<and> U \<inter> V = {}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1685
  proof (cases rule: linorder_cases)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1686
    assume "x = y" with `x \<noteq> y` show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1687
  next assume "x < y" from *[OF this] show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1688
  next assume "y < x" from *[OF this] show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1689
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1690
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1691
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1692
definition (in complete_lattice) isoton :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<up>" 50) where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1693
  "A \<up> X \<longleftrightarrow> (\<forall>i. A i \<le> A (Suc i)) \<and> (SUP i. A i) = X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1694
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1695
definition (in complete_lattice) antiton (infix "\<down>" 50) where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1696
  "A \<down> X \<longleftrightarrow> (\<forall>i. A i \<ge> A (Suc i)) \<and> (INF i. A i) = X"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1697
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1698
lemma isotoneI[intro?]: "\<lbrakk> \<And>i. f i \<le> f (Suc i) ; (SUP i. f i) = F \<rbrakk> \<Longrightarrow> f \<up> F"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1699
  unfolding isoton_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1700
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1701
lemma (in complete_lattice) isotonD[dest]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1702
  assumes "A \<up> X" shows "A i \<le> A (Suc i)" "(SUP i. A i) = X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1703
  using assms unfolding isoton_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1704
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1705
lemma isotonD'[dest]:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1706
  assumes "(A::_=>_) \<up> X" shows "A i x \<le> A (Suc i) x" "(SUP i. A i) = X"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1707
  using assms unfolding isoton_def le_fun_def by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1708
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1709
lemma isoton_mono_le:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1710
  assumes "f \<up> x" "i \<le> j"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1711
  shows "f i \<le> f j"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1712
  using `f \<up> x`[THEN isotonD(1)] lift_Suc_mono_le[of f, OF _ `i \<le> j`] by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1713
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1714
lemma isoton_const:
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1715
  shows "(\<lambda> i. c) \<up> c"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1716
unfolding isoton_def by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1717
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1718
lemma isoton_cmult_right:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1719
  assumes "f \<up> (x::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1720
  shows "(\<lambda>i. c * f i) \<up> (c * x)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1721
  using assms unfolding isoton_def pextreal_SUP_cmult
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1722
  by (auto intro: pextreal_mult_cancel)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1723
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1724
lemma isoton_cmult_left:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1725
  "f \<up> (x::pextreal) \<Longrightarrow> (\<lambda>i. f i * c) \<up> (x * c)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1726
  by (subst (1 2) mult_commute) (rule isoton_cmult_right)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1727
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1728
lemma isoton_add:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1729
  assumes "f \<up> (x::pextreal)" and "g \<up> y"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1730
  shows "(\<lambda>i. f i + g i) \<up> (x + y)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1731
  using assms unfolding isoton_def
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1732
  by (auto intro: pextreal_mult_cancel add_mono simp: SUPR_pextreal_add)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1733
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1734
lemma isoton_fun_expand:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1735
  "f \<up> x \<longleftrightarrow> (\<forall>i. (\<lambda>j. f j i) \<up> (x i))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1736
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1737
  have "\<And>i. {y. \<exists>f'\<in>range f. y = f' i} = range (\<lambda>j. f j i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1738
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1739
  with assms show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1740
    by (auto simp add: isoton_def le_fun_def Sup_fun_def SUPR_def)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1741
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1742
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1743
lemma isoton_indicator:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1744
  assumes "f \<up> g"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1745
  shows "(\<lambda>i x. f i x * indicator A x) \<up> (\<lambda>x. g x * indicator A x :: pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1746
  using assms unfolding isoton_fun_expand by (auto intro!: isoton_cmult_left)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1747
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1748
lemma isoton_setsum:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1749
  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> pextreal"
40859
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1750
  assumes "finite A" "A \<noteq> {}"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1751
  assumes "\<And> x. x \<in> A \<Longrightarrow> f x \<up> y x"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1752
  shows "(\<lambda> i. (\<Sum> x \<in> A. f x i)) \<up> (\<Sum> x \<in> A. y x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1753
using assms
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1754
proof (induct A rule:finite_ne_induct)
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1755
  case singleton thus ?case by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1756
next
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1757
  case (insert a A) note asms = this
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1758
  hence *: "(\<lambda> i. \<Sum> x \<in> A. f x i) \<up> (\<Sum> x \<in> A. y x)" by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1759
  have **: "(\<lambda> i. f a i) \<up> y a" using asms by simp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1760
  have "(\<lambda> i. f a i + (\<Sum> x \<in> A. f x i)) \<up> (y a + (\<Sum> x \<in> A. y x))"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1761
    using * ** isoton_add by auto
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1762
  thus "(\<lambda> i. \<Sum> x \<in> insert a A. f x i) \<up> (\<Sum> x \<in> insert a A. y x)"
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1763
    using asms by fastsimp
de0b30e6c2d2 Support product spaces on sigma finite measures.
hoelzl
parents: 39302
diff changeset
  1764
qed
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1765
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1766
lemma isoton_Sup:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1767
  assumes "f \<up> u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1768
  shows "f i \<le> u"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1769
  using le_SUPI[of i UNIV f] assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1770
  unfolding isoton_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1771
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1772
lemma isoton_mono:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1773
  assumes iso: "x \<up> a" "y \<up> b" and *: "\<And>n. x n \<le> y (N n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1774
  shows "a \<le> b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1775
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1776
  from iso have "a = (SUP n. x n)" "b = (SUP n. y n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1777
    unfolding isoton_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1778
  with * show ?thesis by (auto intro!: SUP_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1779
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1780
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1781
lemma pextreal_le_mult_one_interval:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1782
  fixes x y :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1783
  assumes "\<And>z. \<lbrakk> 0 < z ; z < 1 \<rbrakk> \<Longrightarrow> z * x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1784
  shows "x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1785
proof (cases x, cases y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1786
  assume "x = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1787
  with assms[of "1 / 2"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1788
  show "x \<le> y" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1789
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1790
  fix r p assume *: "y = Real p" "x = Real r" and **: "0 \<le> r" "0 \<le> p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1791
  have "r \<le> p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1792
  proof (rule field_le_mult_one_interval)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1793
    fix z :: real assume "0 < z" and "z < 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1794
    with assms[of "Real z"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1795
    show "z * r \<le> p" using ** * by (auto simp: zero_le_mult_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1796
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1797
  thus "x \<le> y" using ** * by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1798
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1799
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1800
lemma pextreal_greater_0[intro]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1801
  fixes a :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1802
  assumes "a \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1803
  shows "a > 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1804
using assms apply (cases a) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1805
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1806
lemma pextreal_mult_strict_right_mono:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1807
  assumes "a < b" and "0 < c" "c < \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1808
  shows "a * c < b * c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1809
  using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1810
  by (cases a, cases b, cases c)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1811
     (auto simp: zero_le_mult_iff pextreal_less_\<omega>)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1812
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1813
lemma minus_pextreal_eq2:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1814
  fixes x y z :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1815
  assumes "y \<le> x" and "y \<noteq> \<omega>" shows "z = x - y \<longleftrightarrow> z + y = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1816
  using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1817
  apply (subst eq_commute)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1818
  apply (subst minus_pextreal_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1819
  by (cases x, cases z, auto simp add: ac_simps not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1820
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1821
lemma pextreal_diff_eq_diff_imp_eq:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1822
  assumes "a \<noteq> \<omega>" "b \<le> a" "c \<le> a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1823
  assumes "a - b = a - c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1824
  shows "b = c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1825
  using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1826
  by (cases a, cases b, cases c) (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1827
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1828
lemma pextreal_inverse_eq_0: "inverse x = 0 \<longleftrightarrow> x = \<omega>"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1829
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1830
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1831
lemma pextreal_mult_inverse:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1832
  "\<lbrakk> x \<noteq> \<omega> ; x \<noteq> 0 \<rbrakk> \<Longrightarrow> x * inverse x = 1"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1833
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1834
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1835
lemma pextreal_zero_less_diff_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1836
  fixes a b :: pextreal shows "0 < a - b \<longleftrightarrow> b < a"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1837
  apply (cases a, cases b)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1838
  apply (auto simp: pextreal_noteq_omega_Ex pextreal_less_\<omega>)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1839
  apply (cases b)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1840
  by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1841
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1842
lemma pextreal_less_Real_Ex:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1843
  fixes a b :: pextreal shows "x < Real r \<longleftrightarrow> (\<exists>p\<ge>0. p < r \<and> x = Real p)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1844
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1845
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1846
lemma open_Real: assumes "open S" shows "open (Real ` ({0..} \<inter> S))"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1847
  unfolding open_pextreal_def apply(rule,rule,rule,rule assms) by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1848
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1849
lemma pextreal_zero_le_diff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1850
  fixes a b :: pextreal shows "a - b = 0 \<longleftrightarrow> a \<le> b"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1851
  by (cases a, cases b, simp_all, cases b, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1852
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1853
lemma lim_Real[simp]: assumes "\<forall>n. f n \<ge> 0" "m\<ge>0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1854
  shows "(\<lambda>n. Real (f n)) ----> Real m \<longleftrightarrow> (\<lambda>n. f n) ----> m" (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1855
proof assume ?l show ?r unfolding Lim_sequentially
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1856
  proof safe fix e::real assume e:"e>0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1857
    note open_ball[of m e] note open_Real[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1858
    note * = `?l`[unfolded tendsto_def,rule_format,OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1859
    have "eventually (\<lambda>x. Real (f x) \<in> Real ` ({0..} \<inter> ball m e)) sequentially"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1860
      apply(rule *) unfolding image_iff using assms(2) e by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1861
    thus "\<exists>N. \<forall>n\<ge>N. dist (f n) m < e" unfolding eventually_sequentially 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1862
      apply safe apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1863
    proof- fix n x assume "Real (f n) = Real x" "0 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1864
      hence *:"f n = x" using assms(1) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1865
      assume "x \<in> ball m e" thus "dist (f n) m < e" unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1866
        by (auto simp add:dist_commute)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1867
    qed qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1868
next assume ?r show ?l unfolding tendsto_def eventually_sequentially 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1869
  proof safe fix S assume S:"open S" "Real m \<in> S"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1870
    guess T y using S(1) apply-apply(erule pextreal_openE) . note T=this
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1871
    have "m\<in>real ` (S - {\<omega>})" unfolding image_iff 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1872
      apply(rule_tac x="Real m" in bexI) using assms(2) S(2) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1873
    hence "m \<in> T" unfolding T(2)[THEN sym] by auto 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1874
    from `?r`[unfolded tendsto_def eventually_sequentially,rule_format,OF T(1) this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1875
    guess N .. note N=this[rule_format]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1876
    show "\<exists>N. \<forall>n\<ge>N. Real (f n) \<in> S" apply(rule_tac x=N in exI) 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1877
    proof safe fix n assume n:"N\<le>n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1878
      have "f n \<in> real ` (S - {\<omega>})" using N[OF n] assms unfolding T(2)[THEN sym] 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1879
        unfolding image_iff apply-apply(rule_tac x="Real (f n)" in bexI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1880
        unfolding real_Real by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1881
      then guess x unfolding image_iff .. note x=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1882
      show "Real (f n) \<in> S" unfolding x apply(subst Real_real) using x by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1883
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1884
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1885
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1886
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1887
lemma pextreal_INFI:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1888
  fixes x :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1889
  assumes "\<And>i. i \<in> A \<Longrightarrow> x \<le> f i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1890
  assumes "\<And>y. (\<And>i. i \<in> A \<Longrightarrow> y \<le> f i) \<Longrightarrow> y \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1891
  shows "(INF i:A. f i) = x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1892
  unfolding INFI_def Inf_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1893
  using assms by (auto intro!: Greatest_equality)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1894
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1895
lemma real_of_pextreal_less:"x < y \<Longrightarrow> y\<noteq>\<omega> \<Longrightarrow> real x < real y"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1896
proof- case goal1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1897
  have *:"y = Real (real y)" "x = Real (real x)" using goal1 Real_real by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1898
  show ?case using goal1 apply- apply(subst(asm) *(1))apply(subst(asm) *(2))
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1899
    unfolding pextreal_less by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1900
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1901
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1902
lemma not_less_omega[simp]:"\<not> x < \<omega> \<longleftrightarrow> x = \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1903
  by (metis antisym_conv3 pextreal_less(3)) 
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1904
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1905
lemma Real_real': assumes "x\<noteq>\<omega>" shows "Real (real x) = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1906
proof- have *:"(THE r. 0 \<le> r \<and> x = Real r) = real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1907
    apply(rule the_equality) using assms unfolding Real_real by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1908
  have "Real (THE r. 0 \<le> r \<and> x = Real r) = x" unfolding *
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1909
    using assms unfolding Real_real by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1910
  thus ?thesis unfolding real_of_pextreal_def of_pextreal_def
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1911
    unfolding pextreal_case_def using assms by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1912
qed 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1913
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1914
lemma Real_less_plus_one:"Real x < Real (max (x + 1) 1)" 
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1915
  unfolding pextreal_less by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1916
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1917
lemma Lim_omega: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1918
proof assume ?r show ?l apply(rule topological_tendstoI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1919
    unfolding eventually_sequentially
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1920
  proof- fix S assume "open S" "\<omega> \<in> S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1921
    from open_omega[OF this] guess B .. note B=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1922
    from `?r`[rule_format,of "(max B 0)+1"] guess N .. note N=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1923
    show "\<exists>N. \<forall>n\<ge>N. f n \<in> S" apply(rule_tac x=N in exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1924
    proof safe case goal1 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1925
      have "Real B < Real ((max B 0) + 1)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1926
      also have "... \<le> f n" using goal1 N by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1927
      finally show ?case using B by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1928
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1929
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1930
next assume ?l show ?r
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1931
  proof fix B::real have "open {Real B<..}" "\<omega> \<in> {Real B<..}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1932
    from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1933
    guess N .. note N=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1934
    show "\<exists>N. \<forall>n\<ge>N. Real B \<le> f n" apply(rule_tac x=N in exI) using N by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1935
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1936
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1937
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1938
lemma Lim_bounded_omgea: assumes lim:"f ----> l" and "\<And>n. f n \<le> Real B" shows "l \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1939
proof(rule ccontr,unfold not_not) let ?B = "max (B + 1) 1" assume as:"l=\<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1940
  from lim[unfolded this Lim_omega,rule_format,of "?B"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1941
  guess N .. note N=this[rule_format,OF le_refl]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1942
  hence "Real ?B \<le> Real B" using assms(2)[of N] by(rule order_trans) 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1943
  hence "Real ?B < Real ?B" using Real_less_plus_one[of B] by(rule le_less_trans)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1944
  thus False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1945
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1946
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1947
lemma incseq_le_pextreal: assumes inc: "\<And>n m. n\<ge>m \<Longrightarrow> X n \<ge> X m"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1948
  and lim: "X ----> (L::pextreal)" shows "X n \<le> L"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1949
proof(cases "L = \<omega>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1950
  case False have "\<forall>n. X n \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1951
  proof(rule ccontr,unfold not_all not_not,safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1952
    case goal1 hence "\<forall>n\<ge>x. X n = \<omega>" using inc[of x] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1953
    hence "X ----> \<omega>" unfolding tendsto_def eventually_sequentially
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1954
      apply safe apply(rule_tac x=x in exI) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1955
    note Lim_unique[OF trivial_limit_sequentially this lim]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1956
    with False show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1957
  qed note * =this[rule_format]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1958
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1959
  have **:"\<forall>m n. m \<le> n \<longrightarrow> Real (real (X m)) \<le> Real (real (X n))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1960
    unfolding Real_real using * inc by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1961
  have "real (X n) \<le> real L" apply-apply(rule incseq_le) defer
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1962
    apply(subst lim_Real[THEN sym]) apply(rule,rule,rule)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1963
    unfolding Real_real'[OF *] Real_real'[OF False] 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1964
    unfolding incseq_def using ** lim by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1965
  hence "Real (real (X n)) \<le> Real (real L)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1966
  thus ?thesis unfolding Real_real using * False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1967
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1968
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1969
lemma SUP_Lim_pextreal: assumes "\<And>n m. n\<ge>m \<Longrightarrow> f n \<ge> f m" "f ----> l"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1970
  shows "(SUP n. f n) = (l::pextreal)" unfolding SUPR_def Sup_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1971
proof (safe intro!: Least_equality)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1972
  fix n::nat show "f n \<le> l" apply(rule incseq_le_pextreal)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1973
    using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1974
next fix y assume y:"\<forall>x\<in>range f. x \<le> y" show "l \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1975
  proof(rule ccontr,cases "y=\<omega>",unfold not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1976
    case False assume as:"y < l"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1977
    have l:"l \<noteq> \<omega>" apply(rule Lim_bounded_omgea[OF assms(2), of "real y"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1978
      using False y unfolding Real_real by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1979
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1980
    have yl:"real y < real l" using as apply-
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1981
      apply(subst(asm) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1982
      apply(subst(asm) Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) 
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  1983
      unfolding pextreal_less apply(subst(asm) if_P) by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1984
    hence "y + (y - l) * Real (1 / 2) < l" apply-
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1985
      apply(subst Real_real'[THEN sym,OF `y\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `y\<noteq>\<omega>`])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1986
      apply(subst Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) apply(subst(2) Real_real'[THEN sym,OF `l\<noteq>\<omega>`]) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1987
    hence *:"l \<in> {y + (y - l) / 2<..}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1988
    have "open {y + (y-l)/2 <..}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1989
    note topological_tendstoD[OF assms(2) this *]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1990
    from this[unfolded eventually_sequentially] guess N .. note this[rule_format, of N]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1991
    hence "y + (y - l) * Real (1 / 2) < y" using y[rule_format,of "f N"] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1992
    hence "Real (real y) + (Real (real y) - Real (real l)) * Real (1 / 2) < Real (real y)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1993
      unfolding Real_real using `y\<noteq>\<omega>` `l\<noteq>\<omega>` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1994
    thus False using yl by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1995
  qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1996
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1997
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1998
lemma Real_max':"Real x = Real (max x 0)" 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  1999
proof(cases "x < 0") case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2000
  hence *:"max x 0 = 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2001
  show ?thesis unfolding * using True by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2002
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2003
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2004
lemma lim_pextreal_increasing: assumes "\<forall>n m. n\<ge>m \<longrightarrow> f n \<ge> f m"
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2005
  obtains l where "f ----> (l::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2006
proof(cases "\<exists>B. \<forall>n. f n < Real B")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2007
  case False thus thesis apply- apply(rule that[of \<omega>]) unfolding Lim_omega not_ex not_all
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2008
    apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2009
    apply(rule order_trans[OF _ assms[rule_format]]) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2010
next case True then guess B .. note B = this[rule_format]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2011
  hence *:"\<And>n. f n < \<omega>" apply-apply(rule less_le_trans,assumption) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2012
  have *:"\<And>n. f n \<noteq> \<omega>" proof- case goal1 show ?case using *[of n] by auto qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2013
  have B':"\<And>n. real (f n) \<le> max 0 B" proof- case goal1 thus ?case
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2014
      using B[of n] apply-apply(subst(asm) Real_real'[THEN sym]) defer
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2015
      apply(subst(asm)(2) Real_max') unfolding pextreal_less apply(subst(asm) if_P) using *[of n] by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2016
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2017
  have "\<exists>l. (\<lambda>n. real (f n)) ----> l" apply(rule Topology_Euclidean_Space.bounded_increasing_convergent)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2018
  proof safe show "bounded {real (f n) |n. True}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2019
      unfolding bounded_def apply(rule_tac x=0 in exI,rule_tac x="max 0 B" in exI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2020
      using B' unfolding dist_norm by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2021
    fix n::nat have "Real (real (f n)) \<le> Real (real (f (Suc n)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2022
      using assms[rule_format,of n "Suc n"] apply(subst Real_real)+
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2023
      using *[of n] *[of "Suc n"] by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2024
    thus "real (f n) \<le> real (f (Suc n))" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2025
  qed then guess l .. note l=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2026
  have "0 \<le> l" apply(rule LIMSEQ_le_const[OF l])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2027
    by(rule_tac x=0 in exI,auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2028
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2029
  thus ?thesis apply-apply(rule that[of "Real l"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2030
    using l apply-apply(subst(asm) lim_Real[THEN sym]) prefer 3
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2031
    unfolding Real_real using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2032
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2033
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2034
lemma setsum_neq_omega: assumes "finite s" "\<And>x. x \<in> s \<Longrightarrow> f x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2035
  shows "setsum f s \<noteq> \<omega>" using assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2036
proof induct case (insert x s)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2037
  show ?case unfolding setsum.insert[OF insert(1-2)] 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2038
    using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2039
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2040
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2041
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2042
lemma real_Real': "0 \<le> x \<Longrightarrow> real (Real x) = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2043
  unfolding real_Real by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2044
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2045
lemma real_pextreal_pos[intro]:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2046
  assumes "x \<noteq> 0" "x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2047
  shows "real x > 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2048
  apply(subst real_Real'[THEN sym,of 0]) defer
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2049
  apply(rule real_of_pextreal_less) using assms by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2050
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2051
lemma Lim_omega_gt: "f ----> \<omega> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n > Real B)" (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2052
proof assume ?l thus ?r unfolding Lim_omega apply safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2053
    apply(erule_tac x="max B 0 +1" in allE,safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2054
    apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2055
    apply(rule_tac y="Real (max B 0 + 1)" in less_le_trans) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2056
next assume ?r thus ?l unfolding Lim_omega apply safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2057
    apply(erule_tac x=B in allE,safe) apply(rule_tac x=N in exI,safe) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2058
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2059
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2060
lemma pextreal_minus_le_cancel:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2061
  fixes a b c :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2062
  assumes "b \<le> a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2063
  shows "c - a \<le> c - b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2064
  using assms by (cases a, cases b, cases c, simp, simp, simp, cases b, cases c, simp_all)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2065
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2066
lemma pextreal_minus_\<omega>[simp]: "x - \<omega> = 0" by (cases x) simp_all
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2067
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2068
lemma pextreal_minus_mono[intro]: "a - x \<le> (a::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2069
proof- have "a - x \<le> a - 0"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2070
    apply(rule pextreal_minus_le_cancel) by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2071
  thus ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2072
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2073
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2074
lemma pextreal_minus_eq_\<omega>[simp]: "x - y = \<omega> \<longleftrightarrow> (x = \<omega> \<and> y \<noteq> \<omega>)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2075
  by (cases x, cases y) (auto, cases y, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2076
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2077
lemma pextreal_less_minus_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2078
  fixes a b c :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2079
  shows "a < b - c \<longleftrightarrow> c + a < b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2080
  by (cases c, cases a, cases b, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2081
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2082
lemma pextreal_minus_less_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2083
  fixes a b c :: pextreal shows "a - c < b \<longleftrightarrow> (0 < b \<and> (c \<noteq> \<omega> \<longrightarrow> a < b + c))"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2084
  by (cases c, cases a, cases b, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2085
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2086
lemma pextreal_le_minus_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2087
  fixes a b c :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2088
  shows "a \<le> c - b \<longleftrightarrow> ((c \<le> b \<longrightarrow> a = 0) \<and> (b < c \<longrightarrow> a + b \<le> c))"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2089
  by (cases a, cases c, cases b, auto simp: pextreal_noteq_omega_Ex)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2090
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2091
lemma pextreal_minus_le_iff:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2092
  fixes a b c :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2093
  shows "a - c \<le> b \<longleftrightarrow> (c \<le> a \<longrightarrow> a \<le> b + c)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2094
  by (cases a, cases c, cases b, auto simp: pextreal_noteq_omega_Ex)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2095
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2096
lemmas pextreal_minus_order = pextreal_minus_le_iff pextreal_minus_less_iff pextreal_le_minus_iff pextreal_less_minus_iff
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2097
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2098
lemma pextreal_minus_strict_mono:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2099
  assumes "a > 0" "x > 0" "a\<noteq>\<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2100
  shows "a - x < (a::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2101
  using assms by(cases x, cases a, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2102
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2103
lemma pextreal_minus':
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2104
  "Real r - Real p = (if 0 \<le> r \<and> p \<le> r then if 0 \<le> p then Real (r - p) else Real r else 0)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2105
  by (auto simp: minus_pextreal_eq not_less)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2106
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2107
lemma pextreal_minus_plus:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2108
  "x \<le> (a::pextreal) \<Longrightarrow> a - x + x = a"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2109
  by (cases a, cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2110
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2111
lemma pextreal_cancel_plus_minus: "b \<noteq> \<omega> \<Longrightarrow> a + b - b = a"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2112
  by (cases a, cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2113
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2114
lemma pextreal_minus_le_cancel_right:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2115
  fixes a b c :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2116
  assumes "a \<le> b" "c \<le> a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2117
  shows "a - c \<le> b - c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2118
  using assms by (cases a, cases b, cases c, auto, cases c, auto)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2119
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2120
lemma real_of_pextreal_setsum':
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2121
  assumes "\<forall>x \<in> S. f x \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2122
  shows "(\<Sum>x\<in>S. real (f x)) = real (setsum f S)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2123
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2124
  assume "finite S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2125
  from this assms show ?thesis
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2126
    by induct (simp_all add: real_of_pextreal_add setsum_\<omega>)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2127
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2128
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2129
lemma Lim_omega_pos: "f ----> \<omega> \<longleftrightarrow> (\<forall>B>0. \<exists>N. \<forall>n\<ge>N. f n \<ge> Real B)" (is "?l = ?r")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2130
  unfolding Lim_omega apply safe defer
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2131
  apply(erule_tac x="max 1 B" in allE) apply safe defer
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2132
  apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2133
  apply(rule_tac y="Real (max 1 B)" in order_trans) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2134
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2135
lemma pextreal_LimI_finite:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2136
  assumes "x \<noteq> \<omega>" "\<And>r. 0 < r \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. u n < x + r \<and> x < u n + r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2137
  shows "u ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2138
proof (rule topological_tendstoI, unfold eventually_sequentially)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2139
  fix S assume "open S" "x \<in> S"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2140
  then obtain A where "open A" and A_eq: "Real ` (A \<inter> {0..}) = S - {\<omega>}" by (auto elim!: pextreal_openE)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2141
  then have "x \<in> Real ` (A \<inter> {0..})" using `x \<in> S` `x \<noteq> \<omega>` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2142
  then have "real x \<in> A" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2143
  then obtain r where "0 < r" and dist: "\<And>y. dist y (real x) < r \<Longrightarrow> y \<in> A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2144
    using `open A` unfolding open_real_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2145
  then obtain n where
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2146
    upper: "\<And>N. n \<le> N \<Longrightarrow> u N < x + Real r" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2147
    lower: "\<And>N. n \<le> N \<Longrightarrow> x < u N + Real r" using assms(2)[of "Real r"] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2148
  show "\<exists>N. \<forall>n\<ge>N. u n \<in> S"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2149
  proof (safe intro!: exI[of _ n])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2150
    fix N assume "n \<le> N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2151
    from upper[OF this] `x \<noteq> \<omega>` `0 < r`
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2152
    have "u N \<noteq> \<omega>" by (force simp: pextreal_noteq_omega_Ex)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2153
    with `x \<noteq> \<omega>` `0 < r` lower[OF `n \<le> N`] upper[OF `n \<le> N`]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2154
    have "dist (real (u N)) (real x) < r" "u N \<noteq> \<omega>"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2155
      by (auto simp: pextreal_noteq_omega_Ex dist_real_def abs_diff_less_iff field_simps)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2156
    from dist[OF this(1)]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2157
    have "u N \<in> Real ` (A \<inter> {0..})" using `u N \<noteq> \<omega>`
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2158
      by (auto intro!: image_eqI[of _ _ "real (u N)"] simp: pextreal_noteq_omega_Ex Real_real)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2159
    thus "u N \<in> S" using A_eq by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2160
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2161
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2162
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2163
lemma real_Real_max:"real (Real x) = max x 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2164
  unfolding real_Real by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2165
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2166
lemma Sup_lim:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2167
  assumes "\<forall>n. b n \<in> s" "b ----> (a::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2168
  shows "a \<le> Sup s"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2169
proof(rule ccontr,unfold not_le)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2170
  assume as:"Sup s < a" hence om:"Sup s \<noteq> \<omega>" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2171
  have s:"s \<noteq> {}" using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2172
  { presume *:"\<forall>n. b n < a \<Longrightarrow> False"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2173
    show False apply(cases,rule *,assumption,unfold not_all not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2174
    proof- case goal1 then guess n .. note n=this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2175
      thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2176
        using as by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2177
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2178
  } assume b:"\<forall>n. b n < a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2179
  show False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2180
  proof(cases "a = \<omega>")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2181
    case False have *:"a - Sup s > 0" 
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2182
      using False as by(auto simp: pextreal_zero_le_diff)
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2183
    have "(a - Sup s) / 2 \<le> a / 2" unfolding divide_pextreal_def
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2184
      apply(rule mult_right_mono) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2185
    also have "... = Real (real (a / 2))" apply(rule Real_real'[THEN sym])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2186
      using False by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2187
    also have "... < Real (real a)" unfolding pextreal_less using as False
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2188
      by(auto simp add: real_of_pextreal_mult[THEN sym])
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2189
    also have "... = a" apply(rule Real_real') using False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2190
    finally have asup:"a > (a - Sup s) / 2" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2191
    have "\<exists>n. a - b n < (a - Sup s) / 2"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2192
    proof(rule ccontr,unfold not_ex not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2193
      case goal1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2194
      have "(a - Sup s) * Real (1 / 2)  > 0" 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2195
        using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2196
      hence "a - (a - Sup s) * Real (1 / 2) < a"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2197
        apply-apply(rule pextreal_minus_strict_mono)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2198
        using False * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2199
      hence *:"a \<in> {a - (a - Sup s) / 2<..}"using asup by auto 
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2200
      note topological_tendstoD[OF assms(2) open_pextreal_greaterThan,OF *]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2201
      from this[unfolded eventually_sequentially] guess n .. 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2202
      note n = this[rule_format,of n] 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2203
      have "b n + (a - Sup s) / 2 \<le> a" 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2204
        using add_right_mono[OF goal1[rule_format,of n],of "b n"]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2205
        unfolding pextreal_minus_plus[OF less_imp_le[OF b[rule_format]]]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2206
        by(auto simp: add_commute)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2207
      hence "b n \<le> a - (a - Sup s) / 2" unfolding pextreal_le_minus_iff
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2208
        using asup by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2209
      hence "b n \<notin> {a - (a - Sup s) / 2<..}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2210
      thus False using n by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2211
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2212
    then guess n .. note n = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2213
    have "Sup s < a - (a - Sup s) / 2"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2214
      using False as om by (cases a) (auto simp: pextreal_noteq_omega_Ex field_simps)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2215
    also have "... \<le> b n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2216
    proof- note add_right_mono[OF less_imp_le[OF n],of "b n"]
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2217
      note this[unfolded pextreal_minus_plus[OF less_imp_le[OF b[rule_format]]]]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2218
      hence "a - (a - Sup s) / 2 \<le> (a - Sup s) / 2 + b n - (a - Sup s) / 2"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2219
        apply(rule pextreal_minus_le_cancel_right) using asup by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2220
      also have "... = b n + (a - Sup s) / 2 - (a - Sup s) / 2" 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2221
        by(auto simp add: add_commute)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2222
      also have "... = b n" apply(subst pextreal_cancel_plus_minus)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2223
      proof(rule ccontr,unfold not_not) case goal1
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2224
        show ?case using asup unfolding goal1 by auto 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2225
      qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2226
      finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2227
    qed     
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2228
    finally show False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2229
      using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]] by auto  
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2230
  next case True
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2231
    from assms(2)[unfolded True Lim_omega_gt,rule_format,of "real (Sup s)"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2232
    guess N .. note N = this[rule_format,of N]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2233
    thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of N]] 
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2234
      unfolding Real_real using om by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2235
  qed qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2236
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2237
lemma Sup_mono_lim:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2238
  assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> b ----> (a::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2239
  shows "Sup A \<le> Sup B"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2240
  unfolding Sup_le_iff apply(rule) apply(drule assms[rule_format]) apply safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2241
  apply(rule_tac b=b in Sup_lim) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2242
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2243
lemma pextreal_less_add:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2244
  assumes "x \<noteq> \<omega>" "a < b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2245
  shows "x + a < x + b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2246
  using assms by (cases a, cases b, cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2247
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2248
lemma SUPR_lim:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2249
  assumes "\<forall>n. b n \<in> B" "(\<lambda>n. f (b n)) ----> (f a::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2250
  shows "f a \<le> SUPR B f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2251
  unfolding SUPR_def apply(rule Sup_lim[of "\<lambda>n. f (b n)"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2252
  using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2253
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2254
lemma SUP_\<omega>_imp:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2255
  assumes "(SUP i. f i) = \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2256
  shows "\<exists>i. Real x < f i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2257
proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2258
  assume "\<not> ?thesis" hence "\<And>i. f i \<le> Real x" by (simp add: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2259
  hence "(SUP i. f i) \<le> Real x" unfolding SUP_le_iff by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2260
  with assms show False by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2261
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2262
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2263
lemma SUPR_mono_lim:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2264
  assumes "\<forall>a\<in>A. \<exists>b. \<forall>n. b n \<in> B \<and> (\<lambda>n. f (b n)) ----> (f a::pextreal)"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2265
  shows "SUPR A f \<le> SUPR B f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2266
  unfolding SUPR_def apply(rule Sup_mono_lim)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2267
  apply safe apply(drule assms[rule_format],safe)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2268
  apply(rule_tac x="\<lambda>n. f (b n)" in exI) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2269
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2270
lemma real_0_imp_eq_0:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2271
  assumes "x \<noteq> \<omega>" "real x = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2272
  shows "x = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2273
using assms by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2274
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2275
lemma SUPR_mono:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2276
  assumes "\<forall>a\<in>A. \<exists>b\<in>B. f b \<ge> f a"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2277
  shows "SUPR A f \<le> SUPR B f"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2278
  unfolding SUPR_def apply(rule Sup_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2279
  using assms by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2280
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2281
lemma less_add_Real:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2282
  fixes x :: real
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2283
  fixes a b :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2284
  assumes "x \<ge> 0" "a < b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2285
  shows "a + Real x < b + Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2286
using assms by (cases a, cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2287
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2288
lemma le_add_Real:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2289
  fixes x :: real
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2290
  fixes a b :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2291
  assumes "x \<ge> 0" "a \<le> b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2292
  shows "a + Real x \<le> b + Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2293
using assms by (cases a, cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2294
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2295
lemma le_imp_less_pextreal:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2296
  fixes x :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2297
  assumes "x > 0" "a + x \<le> b" "a \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2298
  shows "a < b"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2299
using assms by (cases x, cases a, cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2300
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2301
lemma pextreal_INF_minus:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2302
  fixes f :: "nat \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2303
  assumes "c \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2304
  shows "(INF i. c - f i) = c - (SUP i. f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2305
proof (cases "SUP i. f i")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2306
  case infinite
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2307
  from `c \<noteq> \<omega>` obtain x where [simp]: "c = Real x" by (cases c) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2308
  from SUP_\<omega>_imp[OF infinite] obtain i where "Real x < f i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2309
  have "(INF i. c - f i) \<le> c - f i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2310
    by (auto intro!: complete_lattice_class.INF_leI)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2311
  also have "\<dots> = 0" using `Real x < f i` by (auto simp: minus_pextreal_eq)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2312
  finally show ?thesis using infinite by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2313
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2314
  case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2315
  from `c \<noteq> \<omega>` obtain x where c: "c = Real x" by (cases c) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2316
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2317
  show ?thesis unfolding c
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2318
  proof (rule pextreal_INFI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2319
    fix i have "f i \<le> (SUP i. f i)" by (rule le_SUPI) simp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2320
    thus "Real x - (SUP i. f i) \<le> Real x - f i" by (rule pextreal_minus_le_cancel)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2321
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2322
    fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> y \<le> Real x - f i"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2323
    from this[of 0] obtain p where p: "y = Real p" "0 \<le> p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2324
      by (cases "f 0", cases y, auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2325
    hence "\<And>i. Real p \<le> Real x - f i" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2326
    hence *: "\<And>i. Real x \<le> f i \<Longrightarrow> Real p = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2327
      "\<And>i. f i < Real x \<Longrightarrow> Real p + f i \<le> Real x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2328
      unfolding pextreal_le_minus_iff by auto
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2329
    show "y \<le> Real x - (SUP i. f i)" unfolding p pextreal_le_minus_iff
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2330
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2331
      assume x_less: "Real x \<le> (SUP i. f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2332
      show "Real p = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2333
      proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2334
        assume "Real p \<noteq> 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2335
        hence "0 < Real p" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2336
        from Sup_close[OF this, of "range f"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2337
        obtain i where e: "(SUP i. f i) < f i + Real p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2338
          using preal unfolding SUPR_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2339
        hence "Real x \<le> f i + Real p" using x_less by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2340
        show False
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2341
        proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2342
          assume "\<forall>i. f i < Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2343
          hence "Real p + f i \<le> Real x" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2344
          hence "f i + Real p \<le> (SUP i. f i)" using x_less by (auto simp: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2345
          thus False using e by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2346
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2347
          assume "\<not> (\<forall>i. f i < Real x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2348
          then obtain i where "Real x \<le> f i" by (auto simp: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2349
          from *(1)[OF this] show False using `Real p \<noteq> 0` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2350
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2351
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2352
    next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2353
      have "\<And>i. f i \<le> (SUP i. f i)" by (rule complete_lattice_class.le_SUPI) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2354
      also assume "(SUP i. f i) < Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2355
      finally have "\<And>i. f i < Real x" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2356
      hence *: "\<And>i. Real p + f i \<le> Real x" using * by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2357
      have "Real p \<le> Real x" using *[of 0] by (cases "f 0") (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2358
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2359
      have SUP_eq: "(SUP i. f i) \<le> Real x - Real p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2360
      proof (rule SUP_leI)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2361
        fix i show "f i \<le> Real x - Real p" unfolding pextreal_le_minus_iff
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2362
        proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2363
          assume "Real x \<le> Real p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2364
          with *[of i] show "f i = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2365
            by (cases "f i") (auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2366
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2367
          assume "Real p < Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2368
          show "f i + Real p \<le> Real x" using * by (auto simp: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2369
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2370
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2371
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2372
      show "Real p + (SUP i. f i) \<le> Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2373
      proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2374
        assume "Real x \<le> Real p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2375
        with `Real p \<le> Real x` have [simp]: "Real p = Real x" by (rule antisym)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2376
        { fix i have "f i = 0" using *[of i] by (cases "f i") (auto split: split_if_asm) }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2377
        hence "(SUP i. f i) \<le> 0" by (auto intro!: SUP_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2378
        thus ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2379
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2380
        assume "\<not> Real x \<le> Real p" hence "Real p < Real x" unfolding not_le .
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2381
        with SUP_eq show ?thesis unfolding pextreal_le_minus_iff by (auto simp: field_simps)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2382
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2383
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2384
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2385
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2386
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2387
lemma pextreal_SUP_minus:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2388
  fixes f :: "nat \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2389
  shows "(SUP i. c - f i) = c - (INF i. f i)"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2390
proof (rule pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2391
  fix i have "(INF i. f i) \<le> f i" by (rule INF_leI) simp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2392
  thus "c - f i \<le> c - (INF i. f i)" by (rule pextreal_minus_le_cancel)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2393
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2394
  fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> c - f i \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2395
  show "c - (INF i. f i) \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2396
  proof (cases y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2397
    case (preal p)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2398
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2399
    show ?thesis unfolding pextreal_minus_le_iff preal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2400
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2401
      assume INF_le_x: "(INF i. f i) \<le> c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2402
      from * have *: "\<And>i. f i \<le> c \<Longrightarrow> c \<le> Real p + f i"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2403
        unfolding pextreal_minus_le_iff preal by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2404
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2405
      have INF_eq: "c - Real p \<le> (INF i. f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2406
      proof (rule le_INFI)
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2407
        fix i show "c - Real p \<le> f i" unfolding pextreal_minus_le_iff
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2408
        proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2409
          assume "Real p \<le> c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2410
          show "c \<le> f i + Real p"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2411
          proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2412
            assume "f i \<le> c" from *[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2413
            show ?thesis by (simp add: field_simps)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2414
          next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2415
            assume "\<not> f i \<le> c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2416
            hence "c \<le> f i" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2417
            also have "\<dots> \<le> f i + Real p" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2418
            finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2419
          qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2420
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2421
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2422
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2423
      show "c \<le> Real p + (INF i. f i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2424
      proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2425
        assume "Real p \<le> c"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2426
        with INF_eq show ?thesis unfolding pextreal_minus_le_iff by (auto simp: field_simps)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2427
      next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2428
        assume "\<not> Real p \<le> c"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2429
        hence "c \<le> Real p" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2430
        also have "Real p \<le> Real p + (INF i. f i)" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2431
        finally show ?thesis .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2432
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2433
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2434
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2435
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2436
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2437
lemma pextreal_le_minus_imp_0:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2438
  fixes a b :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2439
  shows "a \<le> a - b \<Longrightarrow> a \<noteq> 0 \<Longrightarrow> a \<noteq> \<omega> \<Longrightarrow> b = 0"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2440
  by (cases a, cases b, auto split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2441
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2442
lemma lim_INF_eq_lim_SUP:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2443
  fixes X :: "nat \<Rightarrow> real"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2444
  assumes "\<And>i. 0 \<le> X i" and "0 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2445
  and lim_INF: "(SUP n. INF m. Real (X (n + m))) = Real x" (is "(SUP n. ?INF n) = _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2446
  and lim_SUP: "(INF n. SUP m. Real (X (n + m))) = Real x" (is "(INF n. ?SUP n) = _")
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2447
  shows "X ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2448
proof (rule LIMSEQ_I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2449
  fix r :: real assume "0 < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2450
  hence "0 \<le> r" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2451
  from Sup_close[of "Real r" "range ?INF"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2452
  obtain n where inf: "Real x < ?INF n + Real r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2453
    unfolding SUPR_def lim_INF[unfolded SUPR_def] using `0 < r` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2454
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2455
  from Inf_close[of "range ?SUP" "Real r"]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2456
  obtain n' where sup: "?SUP n' < Real x + Real r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2457
    unfolding INFI_def lim_SUP[unfolded INFI_def] using `0 < r` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2458
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2459
  show "\<exists>N. \<forall>n\<ge>N. norm (X n - x) < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2460
  proof (safe intro!: exI[of _ "max n n'"])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2461
    fix m assume "max n n' \<le> m" hence "n \<le> m" "n' \<le> m" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2462
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2463
    note inf
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2464
    also have "?INF n + Real r \<le> Real (X (n + (m - n))) + Real r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2465
      by (rule le_add_Real, auto simp: `0 \<le> r` intro: INF_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2466
    finally have up: "x < X m + r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2467
      using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n \<le> m` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2468
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2469
    have "Real (X (n' + (m - n'))) \<le> ?SUP n'"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2470
      by (auto simp: `0 \<le> r` intro: le_SUPI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2471
    also note sup
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2472
    finally have down: "X m < x + r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2473
      using `0 \<le> X m` `0 \<le> x` `0 \<le> r` `n' \<le> m` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2474
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2475
    show "norm (X m - x) < r" using up down by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2476
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2477
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2478
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2479
lemma Sup_countable_SUPR:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2480
  assumes "Sup A \<noteq> \<omega>" "A \<noteq> {}"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2481
  shows "\<exists> f::nat \<Rightarrow> pextreal. range f \<subseteq> A \<and> Sup A = SUPR UNIV f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2482
proof -
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2483
  have "\<And>n. 0 < 1 / (of_nat n :: pextreal)" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2484
  from Sup_close[OF this assms]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2485
  have "\<forall>n. \<exists>x. x \<in> A \<and> Sup A < x + 1 / of_nat n" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2486
  from choice[OF this] obtain f where "range f \<subseteq> A" and
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2487
    epsilon: "\<And>n. Sup A < f n + 1 / of_nat n" by blast
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2488
  have "SUPR UNIV f = Sup A"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2489
  proof (rule pextreal_SUPI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2490
    fix i show "f i \<le> Sup A" using `range f \<subseteq> A`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2491
      by (auto intro!: complete_lattice_class.Sup_upper)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2492
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2493
    fix y assume bound: "\<And>i. i \<in> UNIV \<Longrightarrow> f i \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2494
    show "Sup A \<le> y"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2495
    proof (rule pextreal_le_epsilon)
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2496
      fix e :: pextreal assume "0 < e"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2497
      show "Sup A \<le> y + e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2498
      proof (cases e)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2499
        case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2500
        hence "0 < r" using `0 < e` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2501
        then obtain n where *: "inverse (of_nat n) < r" "0 < n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2502
          using ex_inverse_of_nat_less by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2503
        have "Sup A \<le> f n + 1 / of_nat n" using epsilon[of n] by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2504
        also have "1 / of_nat n \<le> e" using preal * by (auto simp: real_eq_of_nat)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2505
        with bound have "f n + 1 / of_nat n \<le> y + e" by (rule add_mono) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2506
        finally show "Sup A \<le> y + e" .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2507
      qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2508
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2509
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2510
  with `range f \<subseteq> A` show ?thesis by (auto intro!: exI[of _ f])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2511
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2512
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2513
lemma SUPR_countable_SUPR:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2514
  assumes "SUPR A g \<noteq> \<omega>" "A \<noteq> {}"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2515
  shows "\<exists> f::nat \<Rightarrow> pextreal. range f \<subseteq> g`A \<and> SUPR A g = SUPR UNIV f"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2516
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2517
  have "Sup (g`A) \<noteq> \<omega>" "g`A \<noteq> {}" using assms unfolding SUPR_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2518
  from Sup_countable_SUPR[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2519
  show ?thesis unfolding SUPR_def .
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2520
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2521
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2522
lemma pextreal_setsum_subtractf:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2523
  assumes "\<And>i. i\<in>A \<Longrightarrow> g i \<le> f i" and "\<And>i. i\<in>A \<Longrightarrow> f i \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2524
  shows "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2525
proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2526
  assume "finite A" from this assms show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2527
  proof induct
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2528
    case (insert x A)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2529
    hence hyp: "(\<Sum>i\<in>A. f i - g i) = (\<Sum>i\<in>A. f i) - (\<Sum>i\<in>A. g i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2530
      by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2531
    { fix i assume *: "i \<in> insert x A"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2532
      hence "g i \<le> f i" using insert by simp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2533
      also have "f i < \<omega>" using * insert by (simp add: pextreal_less_\<omega>)
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2534
      finally have "g i \<noteq> \<omega>" by (simp add: pextreal_less_\<omega>) }
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2535
    hence "setsum g A \<noteq> \<omega>" "g x \<noteq> \<omega>" by (auto simp: setsum_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2536
    moreover have "setsum f A \<noteq> \<omega>" "f x \<noteq> \<omega>" using insert by (auto simp: setsum_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2537
    moreover have "g x \<le> f x" using insert by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2538
    moreover have "(\<Sum>i\<in>A. g i) \<le> (\<Sum>i\<in>A. f i)" using insert by (auto intro!: setsum_mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2539
    ultimately show ?case using `finite A` `x \<notin> A` hyp
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2540
      by (auto simp: pextreal_noteq_omega_Ex)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2541
  qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2542
qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2543
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2544
lemma real_of_pextreal_diff:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2545
  "y \<le> x \<Longrightarrow> x \<noteq> \<omega> \<Longrightarrow> real x - real y = real (x - y)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2546
  by (cases x, cases y) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2547
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2548
lemma psuminf_minus:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2549
  assumes ord: "\<And>i. g i \<le> f i" and fin: "psuminf g \<noteq> \<omega>" "psuminf f \<noteq> \<omega>"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2550
  shows "(\<Sum>\<^isub>\<infinity> i. f i - g i) = psuminf f - psuminf g"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2551
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2552
  have [simp]: "\<And>i. f i \<noteq> \<omega>" using fin by (auto intro: psuminf_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2553
  from fin have "(\<lambda>x. real (f x)) sums real (\<Sum>\<^isub>\<infinity>x. f x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2554
    and "(\<lambda>x. real (g x)) sums real (\<Sum>\<^isub>\<infinity>x. g x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2555
    by (auto intro: psuminf_imp_suminf)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2556
  from sums_diff[OF this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2557
  have "(\<lambda>n. real (f n - g n)) sums (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))" using fin ord
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2558
    by (subst (asm) (1 2) real_of_pextreal_diff) (auto simp: psuminf_\<omega> psuminf_le)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2559
  hence "(\<Sum>\<^isub>\<infinity> i. Real (real (f i - g i))) = Real (real ((\<Sum>\<^isub>\<infinity>x. f x) - (\<Sum>\<^isub>\<infinity>x. g x)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2560
    by (rule suminf_imp_psuminf) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2561
  thus ?thesis using fin by (simp add: Real_real psuminf_\<omega>)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2562
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2563
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2564
lemma INF_eq_LIMSEQ:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2565
  assumes "mono (\<lambda>i. - f i)" and "\<And>n. 0 \<le> f n" and "0 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2566
  shows "(INF n. Real (f n)) = Real x \<longleftrightarrow> f ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2567
proof
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2568
  assume x: "(INF n. Real (f n)) = Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2569
  { fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2570
    have "Real x \<le> Real (f n)" using x[symmetric] by (auto intro: INF_leI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2571
    hence "x \<le> f n" using assms by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2572
    hence "\<bar>f n - x\<bar> = f n - x" by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2573
  note abs_eq = this
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2574
  show "f ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2575
  proof (rule LIMSEQ_I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2576
    fix r :: real assume "0 < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2577
    show "\<exists>no. \<forall>n\<ge>no. norm (f n - x) < r"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2578
    proof (rule ccontr)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2579
      assume *: "\<not> ?thesis"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2580
      { fix N
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2581
        from * obtain n where *: "N \<le> n" "r \<le> f n - x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2582
          using abs_eq by (auto simp: not_less)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2583
        hence "x + r \<le> f n" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2584
        also have "f n \<le> f N" using `mono (\<lambda>i. - f i)` * by (auto dest: monoD)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2585
        finally have "Real (x + r) \<le> Real (f N)" using `0 \<le> f N` by auto }
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2586
      hence "Real x < Real (x + r)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2587
        and "Real (x + r) \<le> (INF n. Real (f n))" using `0 < r` `0 \<le> x` by (auto intro: le_INFI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2588
      hence "Real x < (INF n. Real (f n))" by (rule less_le_trans)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2589
      thus False using x by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2590
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2591
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2592
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2593
  assume "f ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2594
  show "(INF n. Real (f n)) = Real x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2595
  proof (rule pextreal_INFI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2596
    fix n
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2597
    from decseq_le[OF _ `f ----> x`] assms
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2598
    show "Real x \<le> Real (f n)" unfolding decseq_eq_incseq incseq_mono by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2599
  next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2600
    fix y assume *: "\<And>n. n\<in>UNIV \<Longrightarrow> y \<le> Real (f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2601
    thus "y \<le> Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2602
    proof (cases y)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2603
      case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2604
      with * have "\<exists>N. \<forall>n\<ge>N. r \<le> f n" using assms by fastsimp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2605
      from LIMSEQ_le_const[OF `f ----> x` this]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2606
      show "y \<le> Real x" using `0 \<le> x` preal by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2607
    qed simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2608
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2609
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2610
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2611
lemma INFI_bound:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2612
  assumes "\<forall>N. x \<le> f N"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2613
  shows "x \<le> (INF n. f n)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2614
  using assms by (simp add: INFI_def le_Inf_iff)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2615
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2616
lemma LIMSEQ_imp_lim_INF:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2617
  assumes pos: "\<And>i. 0 \<le> X i" and lim: "X ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2618
  shows "(SUP n. INF m. Real (X (n + m))) = Real x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2619
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2620
  have "0 \<le> x" using assms by (auto intro!: LIMSEQ_le_const)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2621
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2622
  have "\<And>n. (INF m. Real (X (n + m))) \<le> Real (X (n + 0))" by (rule INF_leI) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2623
  also have "\<And>n. Real (X (n + 0)) < \<omega>" by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2624
  finally have "\<forall>n. \<exists>r\<ge>0. (INF m. Real (X (n + m))) = Real r"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2625
    by (auto simp: pextreal_less_\<omega> pextreal_noteq_omega_Ex)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2626
  from choice[OF this] obtain r where r: "\<And>n. (INF m. Real (X (n + m))) = Real (r n)" "\<And>n. 0 \<le> r n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2627
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2628
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2629
  show ?thesis unfolding r
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2630
  proof (subst SUP_eq_LIMSEQ)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2631
    show "mono r" unfolding mono_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2632
    proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2633
      fix x y :: nat assume "x \<le> y"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2634
      have "Real (r x) \<le> Real (r y)" unfolding r(1)[symmetric] using pos
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2635
      proof (safe intro!: INF_mono bexI)
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2636
        fix m have "x + (m + y - x) = y + m"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2637
          using `x \<le> y` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2638
        thus "Real (X (x + (m + y - x))) \<le> Real (X (y + m))" by simp
38705
aaee86c0e237 moved generic lemmas in Probability to HOL
hoelzl
parents: 38656
diff changeset
  2639
      qed simp
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2640
      thus "r x \<le> r y" using r by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2641
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2642
    show "\<And>n. 0 \<le> r n" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2643
    show "0 \<le> x" by fact
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2644
    show "r ----> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2645
    proof (rule LIMSEQ_I)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2646
      fix e :: real assume "0 < e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2647
      hence "0 < e/2" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2648
      from LIMSEQ_D[OF lim this] obtain N where *: "\<And>n. N \<le> n \<Longrightarrow> \<bar>X n - x\<bar> < e/2"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2649
        by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2650
      show "\<exists>N. \<forall>n\<ge>N. norm (r n - x) < e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2651
      proof (safe intro!: exI[of _ N])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2652
        fix n assume "N \<le> n"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2653
        show "norm (r n - x) < e"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2654
        proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2655
          assume "r n < x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2656
          have "x - r n \<le> e/2"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2657
          proof cases
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2658
            assume e: "e/2 \<le> x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2659
            have "Real (x - e/2) \<le> Real (r n)" unfolding r(1)[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2660
            proof (rule le_INFI)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2661
              fix m show "Real (x - e / 2) \<le> Real (X (n + m))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2662
                using *[of "n + m"] `N \<le> n`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2663
                using pos by (auto simp: field_simps abs_real_def split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2664
            qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2665
            with e show ?thesis using pos `0 \<le> x` r(2) by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2666
          next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2667
            assume "\<not> e/2 \<le> x" hence "x - e/2 < 0" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2668
            with `0 \<le> r n` show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2669
          qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2670
          with `r n < x` show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2671
        next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2672
          assume e: "\<not> r n < x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2673
          have "Real (r n) \<le> Real (X (n + 0))" unfolding r(1)[symmetric]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2674
            by (rule INF_leI) simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2675
          hence "r n - x \<le> X n - x" using r pos by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2676
          also have "\<dots> < e/2" using *[OF `N \<le> n`] by (auto simp: field_simps abs_real_def split: split_if_asm)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2677
          finally have "r n - x < e" using `0 < e` by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2678
          with e show ?thesis by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2679
        qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2680
      qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2681
    qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2682
  qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2683
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2684
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2685
lemma real_of_pextreal_strict_mono_iff:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2686
  "real a < real b \<longleftrightarrow> (b \<noteq> \<omega> \<and> ((a = \<omega> \<and> 0 < b) \<or> (a < b)))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2687
proof (cases a)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2688
  case infinite thus ?thesis by (cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2689
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2690
  case preal thus ?thesis by (cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2691
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2692
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2693
lemma real_of_pextreal_mono_iff:
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2694
  "real a \<le> real b \<longleftrightarrow> (a = \<omega> \<or> (b \<noteq> \<omega> \<and> a \<le> b) \<or> (b = \<omega> \<and> a = 0))"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2695
proof (cases a)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2696
  case infinite thus ?thesis by (cases b) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2697
next
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2698
  case preal thus ?thesis by (cases b)  auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2699
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2700
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2701
lemma ex_pextreal_inverse_of_nat_Suc_less:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2702
  fixes e :: pextreal assumes "0 < e" shows "\<exists>n. inverse (of_nat (Suc n)) < e"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2703
proof (cases e)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2704
  case (preal r)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2705
  with `0 < e` ex_inverse_of_nat_Suc_less[of r]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2706
  obtain n where "inverse (of_nat (Suc n)) < r" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2707
  with preal show ?thesis
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2708
    by (auto simp: real_eq_of_nat[symmetric])
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2709
qed auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2710
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2711
lemma Lim_eq_Sup_mono:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2712
  fixes u :: "nat \<Rightarrow> pextreal" assumes "mono u"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2713
  shows "u ----> (SUP i. u i)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2714
proof -
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2715
  from lim_pextreal_increasing[of u] `mono u`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2716
  obtain l where l: "u ----> l" unfolding mono_def by auto
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2717
  from SUP_Lim_pextreal[OF _ this] `mono u`
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2718
  have "(SUP i. u i) = l" unfolding mono_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2719
  with l show ?thesis by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2720
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2721
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2722
lemma isotone_Lim:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2723
  fixes x :: pextreal assumes "u \<up> x"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2724
  shows "u ----> x" (is ?lim) and "mono u" (is ?mono)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2725
proof -
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2726
  show ?mono using assms unfolding mono_iff_le_Suc isoton_def by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2727
  from Lim_eq_Sup_mono[OF this] `u \<up> x`
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2728
  show ?lim unfolding isoton_def by simp
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2729
qed
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2730
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2731
lemma isoton_iff_Lim_mono:
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2732
  fixes u :: "nat \<Rightarrow> pextreal"
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2733
  shows "u \<up> x \<longleftrightarrow> (mono u \<and> u ----> x)"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2734
proof safe
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2735
  assume "mono u" and x: "u ----> x"
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2736
  with SUP_Lim_pextreal[OF _ x]
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2737
  show "u \<up> x" unfolding isoton_def
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2738
    using `mono u`[unfolded mono_def]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2739
    using `mono u`[unfolded mono_iff_le_Suc]
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2740
    by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2741
qed (auto dest: isotone_Lim)
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2742
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2743
lemma pextreal_inverse_inverse[simp]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2744
  fixes x :: pextreal
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2745
  shows "inverse (inverse x) = x"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2746
  by (cases x) auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2747
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2748
lemma atLeastAtMost_omega_eq_atLeast:
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2749
  shows "{a .. \<omega>} = {a ..}"
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2750
by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2751
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2752
lemma atLeast0AtMost_eq_atMost: "{0 :: pextreal .. a} = {.. a}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2753
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2754
lemma greaterThan_omega_Empty: "{\<omega> <..} = {}" by auto
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2755
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2756
lemma lessThan_0_Empty: "{..< 0 :: pextreal} = {}" by auto
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2757
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2758
lemma real_of_pextreal_inverse[simp]:
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2759
  fixes X :: pextreal
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2760
  shows "real (inverse X) = 1 / real X"
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2761
  by (cases X) (auto simp: inverse_eq_divide)
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2762
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2763
lemma real_of_pextreal_le_0[simp]: "real (X :: pextreal) \<le> 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2764
  by (cases X) auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2765
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2766
lemma real_of_pextreal_less_0[simp]: "\<not> (real (X :: pextreal) < 0)"
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2767
  by (cases X) auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2768
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2769
lemma abs_real_of_pextreal[simp]: "\<bar>real (X :: pextreal)\<bar> = real X"
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2770
  by simp
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2771
41023
9118eb4eb8dc it is known as the extended reals, not the infinite reals
hoelzl
parents: 40874
diff changeset
  2772
lemma zero_less_real_of_pextreal: "0 < real (X :: pextreal) \<longleftrightarrow> X \<noteq> 0 \<and> X \<noteq> \<omega>"
40871
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2773
  by (cases X) auto
688f6ff859e1 Generalized simple_functionD and less_SUP_iff.
hoelzl
parents: 40859
diff changeset
  2774
38656
d5d342611edb Rewrite the Probability theory.
hoelzl
parents:
diff changeset
  2775
end