src/HOL/Analysis/Caratheodory.thy
changeset 68833 fde093888c16
parent 67682 00c436488398
child 69164 74f1b0f10b2b
--- a/src/HOL/Analysis/Caratheodory.thy	Mon Aug 27 22:58:36 2018 +0200
+++ b/src/HOL/Analysis/Caratheodory.thy	Tue Aug 28 13:28:39 2018 +0100
@@ -3,7 +3,7 @@
     Author:     Johannes Hölzl, TU München
 *)
 
-section \<open>Caratheodory Extension Theorem\<close>
+section%important \<open>Caratheodory Extension Theorem\<close>
 
 theory Caratheodory
   imports Measure_Space
@@ -13,7 +13,7 @@
   Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
 \<close>
 
-lemma suminf_ennreal_2dimen:
+lemma%unimportant suminf_ennreal_2dimen:
   fixes f:: "nat \<times> nat \<Rightarrow> ennreal"
   assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
   shows "(\<Sum>i. f (prod_decode i)) = suminf g"
@@ -49,18 +49,18 @@
     by (simp add: suminf_eq_SUP)
 qed
 
-subsection \<open>Characterizations of Measures\<close>
+subsection%important \<open>Characterizations of Measures\<close>
 
-definition outer_measure_space where
+definition%important outer_measure_space where
   "outer_measure_space M f \<longleftrightarrow> positive M f \<and> increasing M f \<and> countably_subadditive M f"
 
-subsubsection \<open>Lambda Systems\<close>
+subsubsection%important \<open>Lambda Systems\<close>
 
-definition lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
+definition%important lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
 where
   "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
 
-lemma (in algebra) lambda_system_eq:
+lemma%unimportant (in algebra) lambda_system_eq:
   "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (x \<inter> l) + f (x - l) = f x}"
 proof -
   have [simp]: "\<And>l x. l \<in> M \<Longrightarrow> x \<in> M \<Longrightarrow> (\<Omega> - l) \<inter> x = x - l"
@@ -69,13 +69,13 @@
     by (auto simp add: lambda_system_def) (metis Int_commute)+
 qed
 
-lemma (in algebra) lambda_system_empty: "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
+lemma%unimportant (in algebra) lambda_system_empty: "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
   by (auto simp add: positive_def lambda_system_eq)
 
-lemma lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
+lemma%unimportant lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
   by (simp add: lambda_system_def)
 
-lemma (in algebra) lambda_system_Compl:
+lemma%unimportant (in algebra) lambda_system_Compl:
   fixes f:: "'a set \<Rightarrow> ennreal"
   assumes x: "x \<in> lambda_system \<Omega> M f"
   shows "\<Omega> - x \<in> lambda_system \<Omega> M f"
@@ -88,7 +88,7 @@
     by (force simp add: lambda_system_def ac_simps)
 qed
 
-lemma (in algebra) lambda_system_Int:
+lemma%unimportant (in algebra) lambda_system_Int:
   fixes f:: "'a set \<Rightarrow> ennreal"
   assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
   shows "x \<inter> y \<in> lambda_system \<Omega> M f"
@@ -122,7 +122,7 @@
   qed
 qed
 
-lemma (in algebra) lambda_system_Un:
+lemma%unimportant (in algebra) lambda_system_Un:
   fixes f:: "'a set \<Rightarrow> ennreal"
   assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
   shows "x \<union> y \<in> lambda_system \<Omega> M f"
@@ -136,7 +136,7 @@
     by (metis lambda_system_Compl lambda_system_Int xl yl)
 qed
 
-lemma (in algebra) lambda_system_algebra:
+lemma%unimportant (in algebra) lambda_system_algebra:
   "positive M f \<Longrightarrow> algebra \<Omega> (lambda_system \<Omega> M f)"
   apply (auto simp add: algebra_iff_Un)
   apply (metis lambda_system_sets set_mp sets_into_space)
