Theory Sigma_Algebra

(*  Title:      HOL/Analysis/Sigma_Algebra.thy
    Author:     Stefan Richter, Markus Wenzel, TU München
    Author:     Johannes Hölzl, TU München
    Plus material from the Hurd/Coble measure theory development,
    translated by Lawrence Paulson.
*)

chapter ‹Measure and Integration Theory›

theory Sigma_Algebra
imports
  Complex_Main
  "HOL-Library.Countable_Set"
  "HOL-Library.FuncSet"
  "HOL-Library.Indicator_Function"
  "HOL-Library.Extended_Nonnegative_Real"
  "HOL-Library.Disjoint_Sets"
begin


section ‹Sigma Algebra›

text ‹Sigma algebras are an elementary concept in measure
  theory. To measure --- that is to integrate --- functions, we first have
  to measure sets. Unfortunately, when dealing with a large universe,
  it is often not possible to consistently assign a measure to every
  subset. Therefore it is necessary to define the set of measurable
  subsets of the universe. A sigma algebra is such a set that has
  three very natural and desirable properties.›

subsection ‹Families of sets›

localetag important› subset_class =
  fixes Ω :: "'a set" and M :: "'a set set"
  assumes space_closed: "M  Pow Ω"

lemma (in subset_class) sets_into_space: "x  M  x  Ω"
  by (metis PowD contra_subsetD space_closed)

subsubsection ‹Semiring of sets›

localetag important› semiring_of_sets = subset_class +
  assumes empty_sets[iff]: "{}  M"
  assumes Int[intro]: "a b. a  M  b  M  a  b  M"
  assumes Diff_cover:
    "a b. a  M  b  M  CM. finite C  disjoint C  a - b = C"

lemma (in semiring_of_sets) finite_INT[intro]:
  assumes "finite I" "I  {}" "i. i  I  A i  M"
  shows "(iI. A i)  M"
  using assms by (induct rule: finite_ne_induct) auto

lemma (in semiring_of_sets) Int_space_eq1 [simp]: "x  M  Ω  x = x"
  by (metis Int_absorb1 sets_into_space)

lemma (in semiring_of_sets) Int_space_eq2 [simp]: "x  M  x  Ω = x"
  by (metis Int_absorb2 sets_into_space)

lemma (in semiring_of_sets) sets_Collect_conj:
  assumes "{xΩ. P x}  M" "{xΩ. Q x}  M"
  shows "{xΩ. Q x  P x}  M"
proof -
  have "{xΩ. Q x  P x} = {xΩ. Q x}  {xΩ. P x}"
    by auto
  with assms show ?thesis by auto
qed

lemma (in semiring_of_sets) sets_Collect_finite_All':
  assumes "i. i  S  {xΩ. P i x}  M" "finite S" "S  {}"
  shows "{xΩ. iS. P i x}  M"
proof -
  have "{xΩ. iS. P i x} = (iS. {xΩ. P i x})"
    using S  {} by auto
  with assms show ?thesis by auto
qed

subsubsection ‹Ring of sets›

localetag important› ring_of_sets = semiring_of_sets +
  assumes Un [intro]: "a b. a  M  b  M  a  b  M"

lemma (in ring_of_sets) finite_Union [intro]:
  "finite X  X  M  X  M"
  by (induct set: finite) (auto simp add: Un)

lemma (in ring_of_sets) finite_UN[intro]:
  assumes "finite I" and "i. i  I  A i  M"
  shows "(iI. A i)  M"
  using assms by induct auto

lemma (in ring_of_sets) Diff [intro]:
  assumes "a  M" "b  M" shows "a - b  M"
  using Diff_cover[OF assms] by auto

lemma ring_of_setsI:
  assumes space_closed: "M  Pow Ω"
  assumes empty_sets[iff]: "{}  M"
  assumes Un[intro]: "a b. a  M  b  M  a  b  M"
  assumes Diff[intro]: "a b. a  M  b  M  a - b  M"
  shows "ring_of_sets Ω M"
proof
  fix a b assume ab: "a  M" "b  M"
  from ab show "CM. finite C  disjoint C  a - b = C"
    by (intro exI[of _ "{a - b}"]) (auto simp: disjoint_def)
  have "a  b = a - (a - b)" by auto
  also have "  M" using ab by auto
  finally show "a  b  M" .
qed fact+

lemma ring_of_sets_iff: "ring_of_sets Ω M  M  Pow Ω  {}  M  (aM. bM. a  b  M)  (aM. bM. a - b  M)"
proof
  assume "ring_of_sets Ω M"
  then interpret ring_of_sets Ω M .
  show "M  Pow Ω  {}  M  (aM. bM. a  b  M)  (aM. bM. a - b  M)"
    using space_closed by auto
qed (auto intro!: ring_of_setsI)

lemma (in ring_of_sets) insert_in_sets:
  assumes "{x}  M" "A  M" shows "insert x A  M"
proof -
  have "{x}  A  M" using assms by (rule Un)
  thus ?thesis by auto
qed

lemma (in ring_of_sets) sets_Collect_disj:
  assumes "{xΩ. P x}  M" "{xΩ. Q x}  M"
  shows "{xΩ. Q x  P x}  M"
proof -
  have "{xΩ. Q x  P x} = {xΩ. Q x}  {xΩ. P x}"
    by auto
  with assms show ?thesis by auto
qed

lemma (in ring_of_sets) sets_Collect_finite_Ex:
  assumes "i. i  S  {xΩ. P i x}  M" "finite S"
  shows "{xΩ. iS. P i x}  M"
proof -
  have "{xΩ. iS. P i x} = (iS. {xΩ. P i x})"
    by auto
  with assms show ?thesis by auto
qed

subsubsection ‹Algebra of sets›

localetag important› algebra = ring_of_sets +
  assumes top [iff]: "Ω  M"

lemma (in algebra) compl_sets [intro]:
  "a  M  Ω - a  M"
  by auto

proposition algebra_iff_Un:
  "algebra Ω M 
    M  Pow Ω 
    {}  M 
    (a  M. Ω - a  M) 
    (a  M.  b  M. a  b  M)" (is "_  ?Un")
proof
  assume "algebra Ω M"
  then interpret algebra Ω M .
  show ?Un using sets_into_space by auto
next
  assume ?Un
  then have "Ω  M" by auto
  interpret ring_of_sets Ω M
  proof (rule ring_of_setsI)
    show Ω: "M  Pow Ω" "{}  M"
      using ?Un by auto
    fix a b assume a: "a  M" and b: "b  M"
    then show "a  b  M" using ?Un by auto
    have "a - b = Ω - ((Ω - a)  b)"
      using Ω a b by auto
    then show "a - b  M"
      using a b  ?Un by auto
  qed
  show "algebra Ω M" proof qed fact
qed

proposition algebra_iff_Int:
     "algebra Ω M 
       M  Pow Ω & {}  M &
       (a  M. Ω - a  M) &
       (a  M.  b  M. a  b  M)" (is "_  ?Int")
proof
  assume "algebra Ω M"
  then interpret algebra Ω M .
  show ?Int using sets_into_space by auto
next
  assume ?Int
  show "algebra Ω M"
  proof (unfold algebra_iff_Un, intro conjI ballI)
    show Ω: "M  Pow Ω" "{}  M"
      using ?Int by auto
    from ?Int show "a. a  M  Ω - a  M" by auto
    fix a b assume M: "a  M" "b  M"
    hence "a  b = Ω - ((Ω - a)  (Ω - b))"
      using Ω by blast
    also have "...  M"
      using M ?Int by auto
    finally show "a  b  M" .
  qed
qed

lemma (in algebra) sets_Collect_neg:
  assumes "{xΩ. P x}  M"
  shows "{xΩ. ¬ P x}  M"
proof -
  have "{xΩ. ¬ P x} = Ω - {xΩ. P x}" by auto
  with assms show ?thesis by auto
qed

lemma (in algebra) sets_Collect_imp:
  "{xΩ. P x}  M  {xΩ. Q x}  M  {xΩ. Q x  P x}  M"
  unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg)

lemma (in algebra) sets_Collect_const:
  "{xΩ. P}  M"
  by (cases P) auto

