996 with V show ?thesis |
996 with V show ?thesis |
997 unfolding sigma_sets_generated_ring_eq |
997 unfolding sigma_sets_generated_ring_eq |
998 by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic) |
998 by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic) |
999 qed |
999 qed |
1000 |
1000 |
|
1001 lemma extend_measure_caratheodory: |
|
1002 fixes G :: "'i \<Rightarrow> 'a set" |
|
1003 assumes M: "M = extend_measure \<Omega> I G \<mu>" |
|
1004 assumes "i \<in> I" |
|
1005 assumes "semiring_of_sets \<Omega> (G ` I)" |
|
1006 assumes empty: "\<And>i. i \<in> I \<Longrightarrow> G i = {} \<Longrightarrow> \<mu> i = 0" |
|
1007 assumes inj: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> G i = G j \<Longrightarrow> \<mu> i = \<mu> j" |
|
1008 assumes nonneg: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> \<mu> i" |
|
1009 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> |
|
1010 (\<Union>i. G (A i)) = G j \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> j" |
|
1011 shows "emeasure M (G i) = \<mu> i" |
|
1012 proof - |
|
1013 interpret semiring_of_sets \<Omega> "G ` I" |
|
1014 by fact |
|
1015 have "\<forall>g\<in>G`I. \<exists>i\<in>I. g = G i" |
|
1016 by auto |
|
1017 then obtain sel where sel: "\<And>g. g \<in> G ` I \<Longrightarrow> sel g \<in> I" "\<And>g. g \<in> G ` I \<Longrightarrow> G (sel g) = g" |
|
1018 by metis |
|
1019 |
|
1020 have "\<exists>\<mu>'. (\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'" |
|
1021 proof (rule caratheodory) |
|
1022 show "positive (G ` I) (\<lambda>s. \<mu> (sel s))" |
|
1023 by (auto simp: positive_def intro!: empty sel nonneg) |
|
1024 show "countably_additive (G ` I) (\<lambda>s. \<mu> (sel s))" |
|
1025 proof (rule countably_additiveI) |
|
1026 fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> G ` I" "disjoint_family A" "(\<Union>i. A i) \<in> G ` I" |
|
1027 then show "(\<Sum>i. \<mu> (sel (A i))) = \<mu> (sel (\<Union>i. A i))" |
|
1028 by (intro add) (auto simp: sel image_subset_iff_funcset comp_def Pi_iff intro!: sel) |
|
1029 qed |
|
1030 qed |
|
1031 then obtain \<mu>' where \<mu>': "\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)" "measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'" |
|
1032 by metis |
|
1033 |
|
1034 show ?thesis |
|
1035 proof (rule emeasure_extend_measure[OF M]) |
|
1036 { fix i assume "i \<in> I" then show "\<mu>' (G i) = \<mu> i" |
|
1037 using \<mu>' by (auto intro!: inj sel) } |
|
1038 show "G ` I \<subseteq> Pow \<Omega>" |
|
1039 by fact |
|
1040 then show "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'" |
|
1041 using \<mu>' by (simp_all add: M sets_extend_measure measure_space_def) |
|
1042 qed fact |
|
1043 qed |
|
1044 |
|
1045 lemma extend_measure_caratheodory_pair: |
|
1046 fixes G :: "'i \<Rightarrow> 'j \<Rightarrow> 'a set" |
|
1047 assumes M: "M = extend_measure \<Omega> {(a, b). P a b} (\<lambda>(a, b). G a b) (\<lambda>(a, b). \<mu> a b)" |
|
1048 assumes "P i j" |
|
1049 assumes semiring: "semiring_of_sets \<Omega> {G a b | a b. P a b}" |
|
1050 assumes empty: "\<And>i j. P i j \<Longrightarrow> G i j = {} \<Longrightarrow> \<mu> i j = 0" |
|
1051 assumes inj: "\<And>i j k l. P i j \<Longrightarrow> P k l \<Longrightarrow> G i j = G k l \<Longrightarrow> \<mu> i j = \<mu> k l" |
|
1052 assumes nonneg: "\<And>i j. P i j \<Longrightarrow> 0 \<le> \<mu> i j" |
|
1053 assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>B::nat \<Rightarrow> 'j. \<And>j k. |
|
1054 (\<And>n. P (A n) (B n)) \<Longrightarrow> P j k \<Longrightarrow> disjoint_family (\<lambda>n. G (A n) (B n)) \<Longrightarrow> |
|
1055 (\<Union>i. G (A i) (B i)) = G j k \<Longrightarrow> (\<Sum>n. \<mu> (A n) (B n)) = \<mu> j k" |
|
1056 shows "emeasure M (G i j) = \<mu> i j" |
|
1057 proof - |
|
1058 have "emeasure M ((\<lambda>(a, b). G a b) (i, j)) = (\<lambda>(a, b). \<mu> a b) (i, j)" |
|
1059 proof (rule extend_measure_caratheodory[OF M]) |
|
1060 show "semiring_of_sets \<Omega> ((\<lambda>(a, b). G a b) ` {(a, b). P a b})" |
|
1061 using semiring by (simp add: image_def conj_commute) |
|
1062 next |
|
1063 fix A :: "nat \<Rightarrow> ('i \<times> 'j)" and j assume "A \<in> UNIV \<rightarrow> {(a, b). P a b}" "j \<in> {(a, b). P a b}" |
|
1064 "disjoint_family ((\<lambda>(a, b). G a b) \<circ> A)" |
|
1065 "(\<Union>i. case A i of (a, b) \<Rightarrow> G a b) = (case j of (a, b) \<Rightarrow> G a b)" |
|
1066 then show "(\<Sum>n. case A n of (a, b) \<Rightarrow> \<mu> a b) = (case j of (a, b) \<Rightarrow> \<mu> a b)" |
|
1067 using add[of "\<lambda>i. fst (A i)" "\<lambda>i. snd (A i)" "fst j" "snd j"] |
|
1068 by (simp add: split_beta' comp_def Pi_iff) |
|
1069 qed (auto split: prod.splits intro: assms) |
|
1070 then show ?thesis by simp |
|
1071 qed |
|
1072 |
|
1073 |
1001 end |
1074 end |