@@ -145,7 +145,7 @@
   apply (metis lambda_system_Un)
   done
 
-lemma (in algebra) lambda_system_strong_additive:
+lemma%unimportant (in algebra) lambda_system_strong_additive:
   assumes z: "z \<in> M" and disj: "x \<inter> y = {}"
       and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
   shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)"
@@ -160,7 +160,7 @@
     by (simp add: lambda_system_eq)
 qed
 
-lemma (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f"
+lemma%unimportant (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f"
 proof (auto simp add: additive_def)
   fix x and y
   assume disj: "x \<inter> y = {}"
@@ -171,13 +171,13 @@
     by (simp add: Un)
 qed
 
-lemma lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
+lemma%unimportant lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
   by (simp add: increasing_def lambda_system_def)
 
-lemma lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
+lemma%unimportant lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
   by (simp add: positive_def lambda_system_def)
 
-lemma (in algebra) lambda_system_strong_sum:
+lemma%unimportant (in algebra) lambda_system_strong_sum:
   fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ennreal"
   assumes f: "positive M f" and a: "a \<in> M"
       and A: "range A \<subseteq> lambda_system \<Omega> M f"
@@ -199,12 +199,12 @@
     by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4])
 qed
 
-lemma (in sigma_algebra) lambda_system_caratheodory:
+lemma%important (in sigma_algebra) lambda_system_caratheodory:
   assumes oms: "outer_measure_space M f"
       and A: "range A \<subseteq> lambda_system \<Omega> M f"
       and disj: "disjoint_family A"
   shows  "(\<Union>i. A i) \<in> lambda_system \<Omega> M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
-proof -
+proof%unimportant -
   have pos: "positive M f" and inc: "increasing M f"
    and csa: "countably_subadditive M f"
     by (metis oms outer_measure_space_def)+
@@ -274,11 +274,11 @@
     by (simp add: lambda_system_eq sums_iff U_eq U_in)
 qed
 
-lemma (in sigma_algebra) caratheodory_lemma:
+lemma%important (in sigma_algebra) caratheodory_lemma:
   assumes oms: "outer_measure_space M f"
   defines "L \<equiv> lambda_system \<Omega> M f"
   shows "measure_space \<Omega> L f"
-proof -
+proof%unimportant -
   have pos: "positive M f"
     by (metis oms outer_measure_space_def)
   have alg: "algebra \<Omega> L"
@@ -297,11 +297,11 @@
     using pos by (simp add: measure_space_def)
 qed
 
-definition outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
+definition%important outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
    "outer_measure M f X =
      (INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))"
 
-lemma (in ring_of_sets) outer_measure_agrees:
+lemma%unimportant (in ring_of_sets) outer_measure_agrees:
   assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \<in> M"
   shows "outer_measure M f s = f s"
   unfolding outer_measure_def
@@ -326,19 +326,19 @@
        (auto simp: disjoint_family_on_def)
 qed
 
-lemma outer_measure_empty:
+lemma%unimportant outer_measure_empty:
   "positive M f \<Longrightarrow> {} \<in> M \<Longrightarrow> outer_measure M f {} = 0"
   unfolding outer_measure_def
   by (intro antisym INF_lower2[of  "\<lambda>_. {}"]) (auto simp: disjoint_family_on_def positive_def)
 
-lemma (in ring_of_sets) positive_outer_measure:
+lemma%unimportant (in ring_of_sets) positive_outer_measure:
   assumes "positive M f" shows "positive (Pow \<Omega>) (outer_measure M f)"
   unfolding positive_def by (auto simp: assms outer_measure_empty)
 
-lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
+lemma%unimportant (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
   by (force simp: increasing_def outer_measure_def intro!: INF_greatest intro: INF_lower)
 
-lemma (in ring_of_sets) outer_measure_le:
+lemma%unimportant (in ring_of_sets) outer_measure_le:
   assumes pos: "positive M f" and inc: "increasing M f" and A: "range A \<subseteq> M" and X: "X \<subseteq> (\<Union>i. A i)"
   shows "outer_measure M f X \<le> (\<Sum>i. f (A i))"
   unfolding outer_measure_def
@@ -351,11 +351,11 @@
     by (blast intro!: suminf_le)
 qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed)
 