lemma algebra_single_set:
  "X  S  algebra S { {}, X, S - X, S }"
  by (auto simp: algebra_iff_Int)

subsubsectiontag unimportant› ‹Restricted algebras›

abbreviation (in algebra)
  "restricted_space A  ((∩) A) ` M"

lemma (in algebra) restricted_algebra:
  assumes "A  M" shows "algebra A (restricted_space A)"
  using assms by (auto simp: algebra_iff_Int)

subsubsection ‹Sigma Algebras›

localetag important› sigma_algebra = algebra +
  assumes countable_nat_UN [intro]: "A. range A  M  (i::nat. A i)  M"

lemma (in algebra) is_sigma_algebra:
  assumes "finite M"
  shows "sigma_algebra Ω M"
proof
  fix A :: "nat  'a set" assume "range A  M"
  then have "(i. A i) = (sM  range A. s)"
    by auto
  also have "(sM  range A. s)  M"
    using finite M by auto
  finally show "(i. A i)  M" .
qed

lemma countable_UN_eq:
  fixes A :: "'i::countable  'a set"
  shows "(range A  M  (i. A i)  M) 
    (range (A  from_nat)  M  (i. (A  from_nat) i)  M)"
proof -
  let ?A' = "A  from_nat"
  have *: "(i. ?A' i) = (i. A i)" (is "?l = ?r")
  proof safe
    fix x i assume "x  A i" thus "x  ?l"
      by (auto intro!: exI[of _ "to_nat i"])
  next
    fix x i assume "x  ?A' i" thus "x  ?r"
      by (auto intro!: exI[of _ "from_nat i"])
  qed
  have "A ` range from_nat = range A"
    using surj_from_nat by simp
  then have **: "range ?A' = range A"
    by (simp only: image_comp [symmetric])
  show ?thesis unfolding * ** ..
qed

lemma (in sigma_algebra) countable_Union [intro]:
  assumes "countable X" "X  M" shows "X  M"
proof cases
  assume "X  {}"
  hence "X = (n. from_nat_into X n)"
    using assms by (auto cong del: SUP_cong)
  also have "  M" using assms
    by (auto intro!: countable_nat_UN) (metis X  {} from_nat_into subsetD)
  finally show ?thesis .
qed simp

lemma (in sigma_algebra) countable_UN[intro]:
  fixes A :: "'i::countable  'a set"
  assumes "A`X  M"
  shows  "(xX. A x)  M"
proof -
  let ?A = "λi. if i  X then A i else {}"
  from assms have "range ?A  M" by auto
  with countable_nat_UN[of "?A  from_nat"] countable_UN_eq[of ?A M]
  have "(x. ?A x)  M" by auto
  moreover have "(x. ?A x) = (xX. A x)" by (auto split: if_split_asm)
  ultimately show ?thesis by simp
qed

lemma (in sigma_algebra) countable_UN':
  fixes A :: "'i  'a set"
  assumes X: "countable X"
  assumes A: "A`X  M"
  shows  "(xX. A x)  M"
proof -
  have "(xX. A x) = (ito_nat_on X ` X. A (from_nat_into X i))"
    using X by auto
  also have "  M"
    using A X
    by (intro countable_UN) auto
  finally show ?thesis .
qed

lemma (in sigma_algebra) countable_UN'':
  " countable X; x y. x  X  A x  M   (xX. A x)  M"
