diff -r 9e71155e3666 -r 0135a0c77b64 src/HOL/Analysis/Finite_Product_Measure.thy --- a/src/HOL/Analysis/Finite_Product_Measure.thy Fri Sep 24 13:40:14 2021 +0200 +++ b/src/HOL/Analysis/Finite_Product_Measure.thy Fri Sep 24 22:23:26 2021 +0200 @@ -269,11 +269,24 @@ qed lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)" -proof (unfold Int_stable_def, safe) + unfolding Int_stable_def +proof safe fix A assume "A \ prod_algebra I M" - from prod_algebraE[OF this] guess J E . note A = this + from prod_algebraE[OF this] obtain J E where A: + "A = prod_emb I M J (Pi\<^sub>E J E)" + "finite J" + "J \ {} \ I = {}" + "J \ I" + "\i. i \ J \ E i \ sets (M i)" + by auto fix B assume "B \ prod_algebra I M" - from prod_algebraE[OF this] guess K F . note B = this + from prod_algebraE[OF this] obtain K F where B: + "B = prod_emb I M K (Pi\<^sub>E K F)" + "finite K" + "K \ {} \ I = {}" + "K \ I" + "\i. i \ K \ F i \ sets (M i)" + by auto have "A \ B = prod_emb I M (J \ K) (\\<^sub>E i\J \ K. (if i \ J then E i else space (M i)) \ (if i \ K then F i else space (M i)))" unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4) @@ -360,7 +373,13 @@ proof (rule sigma_sets_eqI) interpret R: sigma_algebra ?\ "sigma_sets ?\ ?R" by (rule sigma_algebra_sigma_sets) auto fix A assume "A \ prod_algebra I M" - from prod_algebraE[OF this] guess J X . note X = this + from prod_algebraE[OF this] obtain J X where X: + "A = prod_emb I M J (Pi\<^sub>E J X)" + "finite J" + "J \ {} \ I = {}" + "J \ I" + "\i. i \ J \ X i \ sets (M i)" + by auto show "A \ sigma_sets ?\ ?R" proof cases assume "I = {}" @@ -525,7 +544,13 @@ using sets_PiM prod_algebra_sets_into_space space proof (rule measurable_sigma_sets) fix A assume "A \ prod_algebra I M" - from prod_algebraE[OF this] guess J X . + from prod_algebraE[OF this] obtain J X where + "A = prod_emb I M J (Pi\<^sub>E J X)" + "finite J" + "J \ {} \ I = {}" + "J \ I" + "\i. i \ J \ X i \ sets (M i)" + by auto with sets[of J X] show "f -` A \ space N \ sets N" by auto qed @@ -537,7 +562,14 @@ using sets_PiM prod_algebra_sets_into_space space proof (rule measurable_sigma_sets) fix A assume "A \ prod_algebra I M" - from prod_algebraE[OF this] guess J X . note X = this + from prod_algebraE[OF this] obtain J X + where X: + "A = prod_emb I M J (Pi\<^sub>E J X)" + "finite J" + "J \ {} \ I = {}" + "J \ I" + "\i. i \ J \ X i \ sets (M i)" + by auto then have "f -` A \ space N = {\ \ space N. \i\J. f \ i \ X i}" using space by (auto simp: prod_emb_def del: PiE_I) also have "\ \ sets N" using X(3,2,4,5) by (rule sets) @@ -844,7 +876,9 @@ proof - have "\i::'i. \F::nat \ 'a set. range F \ sets (M i) \ incseq F \ (\i. F i) = space (M i) \ (\k. emeasure (M i) (F k) \ \)" using M.sigma_finite_incseq by metis - from choice[OF this] guess F :: "'i \ nat \ 'a set" .. + then obtain F :: "'i \ nat \ 'a set" + where "\x. range (F x) \ sets (M x) \ incseq (F x) \ \ (range (F x)) = space (M x) \ (\k. emeasure (M x) (F x k) \ \)" + by metis then have F: "\i. range (F i) \ sets (M i)" "\i. incseq (F i)" "\i. (\j. F i j) = space (M i)" "\i k. emeasure (M i) (F i k) \ \" by auto let ?F = "\k. \\<^sub>E i\I. F i k" @@ -950,8 +984,11 @@ shows "P = PiM I M" proof - interpret finite_product_sigma_finite M I - proof qed fact - from sigma_finite_pairs guess C .. note C = this + by standard fact + from sigma_finite_pairs obtain C where C: + "\i\I. range (C i) \ sets (M i)" "\k. \i\I. emeasure (M i) (C i k) \ \" + "incseq (\k. \\<^sub>E i\I. C i k)" "(\k. \\<^sub>E i\I. C i k) = space (Pi\<^sub>M I M)" + by auto show ?thesis proof (rule measure_eqI_PiM_finite[OF I refl P, symmetric]) show "(\i. i \ I \ A i \ sets (M i)) \ (Pi\<^sub>M I M) (Pi\<^sub>E I A) = P (Pi\<^sub>E I A)" for A