-lemma (in ring_of_sets) outer_measure_close:
+lemma%unimportant (in ring_of_sets) outer_measure_close:
   "outer_measure M f X < e \<Longrightarrow> \<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) < e"
   unfolding outer_measure_def INF_less_iff by auto
 
-lemma (in ring_of_sets) countably_subadditive_outer_measure:
+lemma%unimportant (in ring_of_sets) countably_subadditive_outer_measure:
   assumes posf: "positive M f" and inc: "increasing M f"
   shows "countably_subadditive (Pow \<Omega>) (outer_measure M f)"
 proof (simp add: countably_subadditive_def, safe)
@@ -398,12 +398,12 @@
   qed
 qed
 
-lemma (in ring_of_sets) outer_measure_space_outer_measure:
+lemma%unimportant (in ring_of_sets) outer_measure_space_outer_measure:
   "positive M f \<Longrightarrow> increasing M f \<Longrightarrow> outer_measure_space (Pow \<Omega>) (outer_measure M f)"
   by (simp add: outer_measure_space_def
     positive_outer_measure increasing_outer_measure countably_subadditive_outer_measure)
 
-lemma (in ring_of_sets) algebra_subset_lambda_system:
+lemma%unimportant (in ring_of_sets) algebra_subset_lambda_system:
   assumes posf: "positive M f" and inc: "increasing M f"
       and add: "additive M f"
   shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (outer_measure M f)"
@@ -457,15 +457,15 @@
     by (rule order_antisym)
 qed
 
-lemma measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
+lemma%unimportant measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
   by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq)
 
-subsection \<open>Caratheodory's theorem\<close>
+subsection%important \<open>Caratheodory's theorem\<close>
 
-theorem (in ring_of_sets) caratheodory':
+theorem%important (in ring_of_sets) caratheodory':
   assumes posf: "positive M f" and ca: "countably_additive M f"
   shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
-proof -
+proof%unimportant -
   have inc: "increasing M f"
     by (metis additive_increasing ca countably_additive_additive posf)
   let ?O = "outer_measure M f"
@@ -489,37 +489,37 @@
     by (intro exI[of _ ?O]) auto
 qed
 
-lemma (in ring_of_sets) caratheodory_empty_continuous:
+lemma%important (in ring_of_sets) caratheodory_empty_continuous:
   assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>"
   assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
   shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