by(erule countable_UN')(auto)

lemma (in sigma_algebra) countable_INT [intro]:
  fixes A :: "'i::countable  'a set"
  assumes A: "A`X  M" "X  {}"
  shows "(iX. A i)  M"
proof -
  from A have "iX. A i  M" by fast
  hence "Ω - (iX. Ω - A i)  M" by blast
  moreover
  have "(iX. A i) = Ω - (iX. Ω - A i)" using space_closed A
    by blast
  ultimately show ?thesis by metis
qed

lemma (in sigma_algebra) countable_INT':
  fixes A :: "'i  'a set"
  assumes X: "countable X" "X  {}"
  assumes A: "A`X  M"
  shows  "(xX. A x)  M"
proof -
  have "(xX. A x) = (ito_nat_on X ` X. A (from_nat_into X i))"
    using X by auto
  also have "  M"
    using A X
    by (intro countable_INT) auto
  finally show ?thesis .
qed

lemma (in sigma_algebra) countable_INT'':
  "UNIV  M  countable I  (i. i  I  F i  M)  (iI. F i)  M"
  by (cases "I = {}") (auto intro: countable_INT')

lemma (in sigma_algebra) countable:
  assumes "a. a  A  {a}  M" "countable A"
  shows "A  M"
proof -
  have "(aA. {a})  M"
    using assms by (intro countable_UN') auto
  also have "(aA. {a}) = A" by auto
  finally show ?thesis by auto
qed

lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)"
  by (auto simp: ring_of_sets_iff)

lemma algebra_Pow: "algebra sp (Pow sp)"
  by (auto simp: algebra_iff_Un)

lemma sigma_algebra_iff:
  "sigma_algebra Ω M 
    algebra Ω M  (A. range A  M  (i::nat. A i)  M)"
  by (simp add: sigma_algebra_def sigma_algebra_axioms_def)

lemma sigma_algebra_Pow: "sigma_algebra sp (Pow sp)"
  by (auto simp: sigma_algebra_iff algebra_iff_Int)

lemma (in sigma_algebra) sets_Collect_countable_All:
  assumes "i. {xΩ. P i x}  M"
  shows "{xΩ. i::'i::countable. P i x}  M"
proof -
  have "{xΩ. i::'i::countable. P i x} = (i. {xΩ. P i x})" by auto
  with assms show ?thesis by auto
qed

lemma (in sigma_algebra) sets_Collect_countable_Ex:
  assumes "i. {xΩ. P i x}  M"
  shows "{xΩ. i::'i::countable. P i x}  M"
proof -
  have "{xΩ. i::'i::countable. P i x} = (i. {xΩ. P i x})" by auto
  with assms show ?thesis by auto
qed

lemma (in sigma_algebra) sets_Collect_countable_Ex':
  assumes "i. i  I  {xΩ. P i x}  M"
  assumes "countable I"
  shows "{xΩ. iI. P i x}  M"
proof -
  have "{xΩ. iI. P i x} = (iI. {xΩ. P i x})" by auto
  with assms show ?thesis
    by (auto intro!: countable_UN')
qed

lemma (in sigma_algebra) sets_Collect_countable_All':
  assumes "i. i  I  {xΩ. P i x}  M"
  assumes "countable I"
  shows "{xΩ. iI. P i x}  M"
proof -
  have "{xΩ. iI. P i x} = (iI. {xΩ. P i x})  Ω" by auto
  with assms show ?thesis
    by (cases "I = {}") (auto intro!: countable_INT')
qed

lemma (in sigma_algebra) sets_Collect_countable_Ex1':
  assumes "i. i  I  {xΩ. P i x}  M"
  assumes "countable I"
  shows "{xΩ. ∃!iI. P i x}  M"
proof -
  have "{xΩ. ∃!iI. P i x} = {xΩ. iI. P i x  (jI. P j x  i = j)}"
    by auto
  with assms show ?thesis
    by (auto intro!: sets_Collect_countable_All' sets_Collect_countable_Ex' sets_Collect_conj sets_Collect_imp sets_Collect_const)
qed

lemmas (in sigma_algebra) sets_Collect =
  sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const
  sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All

lemma (in sigma_algebra) sets_Collect_countable_Ball:
  assumes "i. {xΩ. P i x}  M"
  shows "{xΩ. i::'i::countableX. P i x}  M"
  unfolding Ball_def by (intro sets_Collect assms)

lemma (in sigma_algebra) sets_Collect_countable_Bex:
  assumes "i. {xΩ. P i x}  M"
  shows "{xΩ. i::'i::countableX. P i x}  M"
  unfolding Bex_def by (intro sets_Collect assms)

lemma sigma_algebra_single_set:
  assumes "X  S"
  shows "sigma_algebra S { {}, X, S - X, S }"
  using algebra.is_sigma_algebra[OF algebra_single_set[OF X  S]] by simp

subsubsectiontag unimportant› ‹Binary Unions›

definition binary :: "'a  'a  nat  'a"
  where "binary a b =  (λx. b)(0 := a)"

lemma range_binary_eq: "range(binary a b) = {a,b}"
  by (auto simp add: binary_def)

lemma Un_range_binary: "a  b = (i::nat. binary a b i)"
  by (simp add: range_binary_eq cong del: SUP_cong_simp)

lemma Int_range_binary: "a  b = (i::nat. binary a b i)"
  by (simp add: range_binary_eq cong del: INF_cong_simp)

lemma sigma_algebra_iff2:
  "sigma_algebra Ω M 
    M  Pow Ω  {}  M  (s  M. Ω - s  M)
     (A. range A  M ( i::nat. A i)  M)" (is "?P  ?R  ?S  ?V  ?W")
proof
  assume ?P
  then interpret sigma_algebra Ω M .
  from space_closed show "?R  ?S  ?V  ?W"
    by auto
next
  assume "?R  ?S  ?V  ?W"
  then have ?R ?S ?V ?W
    by simp_all
  show ?P
  proof (rule sigma_algebra.intro)
    show "sigma_algebra_axioms M"
      by standard (use ?W in simp)
    from ?W have *: "range (binary a b)  M   (range (binary a b))  M" for a b
      by auto
    show "algebra Ω M"
      unfolding algebra_iff_Un using ?R ?S ?V *
      by (auto simp add: range_binary_eq)
  qed
qed


subsubsection ‹Initial Sigma Algebra›

texttag important› ‹Sigma algebras can naturally be created as the closure of any set of
  M with regard to the properties just postulated.›

inductive_settag important› sigma_sets :: "'a set  'a set set  'a set set"
  for sp :: "'a set" and A :: "'a set set"
  where
    Basic[intro, simp]: "a  A  a  sigma_sets sp A"
  | Empty: "{}  sigma_sets sp A"
  | Compl: "a  sigma_sets sp A  sp - a  sigma_sets sp A"
  | Union: "(i::nat. a i  sigma_sets sp A)  (i. a i)  sigma_sets sp A"

lemma (in sigma_algebra) sigma_sets_subset:
  assumes a: "a  M"
  shows "sigma_sets Ω a  M"
proof
  fix x
  assume "x  sigma_sets Ω a"
  from this show "x  M"
    by (induct rule: sigma_sets.induct, auto) (metis a subsetD)
qed

lemma sigma_sets_into_sp: "A  Pow sp  x  sigma_sets sp A  x  sp"
  by (erule sigma_sets.induct, auto)

lemma sigma_algebra_sigma_sets:
     "a  Pow Ω  sigma_algebra Ω (sigma_sets Ω a)"
  by (auto simp add: sigma_algebra_iff2 dest: sigma_sets_into_sp
           intro!: sigma_sets.Union sigma_sets.Empty sigma_sets.Compl)

lemma sigma_sets_least_sigma_algebra:
  assumes "A  Pow S"
  shows "sigma_sets S A = {B. A  B  sigma_algebra S B}"
proof safe
  fix B X assume "A  B" and sa: "sigma_algebra S B"
    and X: "X  sigma_sets S A"
  from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF A  B] X
  show "X  B" by auto
next
  fix X assume "X  {B. A  B  sigma_algebra S B}"
  then have [intro!]: "B. A  B  sigma_algebra S B  X  B"
     by simp
  have "A  sigma_sets S A" using assms by auto
  moreover have "sigma_algebra S (sigma_sets S A)"
    using assms by (intro sigma_algebra_sigma_sets[of A]) auto
  ultimately show "X  sigma_sets S A" by auto
qed

lemma sigma_sets_top: "sp  sigma_sets sp A"
  by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty)

lemma binary_in_sigma_sets:
  "binary a b i  sigma_sets sp A" if "a  sigma_sets sp A" and "b  sigma_sets sp A"
  using that by (simp add: binary_def)

lemma sigma_sets_Un:
  "a  b  sigma_sets sp A" if "a  sigma_sets sp A" and "b  sigma_sets sp A"
  using that by (simp add: Un_range_binary binary_in_sigma_sets Union)

lemma sigma_sets_Inter:
  assumes Asb: "A  Pow sp"
  shows "(i::nat. a i  sigma_sets sp A)  (i. a i)  sigma_sets sp A"
proof -
  assume ai: "i::nat. a i  sigma_sets sp A"
  hence "i::nat. sp-(a i)  sigma_sets sp A"
    by (rule sigma_sets.Compl)
  hence "(i. sp-(a i))  sigma_sets sp A"
    by (rule sigma_sets.Union)
  hence "sp-(i. sp-(a i))  sigma_sets sp A"
    by (rule sigma_sets.Compl)
  also have "sp-(i. sp-(a i)) = sp Int (i. a i)"
    by auto
  also have "... = (i. a i)" using ai
    by (blast dest: sigma_sets_into_sp [OF Asb])
  finally show ?thesis .
qed

lemma sigma_sets_INTER:
  assumes Asb: "A  Pow sp"
      and ai: "i::nat. i  S  a i  sigma_sets sp A" and non: "S  {}"
  shows "(iS. a i)  sigma_sets sp A"
proof -
  from ai have "i. (if iS then a i else sp)  sigma_sets sp A"
    by (simp add: sigma_sets.intros(2-) sigma_sets_top)
  hence "(i. (if iS then a i else sp))  sigma_sets sp A"
    by (rule sigma_sets_Inter [OF Asb])
  also have "(i. (if iS then a i else sp)) = (iS. a i)"
    by auto (metis ai non sigma_sets_into_sp subset_empty subset_iff Asb)+
  finally show ?thesis .
qed

lemma sigma_sets_UNION:
  "countable B  (b. b  B  b  sigma_sets X A)   B  sigma_sets X A"
  using from_nat_into [of B] range_from_nat_into [of B] sigma_sets.Union [of "from_nat_into B" X A]
  by (cases "B = {}") (simp_all add: sigma_sets.Empty cong del: SUP_cong)

lemma (in sigma_algebra) sigma_sets_eq:
     "sigma_sets Ω M = M"
proof
  show "M  sigma_sets Ω M"
    by (metis Set.subsetI sigma_sets.Basic)
  next
  show "sigma_sets Ω M  M"
    by (metis sigma_sets_subset subset_refl)
qed

lemma sigma_sets_eqI:
  assumes A: "a. a  A  a  sigma_sets M B"
  assumes B: "b. b  B  b  sigma_sets M A"
  shows "sigma_sets M A = sigma_sets M B"
proof (intro set_eqI iffI)
  fix a assume "a  sigma_sets M A"
  from this A show "a  sigma_sets M B"
    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
next
  fix b assume "b  sigma_sets M B"
  from this B show "b  sigma_sets M A"
    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
qed

lemma sigma_sets_subseteq: assumes "A  B" shows "sigma_sets X A  sigma_sets X B"
proof
  fix x assume "x  sigma_sets X A" then show "x  sigma_sets X B"
    by induct (insert A  B, auto intro: sigma_sets.intros(2-))
qed

lemma sigma_sets_mono: assumes "A  sigma_sets X B" shows "sigma_sets X A  sigma_sets X B"
proof
  fix x assume "x  sigma_sets X A" then show "x  sigma_sets X B"
    by induct (insert A  sigma_sets X B, auto intro: sigma_sets.intros(2-))
qed

lemma sigma_sets_mono': assumes "A  B" shows "sigma_sets X A  sigma_sets X B"
proof
  fix x assume "x  sigma_sets X A" then show "x  sigma_sets X B"
    by induct (insert A  B, auto intro: sigma_sets.intros(2-))
qed

lemma sigma_sets_superset_generator: "A  sigma_sets X A"
  by (auto intro: sigma_sets.Basic)

lemma (in sigma_algebra) restriction_in_sets:
  fixes A :: "nat  'a set"
  assumes "S  M"
  and *: "range A  (λA. S  A) ` M" (is "_  ?r")
  shows "range A  M" "(i. A i)  (λA. S  A) ` M"
proof -
  { fix i have "A i  ?r" using * by auto
    hence "B. A i = B  S  B  M" by auto
    hence "A i  S" "A i  M" using S  M by auto }
  thus "range A  M" "(i. A i)  (λA. S  A) ` M"
    by (auto intro!: image_eqI[of _ _ "(i. A i)"])
qed

lemma (in sigma_algebra) restricted_sigma_algebra:
  assumes "S  M"
  shows "sigma_algebra S (restricted_space S)"
  unfolding sigma_algebra_def sigma_algebra_axioms_def
proof safe
  show "algebra S (restricted_space S)" using restricted_algebra[OF assms] .
next
  fix A :: "nat  'a set" assume "range A  restricted_space S"
  from restriction_in_sets[OF assms this[simplified]]
  show "(i. A i)  restricted_space S" by simp
qed

lemma sigma_sets_Int:
  assumes "A  sigma_sets sp st" "A  sp"
  shows "(∩) A ` sigma_sets sp st = sigma_sets A ((∩) A ` st)"
proof (intro equalityI subsetI)
  fix x assume "x  (∩) A ` sigma_sets sp st"
  then obtain y where "y  sigma_sets sp st" "x = y  A" by auto
  then have "x  sigma_sets (A  sp) ((∩) A ` st)"
  proof (induct arbitrary: x)
    case (Compl a)
    then show ?case
      by (force intro!: sigma_sets.Compl simp: Diff_Int_distrib ac_simps)
  next
    case (Union a)
    then show ?case
      by (auto intro!: sigma_sets.Union
               simp add: UN_extend_simps simp del: UN_simps)
  qed (auto intro!: sigma_sets.intros(2-))
  then show "x  sigma_sets A ((∩) A ` st)"
    using A  sp by (simp add: Int_absorb2)
next
  fix x assume "x  sigma_sets A ((∩) A ` st)"
  then show "x  (∩) A ` sigma_sets sp st"
  proof induct
    case (Compl a)
    then obtain x where "a = A  x" "x  sigma_sets sp st" by auto
    then show ?case using A  sp
      by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl)
  next
    case (Union a)
    then have "i. x. x  sigma_sets sp st  a i = A  x"
      by (auto simp: image_iff Bex_def)
    then obtain f where "x. f x  sigma_sets sp st  a x = A  f x"
      by metis
    then show ?case
      by (auto intro!: bexI[of _ "(x. f x)"] sigma_sets.Union
               simp add: image_iff)
  qed (auto intro!: sigma_sets.intros(2-))
qed

lemma sigma_sets_empty_eq: "sigma_sets A {} = {{}, A}"
proof (intro set_eqI iffI)
  fix a assume "a  sigma_sets A {}" then show "a  {{}, A}"
    by induct blast+
qed (auto intro: sigma_sets.Empty sigma_sets_top)

lemma sigma_sets_single[simp]: "sigma_sets A {A} = {{}, A}"
proof (intro set_eqI iffI)
  fix x assume "x  sigma_sets A {A}"
  then show "x  {{}, A}"
    by induct blast+
next
  fix x assume "x  {{}, A}"
  then show "x  sigma_sets A {A}"
    by (auto intro: sigma_sets.Empty sigma_sets_top)
qed

lemma sigma_sets_sigma_sets_eq:
  "M  Pow S  sigma_sets S (sigma_sets S M) = sigma_sets S M"
  by (rule sigma_algebra.sigma_sets_eq[OF sigma_algebra_sigma_sets, of M S]) auto

lemma sigma_sets_singleton:
  assumes "X  S"
  shows "sigma_sets S { X } = { {}, X, S - X, S }"
proof -
  interpret sigma_algebra S "{ {}, X, S - X, S }"
    by (rule sigma_algebra_single_set) fact
  have "sigma_sets S { X }  sigma_sets S { {}, X, S - X, S }"
    by (rule sigma_sets_subseteq) simp
  moreover have " = { {}, X, S - X, S }"
    using sigma_sets_eq by simp
  moreover
  { fix A assume "A  { {}, X, S - X, S }"
    then have "A  sigma_sets S { X }"
      by (auto intro: sigma_sets.intros(2-) sigma_sets_top) }
  ultimately have "sigma_sets S { X } = sigma_sets S { {}, X, S - X, S }"
    by (intro antisym) auto
  with sigma_sets_eq show ?thesis by simp
qed

lemma restricted_sigma:
  assumes S: "S  sigma_sets Ω M" and M: "M  Pow Ω"
  shows "algebra.restricted_space (sigma_sets Ω M) S =
    sigma_sets S (algebra.restricted_space M S)"
proof -
  from S sigma_sets_into_sp[OF M]
  have "S  sigma_sets Ω M" "S  Ω" by auto
  from sigma_sets_Int[OF this]
  show ?thesis by simp
qed

lemma sigma_sets_vimage_commute:
  assumes X: "X  Ω  Ω'"
  shows "{X -` A  Ω |A. A  sigma_sets Ω' M'}
       = sigma_sets Ω {X -` A  Ω |A. A  M'}" (is "?L = ?R")
proof
  show "?L  ?R"
  proof clarify
    fix A assume "A  sigma_sets Ω' M'"
    then show "X -` A  Ω  ?R"
    proof induct
      case Empty then show ?case
        by (auto intro!: sigma_sets.Empty)
    next
      case (Compl B)
      have [simp]: "X -` (Ω' - B)  Ω = Ω - (X -` B  Ω)"
        by (auto simp add: funcset_mem [OF X])
      with Compl show ?case
        by (auto intro!: sigma_sets.Compl)
    next
      case (Union F)
      then show ?case
        by (auto simp add: vimage_UN UN_extend_simps(4) simp del: UN_simps
                 intro!: sigma_sets.Union)
    qed auto
  qed
  show "?R  ?L"
  proof clarify
    fix A assume "A  ?R"
    then show "B. A = X -` B  Ω  B  sigma_sets Ω' M'"
    proof induct
      case (Basic B) then show ?case by auto
    next
      case Empty then show ?case
        by (auto intro!: sigma_sets.Empty exI[of _ "{}"])
    next
      case (Compl B)
      then obtain A where A: "B = X -` A  Ω" "A  sigma_sets Ω' M'" by auto
      then have [simp]: "Ω - B = X -` (Ω' - A)  Ω"
        by (auto simp add: funcset_mem [OF X])
      with A(2) show ?case
        by (auto intro: sigma_sets.Compl)
    next
      case (Union F)
      then have "i. B. F i = X -` B  Ω  B  sigma_sets Ω' M'" by auto
      then obtain A where "x. F x = X -` A x  Ω  A x  sigma_sets Ω' M'"
        by metis
      then show ?case
        by (auto simp: vimage_UN[symmetric] intro: sigma_sets.Union)
    qed
  qed
qed

lemma (in ring_of_sets) UNION_in_sets:
  fixes A:: "nat  'a set"
  assumes A: "range A  M"
  shows  "(i{0..<n}. A i)  M"
proof (induct n)
  case 0 show ?case by simp
next
  case (Suc n)
  thus ?case
    by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff)
qed

lemma (in ring_of_sets) range_disjointed_sets:
  assumes A: "range A  M"
  shows  "range (disjointed A)  M"
proof (auto simp add: disjointed_def)
  fix n
  show "A n - (i{0..<n}. A i)  M" using UNION_in_sets
    by (metis A Diff UNIV_I image_subset_iff)
qed

lemma (in algebra) range_disjointed_sets':
  "range A  M  range (disjointed A)  M"
  using range_disjointed_sets .

lemma sigma_algebra_disjoint_iff:
  "sigma_algebra Ω M  algebra Ω M 
    (A. range A  M  disjoint_family A  (i::nat. A i)  M)"
proof (auto simp add: sigma_algebra_iff)
  fix A :: "nat  'a set"
  assume M: "algebra Ω M"
     and A: "range A  M"
     and UnA: "A. range A  M  disjoint_family A  (i::nat. A i)  M"
  hence "range (disjointed A)  M 
         disjoint_family (disjointed A) 
         (i. disjointed A i)  M" by blast
  hence "(i. disjointed A i)  M"
    by (simp add: algebra.range_disjointed_sets'[of Ω] M A disjoint_family_disjointed)
  thus "(i::nat. A i)  M" by (simp add: UN_disjointed_eq)
qed

subsubsectiontag unimportant› ‹Ring generated by a semiring›

definition (in semiring_of_sets) generated_ring :: "'a set set" where
  "generated_ring = { C | C. C  M  finite C  disjoint C }"

lemma (in semiring_of_sets) generated_ringE[elim?]:
  assumes "a  generated_ring"
  obtains C where "finite C" "disjoint C" "C  M" "a = C"
  using assms unfolding generated_ring_def by auto

lemma (in semiring_of_sets) generated_ringI[intro?]:
  assumes "finite C" "disjoint C" "C  M" "a = C"
  shows "a  generated_ring"
  using assms unfolding generated_ring_def by auto

lemma (in semiring_of_sets) generated_ringI_Basic:
  "A  M  A  generated_ring"
  by (rule generated_ringI[of "{A}"]) (auto simp: disjoint_def)

lemma (in semiring_of_sets) generated_ring_disjoint_Un[intro]:
  assumes a: "a  generated_ring" and b: "b  generated_ring"
  and "a  b = {}"
  shows "a  b  generated_ring"
proof -
  from a b obtain Ca Cb
    where Ca: "finite Ca" "disjoint Ca" "Ca  M" "a =  Ca"
      and Cb: "finite Cb" "disjoint Cb" "Cb  M" "b =  Cb"
    using generated_ringE by metis
  show ?thesis
  proof
    from a  b = {} Ca Cb show "disjoint (Ca  Cb)"
      by (auto intro!: disjoint_union)
  qed (use Ca Cb in auto)
qed

lemma (in semiring_of_sets) generated_ring_empty: "{}  generated_ring"
  by (auto simp: generated_ring_def disjoint_def)

lemma (in semiring_of_sets) generated_ring_disjoint_Union:
  assumes "finite A" shows "A  generated_ring  disjoint A  A  generated_ring"
  using assms by (induct A) (auto simp: disjoint_def intro!: generated_ring_disjoint_Un generated_ring_empty)

lemma (in semiring_of_sets) generated_ring_disjoint_UNION:
  "finite I  disjoint (A ` I)  (i. i  I  A i  generated_ring)  (A ` I)  generated_ring"
  by (intro generated_ring_disjoint_Union) auto

lemma (in semiring_of_sets) generated_ring_Int:
  assumes a: "a  generated_ring" and b: "b  generated_ring"
  shows "a  b  generated_ring"
proof -
  from a b obtain Ca Cb
    where Ca: "finite Ca" "disjoint Ca" "Ca  M" "a =  Ca"
      and Cb: "finite Cb" "disjoint Cb" "Cb  M" "b =  Cb"
    using generated_ringE by metis
  define C where "C = (λ(a,b). a  b)` (Ca×Cb)"
  show ?thesis
  proof
    show "disjoint C"
    proof (simp add: disjoint_def C_def, intro ballI impI)
      fix a1 b1 a2 b2 assume sets: "a1  Ca" "b1  Cb" "a2  Ca" "b2  Cb"
      assume "a1  b1  a2  b2"
      then have "a1  a2  b1  b2" by auto
      then show "(a1  b1)  (a2  b2) = {}"
      proof
        assume "a1  a2"
        with sets Ca have "a1  a2 = {}"
          by (auto simp: disjoint_def)
        then show ?thesis by auto
      next
        assume "b1  b2"
        with sets Cb have "b1  b2 = {}"
          by (auto simp: disjoint_def)
        then show ?thesis by auto
      qed
    qed
  qed (use Ca Cb in auto simp: C_def)
qed

lemma (in semiring_of_sets) generated_ring_Inter:
  assumes "finite A" "A  {}" shows "A  generated_ring  A  generated_ring"
  using assms by (induct A rule: finite_ne_induct) (auto intro: generated_ring_Int)

lemma (in semiring_of_sets) generated_ring_INTER:
  "finite I  I  {}  (i. i  I  A i  generated_ring)  (A ` I)  generated_ring"
  by (intro generated_ring_Inter) auto

lemma (in semiring_of_sets) generating_ring:
  "ring_of_sets Ω generated_ring"
proof (rule ring_of_setsI)
  let ?R = generated_ring
  show "?R  Pow Ω"
    using sets_into_space by (auto simp: generated_ring_def generated_ring_empty)
  show "{}  ?R" by (rule generated_ring_empty)

  {
    fix a b assume "a  ?R" "b  ?R"
    then obtain Ca Cb
      where Ca: "finite Ca" "disjoint Ca" "Ca  M" "a =  Ca"
        and Cb: "finite Cb" "disjoint Cb" "Cb  M" "b =  Cb"
      using generated_ringE by metis
    show "a - b  ?R"
    proof cases
      assume "Cb = {}" with Cb a  ?R show ?thesis
        by simp
    next
      assume "Cb  {}"
      with Ca Cb have "a - b = (a'Ca. b'Cb. a' - b')" by auto
      also have "  ?R"
      proof (intro generated_ring_INTER generated_ring_disjoint_UNION)
        fix a b assume "a  Ca" "b  Cb"
        with Ca Cb Diff_cover[of a b] show "a - b  ?R"
          by (auto simp add: generated_ring_def)
            (metis DiffI Diff_eq_empty_iff empty_iff)
      next
        show "disjoint ((λa'. b'Cb. a' - b')`Ca)"
          using Ca by (auto simp add: disjoint_def Cb  {})
      next
        show "finite Ca" "finite Cb" "Cb  {}" by fact+
      qed
      finally show "a - b  ?R" .
    qed
  }
  note Diff = this

  fix a b assume sets: "a  ?R" "b  ?R"
  have "a  b = (a - b)  (a  b)  (b - a)" by auto
  also have "  ?R"
    by (intro sets generated_ring_disjoint_Un generated_ring_Int Diff) auto
  finally show "a  b  ?R" .
qed

lemma (in semiring_of_sets) sigma_sets_generated_ring_eq: "sigma_sets Ω generated_ring = sigma_sets Ω M"
proof
  interpret M: sigma_algebra Ω "sigma_sets Ω M"
    using space_closed by (rule sigma_algebra_sigma_sets)
  show "sigma_sets Ω generated_ring  sigma_sets Ω M"
    by (blast intro!: sigma_sets_mono elim: generated_ringE)
qed (auto intro!: generated_ringI_Basic sigma_sets_mono)

subsubsectiontag unimportant› ‹A Two-Element Series›

definition binaryset :: "'a set  'a set  nat  'a set"
  where "binaryset A B = (λx. {})(0 := A, Suc 0 := B)"

lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
  apply (simp add: binaryset_def)
  apply (rule set_eqI)
  apply (auto simp add: image_iff)
  done

lemma UN_binaryset_eq: "(i. binaryset A B i) = A  B"
  by (simp add: range_binaryset_eq cong del: SUP_cong_simp)

subsubsection ‹Closed CDI›

definitiontag important› closed_cdi :: "'a set  'a set set  bool" where
  "closed_cdi Ω M 
   M  Pow Ω &
   (s  M. Ω - s  M) &
   (A. (range A  M) & (A 0 = {}) & (n. A n  A (Suc n)) 
        (i. A i)  M) &
   (A. (range A  M) & disjoint_family A  (i::nat. A i)  M)"

inductive_set
  smallest_ccdi_sets :: "'a set  'a set set  'a set set"
  for Ω M
  where
    Basic [intro]:
      "a  M  a  smallest_ccdi_sets Ω M"
  | Compl [intro]:
      "a  smallest_ccdi_sets Ω M  Ω - a  smallest_ccdi_sets Ω M"
  | Inc:
      "range A  Pow(smallest_ccdi_sets Ω M)  A 0 = {}  (n. A n  A (Suc n))
        (i. A i)  smallest_ccdi_sets Ω M"
  | Disj:
      "range A  Pow(smallest_ccdi_sets Ω M)  disjoint_family A
        (i::nat. A i)  smallest_ccdi_sets Ω M"

lemma (in subset_class) smallest_closed_cdi1: "M  smallest_ccdi_sets Ω M"
  by auto

lemma (in subset_class) smallest_ccdi_sets: "smallest_ccdi_sets Ω M  Pow Ω"
  apply (rule subsetI)
  apply (erule smallest_ccdi_sets.induct)
  apply (auto intro: range_subsetD dest: sets_into_space)
  done

lemma (in subset_class) smallest_closed_cdi2: "closed_cdi Ω (smallest_ccdi_sets Ω M)"
  apply (auto simp add: closed_cdi_def smallest_ccdi_sets)
  apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +
  done

lemma closed_cdi_subset: "closed_cdi Ω M  M  Pow Ω"
  by (simp add: closed_cdi_def)

lemma closed_cdi_Compl: "closed_cdi Ω M  s  M  Ω - s  M"
  by (simp add: closed_cdi_def)

lemma closed_cdi_Inc:
  "closed_cdi Ω M  range A  M  A 0 = {}  (!!n. A n  A (Suc n))  (i. A i)  M"
  by (simp add: closed_cdi_def)

lemma closed_cdi_Disj:
  "closed_cdi Ω M  range A  M  disjoint_family A  (i::nat. A i)  M"
  by (simp add: closed_cdi_def)

lemma closed_cdi_Un:
  assumes cdi: "closed_cdi Ω M" and empty: "{}  M"
      and A: "A  M" and B: "B  M"
      and disj: "A  B = {}"
    shows "A  B  M"
proof -
  have ra: "range (binaryset A B)  M"
   by (simp add: range_binaryset_eq empty A B)
 have di:  "disjoint_family (binaryset A B)" using disj
   by (simp add: disjoint_family_on_def binaryset_def Int_commute)
 from closed_cdi_Disj [OF cdi ra di]
 show ?thesis
   by (simp add: UN_binaryset_eq)
qed

lemma (in algebra) smallest_ccdi_sets_Un:
  assumes A: "A  smallest_ccdi_sets Ω M" and B: "B  smallest_ccdi_sets Ω M"
      and disj: "A  B = {}"
    shows "A  B  smallest_ccdi_sets Ω M"
proof -
  have ra: "range (binaryset A B)  Pow (smallest_ccdi_sets Ω M)"
    by (simp add: range_binaryset_eq  A B smallest_ccdi_sets.Basic)
  have di:  "disjoint_family (binaryset A B)" using disj
    by (simp add: disjoint_family_on_def binaryset_def Int_commute)
  from Disj [OF ra di]
  show ?thesis
    by (simp add: UN_binaryset_eq)
qed

lemma (in algebra) smallest_ccdi_sets_Int1:
  assumes a: "a  M"
  shows "b  smallest_ccdi_sets Ω M  a  b  smallest_ccdi_sets Ω M"
proof (induct rule: smallest_ccdi_sets.induct)
  case (Basic x)
  thus ?case
    by (metis a Int smallest_ccdi_sets.Basic)
next
  case (Compl x)
  have "a  (Ω - x) = Ω - ((Ω - a)  (a  x))"
    by blast
  also have "...  smallest_ccdi_sets Ω M"
    by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2
           Diff_disjoint Int_Diff Int_empty_right smallest_ccdi_sets_Un
           smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl)
  finally show ?case .
next
  case (Inc A)
  have 1: "(i. (λi. a  A i) i) = a  (i. A i)"
    by blast
  have "range (λi. a  A i)  Pow(smallest_ccdi_sets Ω M)" using Inc
    by blast
  moreover have "(λi. a  A i) 0 = {}"
    by (simp add: Inc)
  moreover have "!!n. (λi. a  A i) n  (λi. a  A i) (Suc n)" using Inc
    by blast
  ultimately have 2: "(i. (λi. a  A i) i)  smallest_ccdi_sets Ω M"
    by (rule smallest_ccdi_sets.Inc)
  show ?case
    by (metis 1 2)
next
  case (Disj A)
  have 1: "(i. (λi. a  A i) i) = a  (i. A i)"
    by blast
  have "range (λi. a  A i)  Pow(smallest_ccdi_sets Ω M)" using Disj
    by blast
  moreover have "disjoint_family (λi. a  A i)" using Disj
    by (auto simp add: disjoint_family_on_def)
  ultimately have 2: "(i. (λi. a  A i) i)  smallest_ccdi_sets Ω M"
    by (rule smallest_ccdi_sets.Disj)
  show ?case
    by (metis 1 2)
qed


lemma (in algebra) smallest_ccdi_sets_Int:
  assumes b: "b  smallest_ccdi_sets Ω M"
  shows "a  smallest_ccdi_sets Ω M  a  b  smallest_ccdi_sets Ω M"
proof (induct rule: smallest_ccdi_sets.induct)
  case (Basic x)
  thus ?case
    by (metis b smallest_ccdi_sets_Int1)
next
  case (Compl x)
  have "(Ω - x)  b = Ω - (x  b  (Ω - b))"
    by blast
  also have "...  smallest_ccdi_sets Ω M"
    by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b
           smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)
  finally show ?case .