-proof (intro caratheodory' empty_continuous_imp_countably_additive f)
+proof%unimportant (intro caratheodory' empty_continuous_imp_countably_additive f)
   show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
 qed (rule cont)
 
-subsection \<open>Volumes\<close>
+subsection%important \<open>Volumes\<close>
 
-definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+definition%important volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
   "volume M f \<longleftrightarrow>
   (f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and>
   (\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))"
 
-lemma volumeI:
+lemma%unimportant volumeI:
   assumes "f {} = 0"
   assumes "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> f a"
   assumes "\<And>C. C \<subseteq> M \<Longrightarrow> disjoint C \<Longrightarrow> finite C \<Longrightarrow> \<Union>C \<in> M \<Longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c)"
   shows "volume M f"
   using assms by (auto simp: volume_def)
 
-lemma volume_positive:
+lemma%unimportant volume_positive:
   "volume M f \<Longrightarrow> a \<in> M \<Longrightarrow> 0 \<le> f a"
   by (auto simp: volume_def)
 
-lemma volume_empty:
+lemma%unimportant volume_empty:
   "volume M f \<Longrightarrow> f {} = 0"
   by (auto simp: volume_def)
 
-lemma volume_finite_additive:
+lemma%unimportant volume_finite_additive:
   assumes "volume M f"
   assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "UNION I A \<in> M"
   shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
@@ -540,7 +540,7 @@
     by simp
 qed
 
-lemma (in ring_of_sets) volume_additiveI:
+lemma%unimportant (in ring_of_sets) volume_additiveI:
   assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a"
   assumes [simp]: "\<mu> {} = 0"
   assumes add: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> \<mu> (a \<union> b) = \<mu> a + \<mu> b"
@@ -557,10 +557,10 @@
   qed simp
 qed fact+
 
-lemma (in semiring_of_sets) extend_volume:
+lemma%important (in semiring_of_sets) extend_volume:
   assumes "volume M \<mu>"
   shows "\<exists>\<mu>'. volume generated_ring \<mu>' \<and> (\<forall>a\<in>M. \<mu>' a = \<mu> a)"
-proof -
+proof%unimportant -
   let ?R = generated_ring
   have "\<forall>a\<in>?R. \<exists>m. \<exists>C\<subseteq>M. a = \<Union>C \<and> finite C \<and> disjoint C \<and> m = (\<Sum>c\<in>C. \<mu> c)"
     by (auto simp: generated_ring_def)
@@ -635,12 +635,12 @@
   qed
 qed
 
-subsubsection \<open>Caratheodory on semirings\<close>
+subsubsection%important \<open>Caratheodory on semirings\<close>
 
-theorem (in semiring_of_sets) caratheodory:
+theorem%important (in semiring_of_sets) caratheodory:
   assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
   shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'"
-proof -
+proof%unimportant -
   have "volume M \<mu>"
   proof (rule volumeI)
     { fix a assume "a \<in> M" then show "0 \<le> \<mu> a"
@@ -816,7 +816,7 @@
     by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic)
 qed
 
-lemma extend_measure_caratheodory:
+lemma%important extend_measure_caratheodory:
   fixes G :: "'i \<Rightarrow> 'a set"
   assumes M: "M = extend_measure \<Omega> I G \<mu>"
   assumes "i \<in> I"
@@ -827,7 +827,8 @@
   assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>j. A \<in> UNIV \<rightarrow> I \<Longrightarrow> j \<in> I \<Longrightarrow> disjoint_family (G \<circ> A) \<Longrightarrow>
     (\<Union>i. G (A i)) = G j \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> j"
   shows "emeasure M (G i) = \<mu> i"
-proof -
+
+proof%unimportant -
   interpret semiring_of_sets \<Omega> "G ` I"
     by fact
   have "\<forall>g\<in>G`I. \<exists>i\<in>I. g = G i"
@@ -860,7 +861,7 @@
   qed fact
 qed
 
-lemma extend_measure_caratheodory_pair:
+lemma%important extend_measure_caratheodory_pair:
   fixes G :: "'i \<Rightarrow> 'j \<Rightarrow> 'a set"
   assumes M: "M = extend_measure \<Omega> {(a, b). P a b} (\<lambda>(a, b). G a b) (\<lambda>(a, b). \<mu> a b)"
   assumes "P i j"
@@ -872,7 +873,7 @@
     (\<And>n. P (A n) (B n)) \<Longrightarrow> P j k \<Longrightarrow> disjoint_family (\<lambda>n. G (A n) (B n)) \<Longrightarrow>
     (\<Union>i. G (A i) (B i)) = G j k \<Longrightarrow> (\<Sum>n. \<mu> (A n) (B n)) = \<mu> j k"
   shows "emeasure M (G i j) = \<mu> i j"
-proof -
+proof%unimportant -
   have "emeasure M ((\<lambda>(a, b). G a b) (i, j)) = (\<lambda>(a, b). \<mu> a b) (i, j)"
   proof (rule extend_measure_caratheodory[OF M])
     show "semiring_of_sets \<Omega> ((\<lambda>(a, b). G a b) ` {(a, b). P a b})"