next
  case (Inc A)
  have 1: "(i. (λi. A i  b) i) = (i. A i)  b"
    by blast
  have "range (λi. A i  b)  Pow(smallest_ccdi_sets Ω M)" using Inc
    by blast
  moreover have "(λi. A i  b) 0 = {}"
    by (simp add: Inc)
  moreover have "!!n. (λi. A i  b) n  (λi. A i  b) (Suc n)" using Inc
    by blast
  ultimately have 2: "(i. (λi. A i  b) i)  smallest_ccdi_sets Ω M"
    by (rule smallest_ccdi_sets.Inc)
  show ?case
    by (metis 1 2)
next
  case (Disj A)
  have 1: "(i. (λi. A i  b) i) = (i. A i)  b"
    by blast
  have "range (λi. A i  b)  Pow(smallest_ccdi_sets Ω M)" using Disj
    by blast
  moreover have "disjoint_family (λi. A i  b)" using Disj
    by (auto simp add: disjoint_family_on_def)
  ultimately have 2: "(i. (λi. A i  b) i)  smallest_ccdi_sets Ω M"
    by (rule smallest_ccdi_sets.Disj)
  show ?case
    by (metis 1 2)
qed

lemma (in algebra) sigma_property_disjoint_lemma:
  assumes sbC: "M  C"
      and ccdi: "closed_cdi Ω C"
  shows "sigma_sets Ω M  C"
proof -
  have "smallest_ccdi_sets Ω M  {B . M  B  sigma_algebra Ω B}"
    apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int
            smallest_ccdi_sets_Int)
    apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)
    apply (blast intro: smallest_ccdi_sets.Disj)
    done
  hence "sigma_sets (Ω) (M)  smallest_ccdi_sets Ω M"
    by clarsimp
       (drule sigma_algebra.sigma_sets_subset [where a="M"], auto)
  also have "...   C"
    proof
      fix x
      assume x: "x  smallest_ccdi_sets Ω M"
      thus "x  C"
        proof (induct rule: smallest_ccdi_sets.induct)
          case (Basic x)
          thus ?case
            by (metis Basic subsetD sbC)
        next
          case (Compl x)
          thus ?case
            by (blast intro: closed_cdi_Compl [OF ccdi, simplified])
        next
          case (Inc A)
          thus ?case
               by (auto intro: closed_cdi_Inc [OF ccdi, simplified])
        next
          case (Disj A)
          thus ?case
               by (auto intro: closed_cdi_Disj [OF ccdi, simplified])
        qed
    qed
  finally show ?thesis .
qed

lemma (in algebra) sigma_property_disjoint:
  assumes sbC: "M  C"
      and compl: "!!s. s  C  sigma_sets (Ω) (M)  Ω - s  C"
      and inc: "!!A. range A  C  sigma_sets (Ω) (M)
                      A 0 = {}  (!!n. A n  A (Suc n))
                      (i. A i)  C"
      and disj: "!!A. range A  C  sigma_sets (Ω) (M)
                       disjoint_family A  (i::nat. A i)  C"
  shows "sigma_sets (Ω) (M)  C"
proof -
  have "sigma_sets (Ω) (M)  C  sigma_sets (Ω) (M)"
    proof (rule sigma_property_disjoint_lemma)
      show "M  C  sigma_sets (Ω) (M)"
        by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)
    next
      show "closed_cdi Ω (C  sigma_sets (Ω) (M))"
        by (simp add: closed_cdi_def compl inc disj)
           (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed
             IntE sigma_sets.Compl range_subsetD sigma_sets.Union)
    qed
  thus ?thesis
    by blast
qed

subsubsection ‹Dynkin systems›

localetag important› Dynkin_system = subset_class +
  assumes space: "Ω  M"
    and   compl[intro!]: "A. A  M  Ω - A  M"
    and   UN[intro!]: "A. disjoint_family A  range A  M
                            (i::nat. A i)  M"

lemma (in Dynkin_system) empty[intro, simp]: "{}  M"
  using space compl[of "Ω"] by simp

lemma (in Dynkin_system) diff:
  assumes sets: "D  M" "E  M" and "D  E"
  shows "E - D  M"
proof -
  let ?f = "λx. if x = 0 then D else if x = Suc 0 then Ω - E else {}"
  have "range ?f = {D, Ω - E, {}}"
    by (auto simp: image_iff)
  moreover have "D  (Ω - E) = (i. ?f i)"
    by (auto simp: image_iff split: if_split_asm)
  moreover
  have "disjoint_family ?f" unfolding disjoint_family_on_def
    using D  M[THEN sets_into_space] D  E by auto
  ultimately have "Ω - (D  (Ω - E))  M"
    using sets UN by auto fastforce
  also have "Ω - (D  (Ω - E)) = E - D"
    using assms sets_into_space by auto
  finally show ?thesis .
qed

lemma Dynkin_systemI:
  assumes " A. A  M  A  Ω" "Ω  M"
  assumes " A. A  M  Ω - A  M"
  assumes " A. disjoint_family A  range A  M
           (i::nat. A i)  M"
  shows "Dynkin_system Ω M"
  using assms by (auto simp: Dynkin_system_def Dynkin_system_axioms_def subset_class_def)

lemma Dynkin_systemI':
  assumes 1: " A. A  M  A  Ω"
  assumes empty: "{}  M"
  assumes Diff: " A. A  M  Ω - A  M"
  assumes 2: " A. disjoint_family A  range A  M
           (i::nat. A i)  M"
  shows "Dynkin_system Ω M"
proof -
  from Diff[OF empty] have "Ω  M" by auto
  from 1 this Diff 2 show ?thesis
    by (intro Dynkin_systemI) auto
qed

lemma Dynkin_system_trivial:
  shows "Dynkin_system A (Pow A)"
  by (rule Dynkin_systemI) auto

lemma sigma_algebra_imp_Dynkin_system:
  assumes "sigma_algebra Ω M" shows "Dynkin_system Ω M"
proof -
  interpret sigma_algebra Ω M by fact
  show ?thesis using sets_into_space by (fastforce intro!: Dynkin_systemI)
qed

subsubsection "Intersection sets systems"

definitiontag important› Int_stable :: "'a set set  bool" where
"Int_stable M  ( a  M.  b  M. a  b  M)"

lemma (in algebra) Int_stable: "Int_stable M"
  unfolding Int_stable_def by auto

lemma Int_stableI_image:
  "(i j. i  I  j  I  kI. A i  A j = A k)  Int_stable (A ` I)"
  by (auto simp: Int_stable_def image_def)

lemma Int_stableI:
  "(a b. a  A  b  A  a  b  A)  Int_stable A"
  unfolding Int_stable_def by auto

lemma Int_stableD:
  "Int_stable M  a  M  b  M  a  b  M"
  unfolding Int_stable_def by auto

lemma (in Dynkin_system) sigma_algebra_eq_Int_stable:
  "sigma_algebra Ω M  Int_stable M"
proof
  assume "sigma_algebra Ω M" then show "Int_stable M"
    unfolding sigma_algebra_def using algebra.Int_stable by auto
next
  assume "Int_stable M"
  show "sigma_algebra Ω M"
    unfolding sigma_algebra_disjoint_iff algebra_iff_Un
  proof (intro conjI ballI allI impI)
    show "M  Pow (Ω)" using sets_into_space by auto
  next
    fix A B assume "A  M" "B  M"
    then have "A  B = Ω - ((Ω - A)  (Ω - B))"
              "Ω - A  M" "Ω - B  M"
      using sets_into_space by auto
    then show "A  B  M"
      using Int_stable M unfolding Int_stable_def by auto
  qed auto
qed

subsubsection "Smallest Dynkin systems"

definitiontag important› Dynkin :: "'a set  'a set set  'a set set" where
  "Dynkin Ω M =  ({D. Dynkin_system Ω D  M  D})"

lemma Dynkin_system_Dynkin:
  assumes "M  Pow (Ω)"
  shows "Dynkin_system Ω (Dynkin Ω M)"
proof (rule Dynkin_systemI)
  fix A assume "A  Dynkin Ω M"
  moreover
  { fix D assume "A  D" and d: "Dynkin_system Ω D"
    then have "A  Ω" by (auto simp: Dynkin_system_def subset_class_def) }
  moreover have "{D. Dynkin_system Ω D  M  D}  {}"
    using assms Dynkin_system_trivial by fastforce
  ultimately show "A  Ω"
    unfolding Dynkin_def using assms
    by auto
next
  show "Ω  Dynkin Ω M"
    unfolding Dynkin_def using Dynkin_system.space by fastforce
next
  fix A assume "A  Dynkin Ω M"
  then show "Ω - A  Dynkin Ω M"
    unfolding Dynkin_def using Dynkin_system.compl by force
next
  fix A :: "nat  'a set"
  assume A: "disjoint_family A" "range A  Dynkin Ω M"
  show "(i. A i)  Dynkin Ω M" unfolding Dynkin_def
  proof (simp, safe)
    fix D assume "Dynkin_system Ω D" "M  D"
    with A have "(i. A i)  D"
      by (intro Dynkin_system.UN) (auto simp: Dynkin_def)
    then show "(i. A i)  D" by auto
  qed
qed

lemma Dynkin_Basic[intro]: "A  M  A  Dynkin Ω M"
  unfolding Dynkin_def by auto

lemma (in Dynkin_system) restricted_Dynkin_system:
  assumes "D  M"
  shows "Dynkin_system Ω {Q. Q  Ω  Q  D  M}"
proof (rule Dynkin_systemI, simp_all)
  have "Ω  D = D"
    using D  M sets_into_space by auto
  then show "Ω  D  M"
    using D  M by auto
next
  fix A assume "A  Ω  A  D  M"
  moreover have "(Ω - A)  D = (Ω - (A  D)) - (Ω - D)"
    by auto
  ultimately show "(Ω - A)  D  M"
    using  D  M by (auto intro: diff)
next
  fix A :: "nat  'a set"
  assume "disjoint_family A" "range A  {Q. Q  Ω  Q  D  M}"
  then have "i. A i  Ω" "disjoint_family (λi. A i  D)"
    "range (λi. A i  D)  M" "(x. A x)  D = (x. A x  D)"
    by ((fastforce simp: disjoint_family_on_def)+)
  then show "(x. A x)  Ω  (x. A x)  D  M"
    by (auto simp del: UN_simps)
qed

lemma (in Dynkin_system) Dynkin_subset:
  assumes "N  M"
  shows "Dynkin Ω N  M"
proof -
  have "Dynkin_system Ω M" ..
  then have "Dynkin_system Ω M"
    using assms unfolding Dynkin_system_def Dynkin_system_axioms_def subset_class_def by simp
  with N  M show ?thesis by (auto simp add: Dynkin_def)
qed

lemma sigma_eq_Dynkin:
  assumes sets: "M  Pow Ω"
  assumes "Int_stable M"
  shows "sigma_sets Ω M = Dynkin Ω M"
proof -
  have "Dynkin Ω M  sigma_sets (Ω) (M)"
    using sigma_algebra_imp_Dynkin_system
    unfolding Dynkin_def sigma_sets_least_sigma_algebra[OF sets] by auto
  moreover
  interpret Dynkin_system Ω "Dynkin Ω M"
    using Dynkin_system_Dynkin[OF sets] .
  have "sigma_algebra Ω (Dynkin Ω M)"
    unfolding sigma_algebra_eq_Int_stable Int_stable_def
  proof (intro ballI)
    fix A B assume "A  Dynkin Ω M" "B  Dynkin Ω M"
    let ?D = "λE. {Q. Q  Ω  Q  E  Dynkin Ω M}"
    have "M  ?D B"
    proof
      fix E assume "E  M"
      then have "M  ?D E" "E  Dynkin Ω M"
        using sets_into_space Int_stable M by (auto simp: Int_stable_def)
      then have "Dynkin Ω M  ?D E"
        using restricted_Dynkin_system E  Dynkin Ω M
        by (intro Dynkin_system.Dynkin_subset) simp_all
      then have "B  ?D E"
        using B  Dynkin Ω M by auto
      then have "E  B  Dynkin Ω M"
        by (subst Int_commute) simp
      then show "E  ?D B"
        using sets E  M by auto
    qed
    then have "Dynkin Ω M  ?D B"
      using restricted_Dynkin_system B  Dynkin Ω M
      by (intro Dynkin_system.Dynkin_subset) simp_all
    then show "A  B  Dynkin Ω M"
      using A  Dynkin Ω M sets_into_space by auto
  qed
  from sigma_algebra.sigma_sets_subset[OF this, of "M"]
  have "sigma_sets (Ω) (M)  Dynkin Ω M" by auto
  ultimately have "sigma_sets (Ω) (M) = Dynkin Ω M" by auto
  then show ?thesis
    by (auto simp: Dynkin_def)
qed

lemma (in Dynkin_system) Dynkin_idem:
  "Dynkin Ω M = M"
proof -
  have "Dynkin Ω M = M"
  proof
    show "M  Dynkin Ω M"
      using Dynkin_Basic by auto
    show "Dynkin Ω M  M"
      by (intro Dynkin_subset) auto
  qed
  then show ?thesis
    by (auto simp: Dynkin_def)
qed

lemma (in Dynkin_system) Dynkin_lemma:
  assumes "Int_stable E"
  and E: "E  M" "M  sigma_sets Ω E"
  shows "sigma_sets Ω E = M"
proof -
  have "E  Pow Ω"
    using E sets_into_space by force
  then have *: "sigma_sets Ω E = Dynkin Ω E"
    using Int_stable E by (rule sigma_eq_Dynkin)
  then have "Dynkin Ω E = M"
    using assms Dynkin_subset[OF E(1)] by simp
  with * show ?thesis
    using assms by (auto simp: Dynkin_def)
qed

subsubsection ‹Induction rule for intersection-stable generators›

texttag important› ‹The reason to introduce Dynkin-systems is the following induction rules for σ›-algebras
generated by a generator closed under intersection.›

proposition sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]:
  assumes "Int_stable G"
    and closed: "G  Pow Ω"
    and A: "A  sigma_sets Ω G"
  assumes basic: "A. A  G  P A"
    and empty: "P {}"
    and compl: "A. A  sigma_sets Ω G  P A  P (Ω - A)"
    and union: "A. disjoint_family A  range A  sigma_sets Ω G  (i. P (A i))  P (i::nat. A i)"
  shows "P A"
proof -
  let ?D = "{ A  sigma_sets Ω G. P A }"
  interpret sigma_algebra Ω "sigma_sets Ω G"
    using closed by (rule sigma_algebra_sigma_sets)
  from compl[OF _ empty] closed have space: "P Ω" by simp
  interpret Dynkin_system Ω ?D
    by standard (auto dest: sets_into_space intro!: space compl union)
  have "sigma_sets Ω G = ?D"
    by (rule Dynkin_lemma) (auto simp: basic Int_stable G)
  with A show ?thesis by auto
qed

subsection ‹Measure type›

definitiontag important› positive :: "'a set set  ('a set  ennreal)  bool" where
  "positive M μ  μ {} = 0"

definitiontag important› countably_additive :: "'a set set  ('a set  ennreal)  bool" where
"countably_additive M f 
  (A. range A  M  disjoint_family A  (i. A i)  M 
    (i. f (A i)) = f (i. A i))"

definitiontag important› measure_space :: "'a set  'a set set  ('a set  ennreal)  bool" where
"measure_space Ω A μ 
  sigma_algebra Ω A  positive A μ  countably_additive A μ"

typedeftag important› 'a measure =
  "{(Ω::'a set, A, μ). (a-A. μ a = 0)  measure_space Ω A μ }"
proof
  have "sigma_algebra UNIV {{}, UNIV}"
    by (auto simp: sigma_algebra_iff2)
  then show "(UNIV, {{}, UNIV}, λA. 0)  {(Ω, A, μ). (a-A. μ a = 0)  measure_space Ω A μ} "
    by (auto simp: measure_space_def positive_def countably_additive_def)
qed

definitiontag important› space :: "'a measure  'a set" where
  "space M = fst (Rep_measure M)"

definitiontag important› sets :: "'a measure  'a set set" where
  "sets M = fst (snd (Rep_measure M))"

definitiontag important› emeasure :: "'a measure  'a set  ennreal" where
  "emeasure M = snd (snd (Rep_measure M))"

definitiontag important› measure :: "'a measure  'a set  real" where
  "measure M A = enn2real (emeasure M A)"

declare [[coercion sets]]

declare [[coercion measure]]

declare [[coercion emeasure]]

lemma measure_space: "measure_space (space M) (sets M) (emeasure M)"
  by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse)

interpretation sets: sigma_algebra "space M" "sets M" for M :: "'a measure"
  using measure_space[of M] by (auto simp: measure_space_def)

definitiontag important› measure_of :: "'a set  'a set set  ('a set  ennreal)  'a measure"
  where
"measure_of Ω A μ =
  Abs_measure (Ω, if A  Pow Ω then sigma_sets Ω A else {{}, Ω},
    λa. if a  sigma_sets Ω A  measure_space Ω (sigma_sets Ω A) μ then μ a else 0)"

abbreviation "sigma Ω A  measure_of Ω A (λx. 0)"

lemma measure_space_0: "A  Pow Ω  measure_space Ω (sigma_sets Ω A) (λx. 0)"
  unfolding measure_space_def
  by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def)

lemma sigma_algebra_trivial: "sigma_algebra Ω {{}, Ω}"
by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{Ω}"])+

lemma measure_space_0': "measure_space Ω {{}, Ω} (λx. 0)"
by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial)

lemma measure_space_closed:
  assumes "measure_space Ω M μ"
  shows "M  Pow Ω"
proof -
  interpret sigma_algebra Ω M using assms by(simp add: measure_space_def)
  show ?thesis by(rule space_closed)
qed

lemma (in ring_of_sets) positive_cong_eq:
  "(a. a  M  μ' a = μ a)  positive M μ' = positive M μ"
  by (auto simp add: positive_def)

lemma (in sigma_algebra) countably_additive_eq:
  "(a. a  M  μ' a = μ a)  countably_additive M μ' = countably_additive M μ"
  unfolding countably_additive_def
  by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq)

lemma measure_space_eq:
  assumes closed: "A  Pow Ω" and eq: "a. a  sigma_sets Ω A  μ a = μ' a"
  shows "measure_space Ω (sigma_sets Ω A) μ = measure_space Ω (sigma_sets Ω A) μ'"
proof -
  interpret sigma_algebra Ω "sigma_sets Ω A" using closed by (rule sigma_algebra_sigma_sets)
  from positive_cong_eq[OF eq, of "λi. i"] countably_additive_eq[OF eq, of "λi. i"] show ?thesis
    by (auto simp: measure_space_def)
qed

lemma measure_of_eq:
  assumes closed: "A  Pow Ω" and eq: "(a. a  sigma_sets Ω A  μ a = μ' a)"
  shows "measure_of Ω A μ = measure_of Ω A μ'"
proof -
  have "measure_space Ω (sigma_sets Ω A) μ = measure_space Ω