diff --git a/src/HOL/Analysis/Abstract_Topology_2.thy b/src/HOL/Analysis/Abstract_Topology_2.thy --- a/src/HOL/Analysis/Abstract_Topology_2.thy +++ b/src/HOL/Analysis/Abstract_Topology_2.thy @@ -1,1601 +1,1599 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) section \Abstract Topology 2\ theory Abstract_Topology_2 imports Elementary_Topology Abstract_Topology "HOL-Library.Indicator_Function" begin text \Combination of Elementary and Abstract Topology\ lemma approachable_lt_le2: "(\(d::real) > 0. \x. Q x \ f x < d \ P x) \ (\d>0. \x. f x \ d \ Q x \ P x)" apply auto apply (rule_tac x="d/2" in exI, auto) done lemma triangle_lemma: fixes x y z :: real assumes x: "0 \ x" and y: "0 \ y" and z: "0 \ z" and xy: "x\<^sup>2 \ y\<^sup>2 + z\<^sup>2" shows "x \ y + z" proof - have "y\<^sup>2 + z\<^sup>2 \ y\<^sup>2 + 2 * y * z + z\<^sup>2" using z y by simp with xy have th: "x\<^sup>2 \ (y + z)\<^sup>2" by (simp add: power2_eq_square field_simps) from y z have yz: "y + z \ 0" by arith from power2_le_imp_le[OF th yz] show ?thesis . qed lemma isCont_indicator: fixes x :: "'a::t2_space" shows "isCont (indicator A :: 'a \ real) x = (x \ frontier A)" proof auto fix x assume cts_at: "isCont (indicator A :: 'a \ real) x" and fr: "x \ frontier A" with continuous_at_open have 1: "\V::real set. open V \ indicator A x \ V \ (\U::'a set. open U \ x \ U \ (\y\U. indicator A y \ V))" by auto show False proof (cases "x \ A") assume x: "x \ A" hence "indicator A x \ ({0<..<2} :: real set)" by simp - hence "\U. open U \ x \ U \ (\y\U. indicator A y \ ({0<..<2} :: real set))" - using 1 open_greaterThanLessThan by blast - then guess U .. note U = this + with 1 obtain U where U: "open U" "x \ U" "\y\U. indicator A y \ ({0<..<2} :: real set)" + using open_greaterThanLessThan by metis hence "\y\U. indicator A y > (0::real)" unfolding greaterThanLessThan_def by auto hence "U \ A" using indicator_eq_0_iff by force hence "x \ interior A" using U interiorI by auto thus ?thesis using fr unfolding frontier_def by simp next assume x: "x \ A" hence "indicator A x \ ({-1<..<1} :: real set)" by simp - hence "\U. open U \ x \ U \ (\y\U. indicator A y \ ({-1<..<1} :: real set))" - using 1 open_greaterThanLessThan by blast - then guess U .. note U = this + with 1 obtain U where U: "open U" "x \ U" "\y\U. indicator A y \ ({-1<..<1} :: real set)" + using 1 open_greaterThanLessThan by metis hence "\y\U. indicator A y < (1::real)" unfolding greaterThanLessThan_def by auto hence "U \ -A" by auto hence "x \ interior (-A)" using U interiorI by auto thus ?thesis using fr interior_complement unfolding frontier_def by auto qed next assume nfr: "x \ frontier A" hence "x \ interior A \ x \ interior (-A)" by (auto simp: frontier_def closure_interior) thus "isCont ((indicator A)::'a \ real) x" proof assume int: "x \ interior A" then obtain U where U: "open U" "x \ U" "U \ A" unfolding interior_def by auto hence "\y\U. indicator A y = (1::real)" unfolding indicator_def by auto hence "continuous_on U (indicator A)" by (simp add: indicator_eq_1_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto next assume ext: "x \ interior (-A)" then obtain U where U: "open U" "x \ U" "U \ -A" unfolding interior_def by auto then have "continuous_on U (indicator A)" using continuous_on_topological by (auto simp: subset_iff) thus ?thesis using U continuous_on_eq_continuous_at by auto qed qed lemma closedin_limpt: "closedin (top_of_set T) S \ S \ T \ (\x. x islimpt S \ x \ T \ x \ S)" apply (simp add: closedin_closed, safe) apply (simp add: closed_limpt islimpt_subset) apply (rule_tac x="closure S" in exI, simp) apply (force simp: closure_def) done lemma closedin_closed_eq: "closed S \ closedin (top_of_set S) T \ closed T \ T \ S" by (meson closedin_limpt closed_subset closedin_closed_trans) lemma connected_closed_set: "closed S \ connected S \ (\A B. closed A \ closed B \ A \ {} \ B \ {} \ A \ B = S \ A \ B = {})" unfolding connected_closedin_eq closedin_closed_eq connected_closedin_eq by blast text \If a connnected set is written as the union of two nonempty closed sets, then these sets have to intersect.\ lemma connected_as_closed_union: assumes "connected C" "C = A \ B" "closed A" "closed B" "A \ {}" "B \ {}" shows "A \ B \ {}" by (metis assms closed_Un connected_closed_set) lemma closedin_subset_trans: "closedin (top_of_set U) S \ S \ T \ T \ U \ closedin (top_of_set T) S" by (meson closedin_limpt subset_iff) lemma openin_subset_trans: "openin (top_of_set U) S \ S \ T \ T \ U \ openin (top_of_set T) S" by (auto simp: openin_open) lemma closedin_compact: "\compact S; closedin (top_of_set S) T\ \ compact T" by (metis closedin_closed compact_Int_closed) lemma closedin_compact_eq: fixes S :: "'a::t2_space set" shows "compact S \ (closedin (top_of_set S) T \ compact T \ T \ S)" by (metis closedin_imp_subset closedin_compact closed_subset compact_imp_closed) subsection \Closure\ lemma euclidean_closure_of [simp]: "euclidean closure_of S = closure S" by (auto simp: closure_of_def closure_def islimpt_def) lemma closure_openin_Int_closure: assumes ope: "openin (top_of_set U) S" and "T \ U" shows "closure(S \ closure T) = closure(S \ T)" proof obtain V where "open V" and S: "S = U \ V" using ope using openin_open by metis show "closure (S \ closure T) \ closure (S \ T)" proof (clarsimp simp: S) fix x assume "x \ closure (U \ V \ closure T)" then have "V \ closure T \ A \ x \ closure A" for A by (metis closure_mono subsetD inf.coboundedI2 inf_assoc) then have "x \ closure (T \ V)" by (metis \open V\ closure_closure inf_commute open_Int_closure_subset) then show "x \ closure (U \ V \ T)" by (metis \T \ U\ inf.absorb_iff2 inf_assoc inf_commute) qed next show "closure (S \ T) \ closure (S \ closure T)" by (meson Int_mono closure_mono closure_subset order_refl) qed corollary infinite_openin: fixes S :: "'a :: t1_space set" shows "\openin (top_of_set U) S; x \ S; x islimpt U\ \ infinite S" by (clarsimp simp add: openin_open islimpt_eq_acc_point inf_commute) lemma closure_Int_ballI: assumes "\U. \openin (top_of_set S) U; U \ {}\ \ T \ U \ {}" shows "S \ closure T" proof (clarsimp simp: closure_iff_nhds_not_empty) fix x and A and V assume "x \ S" "V \ A" "open V" "x \ V" "T \ A = {}" then have "openin (top_of_set S) (A \ V \ S)" by (auto simp: openin_open intro!: exI[where x="V"]) moreover have "A \ V \ S \ {}" using \x \ V\ \V \ A\ \x \ S\ by auto ultimately have "T \ (A \ V \ S) \ {}" by (rule assms) with \T \ A = {}\ show False by auto qed subsection \Frontier\ lemma euclidean_interior_of [simp]: "euclidean interior_of S = interior S" by (auto simp: interior_of_def interior_def) lemma euclidean_frontier_of [simp]: "euclidean frontier_of S = frontier S" by (auto simp: frontier_of_def frontier_def) lemma connected_Int_frontier: "\connected s; s \ t \ {}; s - t \ {}\ \ (s \ frontier t \ {})" apply (simp add: frontier_interiors connected_openin, safe) apply (drule_tac x="s \ interior t" in spec, safe) apply (drule_tac [2] x="s \ interior (-t)" in spec) apply (auto simp: disjoint_eq_subset_Compl dest: interior_subset [THEN subsetD]) done subsection \Compactness\ lemma openin_delete: fixes a :: "'a :: t1_space" shows "openin (top_of_set u) s \ openin (top_of_set u) (s - {a})" by (metis Int_Diff open_delete openin_open) lemma compact_eq_openin_cover: "compact S \ (\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D))" proof safe fix C assume "compact S" and "\c\C. openin (top_of_set S) c" and "S \ \C" then have "\c\{T. open T \ S \ T \ C}. open c" and "S \ \{T. open T \ S \ T \ C}" unfolding openin_open by force+ with \compact S\ obtain D where "D \ {T. open T \ S \ T \ C}" and "finite D" and "S \ \D" by (meson compactE) then have "image (\T. S \ T) D \ C \ finite (image (\T. S \ T) D) \ S \ \(image (\T. S \ T) D)" by auto then show "\D\C. finite D \ S \ \D" .. next assume 1: "\C. (\c\C. openin (top_of_set S) c) \ S \ \C \ (\D\C. finite D \ S \ \D)" show "compact S" proof (rule compactI) fix C let ?C = "image (\T. S \ T) C" assume "\t\C. open t" and "S \ \C" then have "(\c\?C. openin (top_of_set S) c) \ S \ \?C" unfolding openin_open by auto with 1 obtain D where "D \ ?C" and "finite D" and "S \ \D" by metis let ?D = "inv_into C (\T. S \ T) ` D" have "?D \ C \ finite ?D \ S \ \?D" proof (intro conjI) from \D \ ?C\ show "?D \ C" by (fast intro: inv_into_into) from \finite D\ show "finite ?D" by (rule finite_imageI) from \S \ \D\ show "S \ \?D" apply (rule subset_trans) by (metis Int_Union Int_lower2 \D \ (\) S ` C\ image_inv_into_cancel) qed then show "\D\C. finite D \ S \ \D" .. qed qed subsection \Continuity\ lemma interior_image_subset: assumes "inj f" "\x. continuous (at x) f" shows "interior (f ` S) \ f ` (interior S)" proof fix x assume "x \ interior (f ` S)" then obtain T where as: "open T" "x \ T" "T \ f ` S" .. then have "x \ f ` S" by auto then obtain y where y: "y \ S" "x = f y" by auto have "open (f -` T)" using assms \open T\ by (simp add: continuous_at_imp_continuous_on open_vimage) moreover have "y \ vimage f T" using \x = f y\ \x \ T\ by simp moreover have "vimage f T \ S" using \T \ image f S\ \inj f\ unfolding inj_on_def subset_eq by auto ultimately have "y \ interior S" .. with \x = f y\ show "x \ f ` interior S" .. qed subsection\<^marker>\tag unimportant\ \Equality of continuous functions on closure and related results\ lemma continuous_closedin_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closedin (top_of_set S) {x \ S. f x = a}" using continuous_closedin_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_closed_preimage_constant: fixes f :: "_ \ 'b::t1_space" shows "continuous_on S f \ closed S \ closed {x \ S. f x = a}" using continuous_closed_preimage[of S f "{a}"] by (simp add: vimage_def Collect_conj_eq) lemma continuous_constant_on_closure: fixes f :: "_ \ 'b::t1_space" assumes "continuous_on (closure S) f" and "\x. x \ S \ f x = a" and "x \ closure S" shows "f x = a" using continuous_closed_preimage_constant[of "closure S" f a] assms closure_minimal[of S "{x \ closure S. f x = a}"] closure_subset unfolding subset_eq by auto lemma image_closure_subset: assumes contf: "continuous_on (closure S) f" and "closed T" and "(f ` S) \ T" shows "f ` (closure S) \ T" proof - have "S \ {x \ closure S. f x \ T}" using assms(3) closure_subset by auto moreover have "closed (closure S \ f -` T)" using continuous_closed_preimage[OF contf] \closed T\ by auto ultimately have "closure S = (closure S \ f -` T)" using closure_minimal[of S "(closure S \ f -` T)"] by auto then show ?thesis by auto qed subsection\<^marker>\tag unimportant\ \A function constant on a set\ definition constant_on (infixl "(constant'_on)" 50) where "f constant_on A \ \y. \x\A. f x = y" lemma constant_on_subset: "\f constant_on A; B \ A\ \ f constant_on B" unfolding constant_on_def by blast lemma injective_not_constant: fixes S :: "'a::{perfect_space} set" shows "\open S; inj_on f S; f constant_on S\ \ S = {}" unfolding constant_on_def by (metis equals0I inj_on_contraD islimpt_UNIV islimpt_def) lemma constant_on_closureI: fixes f :: "_ \ 'b::t1_space" assumes cof: "f constant_on S" and contf: "continuous_on (closure S) f" shows "f constant_on (closure S)" using continuous_constant_on_closure [OF contf] cof unfolding constant_on_def by metis subsection\<^marker>\tag unimportant\ \Continuity relative to a union.\ lemma continuous_on_Un_local: "\closedin (top_of_set (s \ t)) s; closedin (top_of_set (s \ t)) t; continuous_on s f; continuous_on t f\ \ continuous_on (s \ t) f" unfolding continuous_on closedin_limpt by (metis Lim_trivial_limit Lim_within_union Un_iff trivial_limit_within) lemma continuous_on_cases_local: "\closedin (top_of_set (s \ t)) s; closedin (top_of_set (s \ t)) t; continuous_on s f; continuous_on t g; \x. \x \ s \ \P x \ x \ t \ P x\ \ f x = g x\ \ continuous_on (s \ t) (\x. if P x then f x else g x)" by (rule continuous_on_Un_local) (auto intro: continuous_on_eq) lemma continuous_on_cases_le: fixes h :: "'a :: topological_space \ real" assumes "continuous_on {t \ s. h t \ a} f" and "continuous_on {t \ s. a \ h t} g" and h: "continuous_on s h" and "\t. \t \ s; h t = a\ \ f t = g t" shows "continuous_on s (\t. if h t \ a then f(t) else g(t))" proof - have s: "s = (s \ h -` atMost a) \ (s \ h -` atLeast a)" by force have 1: "closedin (top_of_set s) (s \ h -` atMost a)" by (rule continuous_closedin_preimage [OF h closed_atMost]) have 2: "closedin (top_of_set s) (s \ h -` atLeast a)" by (rule continuous_closedin_preimage [OF h closed_atLeast]) have eq: "s \ h -` {..a} = {t \ s. h t \ a}" "s \ h -` {a..} = {t \ s. a \ h t}" by auto show ?thesis apply (rule continuous_on_subset [of s, OF _ order_refl]) apply (subst s) apply (rule continuous_on_cases_local) using 1 2 s assms apply (auto simp: eq) done qed lemma continuous_on_cases_1: fixes s :: "real set" assumes "continuous_on {t \ s. t \ a} f" and "continuous_on {t \ s. a \ t} g" and "a \ s \ f a = g a" shows "continuous_on s (\t. if t \ a then f(t) else g(t))" using assms by (auto intro: continuous_on_cases_le [where h = id, simplified]) subsection\<^marker>\tag unimportant\\Inverse function property for open/closed maps\ lemma continuous_on_inverse_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g (f x) = x" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_open [of T g] gTS) (metis openin_imp_subset fU oo) qed lemma continuous_on_inverse_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "\x. x \ S \ g(f x) = x" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" shows "continuous_on T g" proof - from imf injf have gTS: "g ` T = S" by force from imf injf have fU: "U \ S \ (f ` U) = T \ g -` U" for U by force show ?thesis by (simp add: continuous_on_closed [of T g] gTS) (metis closedin_imp_subset fU oo) qed lemma homeomorphism_injective_open_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_open_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_injective_closed_map: assumes contf: "continuous_on S f" and imf: "f ` S = T" and injf: "inj_on f S" and oo: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" obtains g where "homeomorphism S T f g" proof have "continuous_on T (inv_into S f)" by (metis contf continuous_on_inverse_closed_map imf injf inv_into_f_f oo) with imf injf contf show "homeomorphism S T f (inv_into S f)" by (auto simp: homeomorphism_def) qed lemma homeomorphism_imp_open_map: assumes hom: "homeomorphism S T f g" and oo: "openin (top_of_set S) U" shows "openin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using openin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_open oo) qed lemma homeomorphism_imp_closed_map: assumes hom: "homeomorphism S T f g" and oo: "closedin (top_of_set S) U" shows "closedin (top_of_set T) (f ` U)" proof - from hom oo have [simp]: "f ` U = T \ g -` U" using closedin_subset by (fastforce simp: homeomorphism_def rev_image_eqI) from hom have "continuous_on T g" unfolding homeomorphism_def by blast moreover have "g ` T = S" by (metis hom homeomorphism_def) ultimately show ?thesis by (simp add: continuous_on_closed oo) qed subsection\<^marker>\tag unimportant\ \Seperability\ lemma subset_second_countable: obtains \ :: "'a:: second_countable_topology set set" where "countable \" "{} \ \" "\C. C \ \ \ openin(top_of_set S) C" "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and opeB: "\C. C \ \ \ openin(top_of_set S) C" and \: "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" proof - obtain \ :: "'a set set" where "countable \" and ope: "\C. C \ \ \ open C" and \: "\S. open S \ \U. U \ \ \ S = \U" by (metis univ_second_countable that) show ?thesis proof show "countable ((\C. S \ C) ` \)" by (simp add: \countable \\) show "\C. C \ (\) S ` \ \ openin (top_of_set S) C" using ope by auto show "\T. openin (top_of_set S) T \ \\\(\) S ` \. T = \\" by (metis \ image_mono inf_Sup openin_open) qed qed show ?thesis proof show "countable (\ - {{}})" using \countable \\ by blast show "\C. \C \ \ - {{}}\ \ openin (top_of_set S) C" by (simp add: \\C. C \ \ \ openin (top_of_set S) C\) show "\\\\ - {{}}. T = \\" if "openin (top_of_set S) T" for T using \ [OF that] apply clarify apply (rule_tac x="\ - {{}}" in exI, auto) done qed auto qed lemma Lindelof_openin: fixes \ :: "'a::second_countable_topology set set" assumes "\S. S \ \ \ openin (top_of_set U) S" obtains \' where "\' \ \" "countable \'" "\\' = \\" proof - have "\S. S \ \ \ \T. open T \ S = U \ T" using assms by (simp add: openin_open) then obtain tf where tf: "\S. S \ \ \ open (tf S) \ (S = U \ tf S)" by metis have [simp]: "\\'. \' \ \ \ \\' = U \ \(tf ` \')" using tf by fastforce obtain \ where "countable \ \ \ \ tf ` \" "\\ = \(tf ` \)" using tf by (force intro: Lindelof [of "tf ` \"]) then obtain \' where \': "\' \ \" "countable \'" "\\' = \\" by (clarsimp simp add: countable_subset_image) then show ?thesis .. qed subsection\<^marker>\tag unimportant\\Closed Maps\ lemma continuous_imp_closed_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "closedin (top_of_set S) U" "continuous_on S f" "f ` S = T" "compact S" shows "closedin (top_of_set T) (f ` U)" by (metis assms closedin_compact_eq compact_continuous_image continuous_on_subset subset_image_iff) lemma closed_map_restrict: assumes cloU: "closedin (top_of_set (S \ f -` T')) U" and cc: "\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)" and "T' \ T" shows "closedin (top_of_set T') (f ` U)" proof - obtain V where "closed V" "U = S \ f -` T' \ V" using cloU by (auto simp: closedin_closed) with cc [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: closedin_closed) qed subsection\<^marker>\tag unimportant\\Open Maps\ lemma open_map_restrict: assumes opeU: "openin (top_of_set (S \ f -` T')) U" and oo: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" and "T' \ T" shows "openin (top_of_set T') (f ` U)" proof - obtain V where "open V" "U = S \ f -` T' \ V" using opeU by (auto simp: openin_open) with oo [of "S \ V"] \T' \ T\ show ?thesis by (fastforce simp add: openin_open) qed subsection\<^marker>\tag unimportant\\Quotient maps\ lemma quotient_map_imp_continuous_open: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis by (meson T continuous_on_open_gen ope openin_imp_subset) qed lemma quotient_map_imp_continuous_closed: assumes T: "f ` S \ T" and ope: "\U. U \ T \ (closedin (top_of_set S) (S \ f -` U) \ closedin (top_of_set T) U)" shows "continuous_on S f" proof - have [simp]: "S \ f -` f ` S = S" by auto show ?thesis by (meson T closedin_imp_subset continuous_on_closed_gen ope) qed lemma open_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. openin (top_of_set S) T \ openin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) = openin (top_of_set (f ` S)) T" proof - have "T = f ` (S \ f -` T)" using T by blast then show ?thesis using "ope" contf continuous_on_open by metis qed lemma closed_map_imp_quotient_map: assumes contf: "continuous_on S f" and T: "T \ f ` S" and ope: "\T. closedin (top_of_set S) T \ closedin (top_of_set (f ` S)) (f ` T)" shows "openin (top_of_set S) (S \ f -` T) \ openin (top_of_set (f ` S)) T" (is "?lhs = ?rhs") proof assume ?lhs then have *: "closedin (top_of_set S) (S - (S \ f -` T))" using closedin_diff by fastforce have [simp]: "(f ` S - f ` (S - (S \ f -` T))) = T" using T by blast show ?rhs using ope [OF *, unfolded closedin_def] by auto next assume ?rhs with contf show ?lhs by (auto simp: continuous_on_open) qed lemma continuous_right_inverse_imp_quotient_map: assumes contf: "continuous_on S f" and imf: "f ` S \ T" and contg: "continuous_on T g" and img: "g ` T \ S" and fg [simp]: "\y. y \ T \ f(g y) = y" and U: "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" (is "?lhs = ?rhs") proof - have f: "\Z. openin (top_of_set (f ` S)) Z \ openin (top_of_set S) (S \ f -` Z)" and g: "\Z. openin (top_of_set (g ` T)) Z \ openin (top_of_set T) (T \ g -` Z)" using contf contg by (auto simp: continuous_on_open) show ?thesis proof have "T \ g -` (g ` T \ (S \ f -` U)) = {x \ T. f (g x) \ U}" using imf img by blast also have "... = U" using U by auto finally have eq: "T \ g -` (g ` T \ (S \ f -` U)) = U" . assume ?lhs then have *: "openin (top_of_set (g ` T)) (g ` T \ (S \ f -` U))" by (meson img openin_Int openin_subtopology_Int_subset openin_subtopology_self) show ?rhs using g [OF *] eq by auto next assume rhs: ?rhs show ?lhs by (metis f fg image_eqI image_subset_iff imf img openin_subopen openin_subtopology_self openin_trans rhs) qed qed lemma continuous_left_inverse_imp_quotient_map: assumes "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g(f x) = x" and "U \ f ` S" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set (f ` S)) U" apply (rule continuous_right_inverse_imp_quotient_map) using assms apply force+ done lemma continuous_imp_quotient_map: fixes f :: "'a::t2_space \ 'b::t2_space" assumes "continuous_on S f" "f ` S = T" "compact S" "U \ T" shows "openin (top_of_set S) (S \ f -` U) \ openin (top_of_set T) U" by (metis (no_types, lifting) assms closed_map_imp_quotient_map continuous_imp_closed_map) subsection\<^marker>\tag unimportant\\Pasting lemmas for functions, for of casewise definitions\ subsubsection\on open sets\ lemma pasting_lemma: assumes ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_openin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace by fastforce next fix U assume Y: "openin Y U" have T: "T i \ topspace X" if "i \ I" for i using ope by (simp add: openin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have "\i. i \ I \ openin X (T i \ f i -` U)" using cont unfolding continuous_map_openin_preimage_eq by (metis Y T inf.commute inf_absorb1 ope topspace_subtopology openin_trans_full) then show "openin X (topspace X \ g -` U)" by (auto simp: *) qed lemma pasting_lemma_exists: assumes X: "topspace X \ (\i \ I. T i)" and ope: "\i. i \ I \ openin X (T i)" and cont: "\i. i \ I \ continuous_map (subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof let ?h = "\x. f (SOME i. i \ I \ x \ T i) x" show "continuous_map X Y ?h" apply (rule pasting_lemma [OF ope cont]) apply (blast intro: f)+ by (metis (no_types, lifting) UN_E X subsetD someI_ex) show "f (SOME i. i \ I \ x \ T i) x = f i x" if "i \ I" "x \ topspace X \ T i" for i x by (metis (no_types, lifting) IntD2 IntI f someI_ex that) qed lemma pasting_lemma_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" unfolding continuous_map_closedin_preimage_eq proof (intro conjI allI impI) show "g ` topspace X \ topspace Y" using g cont continuous_map_image_subset_topspace by fastforce next fix U assume Y: "closedin Y U" have T: "T i \ topspace X" if "i \ I" for i using clo by (simp add: closedin_subset that) have *: "topspace X \ g -` U = (\i \ I. T i \ f i -` U)" using f g T by fastforce have cTf: "\i. i \ I \ closedin X (T i \ f i -` U)" using cont unfolding continuous_map_closedin_preimage_eq topspace_subtopology by (simp add: Int_absorb1 T Y clo closedin_closed_subtopology) have sub: "{Z \ (\i. T i \ f i -` U) ` I. Z \ V \ {}} \ (\i. T i \ f i -` U) ` {i \ I. T i \ V \ {}}" for V by auto have 1: "(\i\I. T i \ f i -` U) \ topspace X" using T by blast then have lf: "locally_finite_in X ((\i. T i \ f i -` U) ` I)" unfolding locally_finite_in_def using finite_subset [OF sub] fin by force show "closedin X (topspace X \ g -` U)" apply (subst *) apply (rule closedin_locally_finite_Union) apply (auto intro: cTf lf) done qed subsubsection\Likewise on closed sets, with a finiteness assumption\ lemma pasting_lemma_closed: assumes fin: "finite I" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" shows "continuous_map X Y g" using pasting_lemma_locally_finite [OF _ clo cont f g] fin by auto lemma pasting_lemma_exists_locally_finite: assumes fin: "\x. x \ topspace X \ \V. openin X V \ x \ V \ finite {i \ I. T i \ V \ {}}" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" and g: "\x. x \ topspace X \ \j. j \ I \ x \ T j \ g x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f(@i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_locally_finite [OF fin]) apply (blast intro: assms)+ by (metis (no_types, lifting) UN_E X set_rev_mp someI_ex) next fix x i assume "i \ I" and "x \ topspace X \ T i" show "f (SOME i. i \ I \ x \ T i) x = f i x" apply (rule someI2_ex) using \i \ I\ \x \ topspace X \ T i\ apply blast by (meson Int_iff \i \ I\ \x \ topspace X \ T i\ f) qed lemma pasting_lemma_exists_closed: assumes fin: "finite I" and X: "topspace X \ \(T ` I)" and clo: "\i. i \ I \ closedin X (T i)" and cont: "\i. i \ I \ continuous_map(subtopology X (T i)) Y (f i)" and f: "\i j x. \i \ I; j \ I; x \ topspace X \ T i \ T j\ \ f i x = f j x" obtains g where "continuous_map X Y g" "\x i. \i \ I; x \ topspace X \ T i\ \ g x = f i x" proof show "continuous_map X Y (\x. f (SOME i. i \ I \ x \ T i) x)" apply (rule pasting_lemma_closed [OF \finite I\ clo cont]) apply (blast intro: f)+ by (metis (mono_tags, lifting) UN_iff X someI_ex subset_iff) next fix x i assume "i \ I" "x \ topspace X \ T i" then show "f (SOME i. i \ I \ x \ T i) x = f i x" by (metis (no_types, lifting) IntD2 IntI f someI_ex) qed lemma continuous_map_cases: assumes f: "continuous_map (subtopology X (X closure_of {x. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x. \ P x})) Y g" and fg: "\x. x \ X frontier_of {x. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" proof (rule pasting_lemma_closed) let ?f = "\b. if b then f else g" let ?g = "\x. if P x then f x else g x" let ?T = "\b. if b then X closure_of {x. P x} else X closure_of {x. ~P x}" show "finite {True,False}" by auto have eq: "topspace X - Collect P = topspace X \ {x. \ P x}" by blast show "?f i x = ?f j x" if "i \ {True,False}" "j \ {True,False}" and x: "x \ topspace X \ ?T i \ ?T j" for i j x proof - have "f x = g x" if "i" "\ j" apply (rule fg) unfolding frontier_of_closures eq using x that closure_of_restrict by fastforce moreover have "g x = f x" if "x \ X closure_of {x. \ P x}" "x \ X closure_of Collect P" "\ i" "j" for x apply (rule fg [symmetric]) unfolding frontier_of_closures eq using x that closure_of_restrict by fastforce ultimately show ?thesis using that by (auto simp flip: closure_of_restrict) qed show "\j. j \ {True,False} \ x \ ?T j \ (if P x then f x else g x) = ?f j x" if "x \ topspace X" for x apply simp apply safe apply (metis Int_iff closure_of inf_sup_absorb mem_Collect_eq that) by (metis DiffI eq closure_of_subset_Int contra_subsetD mem_Collect_eq that) qed (auto simp: f g) lemma continuous_map_cases_alt: assumes f: "continuous_map (subtopology X (X closure_of {x \ topspace X. P x})) Y f" and g: "continuous_map (subtopology X (X closure_of {x \ topspace X. ~P x})) Y g" and fg: "\x. x \ X frontier_of {x \ topspace X. P x} \ f x = g x" shows "continuous_map X Y (\x. if P x then f x else g x)" apply (rule continuous_map_cases) using assms apply (simp_all add: Collect_conj_eq closure_of_restrict [symmetric] frontier_of_restrict [symmetric]) done lemma continuous_map_cases_function: assumes contp: "continuous_map X Z p" and contf: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of U}) Y f" and contg: "continuous_map (subtopology X {x \ topspace X. p x \ Z closure_of (topspace Z - U)}) Y g" and fg: "\x. \x \ topspace X; p x \ Z frontier_of U\ \ f x = g x" shows "continuous_map X Y (\x. if p x \ U then f x else g x)" proof (rule continuous_map_cases_alt) show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y f" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of U}" show "continuous_map (subtopology X ?T) Y f" by (simp add: contf) show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule continuous_map_closure_preimage_subset [OF contp]) qed show "continuous_map (subtopology X (X closure_of {x \ topspace X. p x \ U})) Y g" proof (rule continuous_map_from_subtopology_mono) let ?T = "{x \ topspace X. p x \ Z closure_of (topspace Z - U)}" show "continuous_map (subtopology X ?T) Y g" by (simp add: contg) have "X closure_of {x \ topspace X. p x \ U} \ X closure_of {x \ topspace X. p x \ topspace Z - U}" apply (rule closure_of_mono) using continuous_map_closedin contp by fastforce then show "X closure_of {x \ topspace X. p x \ U} \ ?T" by (rule order_trans [OF _ continuous_map_closure_preimage_subset [OF contp]]) qed next show "f x = g x" if "x \ X frontier_of {x \ topspace X. p x \ U}" for x using that continuous_map_frontier_frontier_preimage_subset [OF contp, of U] fg by blast qed subsection \Retractions\ definition\<^marker>\tag important\ retraction :: "('a::topological_space) set \ 'a set \ ('a \ 'a) \ bool" where "retraction S T r \ T \ S \ continuous_on S r \ r ` S \ T \ (\x\T. r x = x)" definition\<^marker>\tag important\ retract_of (infixl "retract'_of" 50) where "T retract_of S \ (\r. retraction S T r)" lemma retraction_idempotent: "retraction S T r \ x \ S \ r (r x) = r x" unfolding retraction_def by auto text \Preservation of fixpoints under (more general notion of) retraction\ lemma invertible_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes contt: "continuous_on T i" and "i ` T \ S" and contr: "continuous_on S r" and "r ` S \ T" and ri: "\y. y \ T \ r (i y) = y" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - have "\x\S. (i \ g \ r) x = x" proof (rule FP) show "continuous_on S (i \ g \ r)" by (meson contt contr assms(4) contg assms(8) continuous_on_compose continuous_on_subset) show "(i \ g \ r) ` S \ S" using assms(2,4,8) by force qed then obtain x where x: "x \ S" "(i \ g \ r) x = x" .. then have *: "g (r x) \ T" using assms(4,8) by auto have "r ((i \ g \ r) x) = r x" using x by auto then show ?thesis using "*" ri that by auto qed lemma homeomorphic_fixpoint_property: fixes S :: "'a::topological_space set" and T :: "'b::topological_space set" assumes "S homeomorphic T" shows "(\f. continuous_on S f \ f ` S \ S \ (\x\S. f x = x)) \ (\g. continuous_on T g \ g ` T \ T \ (\y\T. g y = y))" (is "?lhs = ?rhs") proof - obtain r i where r: "\x\S. i (r x) = x" "r ` S = T" "continuous_on S r" "\y\T. r (i y) = y" "i ` T = S" "continuous_on T i" using assms unfolding homeomorphic_def homeomorphism_def by blast show ?thesis proof assume ?lhs with r show ?rhs by (metis invertible_fixpoint_property[of T i S r] order_refl) next assume ?rhs with r show ?lhs by (metis invertible_fixpoint_property[of S r T i] order_refl) qed qed lemma retract_fixpoint_property: fixes f :: "'a::topological_space \ 'b::topological_space" and S :: "'a set" assumes "T retract_of S" and FP: "\f. \continuous_on S f; f ` S \ S\ \ \x\S. f x = x" and contg: "continuous_on T g" and "g ` T \ T" obtains y where "y \ T" and "g y = y" proof - obtain h where "retraction S T h" using assms(1) unfolding retract_of_def .. then show ?thesis unfolding retraction_def using invertible_fixpoint_property[OF continuous_on_id _ _ _ _ FP] by (metis assms(4) contg image_ident that) qed lemma retraction: "retraction S T r \ T \ S \ continuous_on S r \ r ` S = T \ (\x \ T. r x = x)" by (force simp: retraction_def) lemma retractionE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "retraction S T r" obtains "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof (rule that) from retraction [of S T r] assms have "T \ S" "continuous_on S r" "r ` S = T" and "\x \ T. r x = x" by simp_all then show "T = r ` S" "r ` S \ S" "continuous_on S r" by simp_all from \\x \ T. r x = x\ have "r x = x" if "x \ T" for x using that by simp with \r ` S = T\ show "r (r x) = r x" if "x \ S" for x using that by auto qed lemma retract_ofE: \ \yields properties normalized wrt. simp -- less likely to loop\ assumes "T retract_of S" obtains r where "T = r ` S" "r ` S \ S" "continuous_on S r" "\x. x \ S \ r (r x) = r x" proof - from assms obtain r where "retraction S T r" by (auto simp add: retract_of_def) with that show thesis by (auto elim: retractionE) qed lemma retract_of_imp_extensible: assumes "S retract_of T" and "continuous_on S f" and "f ` S \ U" obtains g where "continuous_on T g" "g ` T \ U" "\x. x \ S \ g x = f x" proof - from \S retract_of T\ obtain r where "retraction T S r" by (auto simp add: retract_of_def) show thesis by (rule that [of "f \ r"]) (use \continuous_on S f\ \f ` S \ U\ \retraction T S r\ in \auto simp: continuous_on_compose2 retraction\) qed lemma idempotent_imp_retraction: assumes "continuous_on S f" and "f ` S \ S" and "\x. x \ S \ f(f x) = f x" shows "retraction S (f ` S) f" by (simp add: assms retraction) lemma retraction_subset: assumes "retraction S T r" and "T \ s'" and "s' \ S" shows "retraction s' T r" unfolding retraction_def by (metis assms continuous_on_subset image_mono retraction) lemma retract_of_subset: assumes "T retract_of S" and "T \ s'" and "s' \ S" shows "T retract_of s'" by (meson assms retract_of_def retraction_subset) lemma retraction_refl [simp]: "retraction S S (\x. x)" by (simp add: retraction) lemma retract_of_refl [iff]: "S retract_of S" unfolding retract_of_def retraction_def using continuous_on_id by blast lemma retract_of_imp_subset: "S retract_of T \ S \ T" by (simp add: retract_of_def retraction_def) lemma retract_of_empty [simp]: "({} retract_of S) \ S = {}" "(S retract_of {}) \ S = {}" by (auto simp: retract_of_def retraction_def) lemma retract_of_singleton [iff]: "({x} retract_of S) \ x \ S" unfolding retract_of_def retraction_def by force lemma retraction_comp: "\retraction S T f; retraction T U g\ \ retraction S U (g \ f)" apply (auto simp: retraction_def intro: continuous_on_compose2) by blast lemma retract_of_trans [trans]: assumes "S retract_of T" and "T retract_of U" shows "S retract_of U" using assms by (auto simp: retract_of_def intro: retraction_comp) lemma closedin_retract: fixes S :: "'a :: t2_space set" assumes "S retract_of T" shows "closedin (top_of_set T) S" proof - obtain r where r: "S \ T" "continuous_on T r" "r ` T \ S" "\x. x \ S \ r x = x" using assms by (auto simp: retract_of_def retraction_def) have "S = {x\T. x = r x}" using r by auto also have "\ = T \ ((\x. (x, r x)) -` ({y. \x. y = (x, x)}))" unfolding vimage_def mem_Times_iff fst_conv snd_conv using r by auto also have "closedin (top_of_set T) \" by (rule continuous_closedin_preimage) (auto intro!: closed_diagonal continuous_on_Pair r) finally show ?thesis . qed lemma closedin_self [simp]: "closedin (top_of_set S) S" by simp lemma retract_of_closed: fixes S :: "'a :: t2_space set" shows "\closed T; S retract_of T\ \ closed S" by (metis closedin_retract closedin_closed_eq) lemma retract_of_compact: "\compact T; S retract_of T\ \ compact S" by (metis compact_continuous_image retract_of_def retraction) lemma retract_of_connected: "\connected T; S retract_of T\ \ connected S" by (metis Topological_Spaces.connected_continuous_image retract_of_def retraction) lemma retraction_openin_vimage_iff: "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" if retraction: "retraction S T r" and "U \ T" using retraction apply (rule retractionE) apply (rule continuous_right_inverse_imp_quotient_map [where g=r]) using \U \ T\ apply (auto elim: continuous_on_subset) done lemma retract_of_Times: "\S retract_of s'; T retract_of t'\ \ (S \ T) retract_of (s' \ t')" apply (simp add: retract_of_def retraction_def Sigma_mono, clarify) apply (rename_tac f g) apply (rule_tac x="\z. ((f \ fst) z, (g \ snd) z)" in exI) apply (rule conjI continuous_intros | erule continuous_on_subset | force)+ done subsection\Retractions on a topological space\ definition retract_of_space :: "'a set \ 'a topology \ bool" (infix "retract'_of'_space" 50) where "S retract_of_space X \ S \ topspace X \ (\r. continuous_map X (subtopology X S) r \ (\x \ S. r x = x))" lemma retract_of_space_retraction_maps: "S retract_of_space X \ S \ topspace X \ (\r. retraction_maps X (subtopology X S) r id)" by (auto simp: retract_of_space_def retraction_maps_def) lemma retract_of_space_section_map: "S retract_of_space X \ S \ topspace X \ section_map (subtopology X S) X id" unfolding retract_of_space_def retraction_maps_def section_map_def by (auto simp: continuous_map_from_subtopology) lemma retract_of_space_imp_subset: "S retract_of_space X \ S \ topspace X" by (simp add: retract_of_space_def) lemma retract_of_space_topspace: "topspace X retract_of_space X" using retract_of_space_def by force lemma retract_of_space_empty [simp]: "{} retract_of_space X \ topspace X = {}" by (auto simp: continuous_map_def retract_of_space_def) lemma retract_of_space_singleton [simp]: "{a} retract_of_space X \ a \ topspace X" proof - have "continuous_map X (subtopology X {a}) (\x. a) \ (\x. a) a = a" if "a \ topspace X" using that by simp then show ?thesis by (force simp: retract_of_space_def) qed lemma retract_of_space_clopen: assumes "openin X S" "closedin X S" "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (cases "S = {}") case False then obtain a where "a \ S" by blast show ?thesis unfolding retract_of_space_def proof (intro exI conjI) show "S \ topspace X" by (simp add: assms closedin_subset) have "continuous_map X X (\x. if x \ S then x else a)" proof (rule continuous_map_cases) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. x)" by (simp add: continuous_map_from_subtopology) show "continuous_map (subtopology X (X closure_of {x. x \ S})) X (\x. a)" using \S \ topspace X\ \a \ S\ by force show "x = a" if "x \ X frontier_of {x. x \ S}" for x using assms that clopenin_eq_frontier_of by fastforce qed then show "continuous_map X (subtopology X S) (\x. if x \ S then x else a)" using \S \ topspace X\ \a \ S\ by (auto simp: continuous_map_in_subtopology) qed auto qed (use assms in auto) lemma retract_of_space_disjoint_union: assumes "openin X S" "openin X T" and ST: "disjnt S T" "S \ T = topspace X" and "S = {} \ topspace X = {}" shows "S retract_of_space X" proof (rule retract_of_space_clopen) have "S \ T = {}" by (meson ST disjnt_def) then have "S = topspace X - T" using ST by auto then show "closedin X S" using \openin X T\ by blast qed (auto simp: assms) lemma retraction_maps_section_image1: assumes "retraction_maps X Y r s" shows "s ` (topspace Y) retract_of_space X" unfolding retract_of_space_section_map proof show "s ` topspace Y \ topspace X" using assms continuous_map_image_subset_topspace retraction_maps_def by blast show "section_map (subtopology X (s ` topspace Y)) X id" unfolding section_map_def using assms retraction_maps_to_retract_maps by blast qed lemma retraction_maps_section_image2: "retraction_maps X Y r s \ subtopology X (s ` (topspace Y)) homeomorphic_space Y" using embedding_map_imp_homeomorphic_space homeomorphic_space_sym section_imp_embedding_map section_map_def by blast subsection\Paths and path-connectedness\ definition pathin :: "'a topology \ (real \ 'a) \ bool" where "pathin X g \ continuous_map (subtopology euclideanreal {0..1}) X g" lemma pathin_compose: "\pathin X g; continuous_map X Y f\ \ pathin Y (f \ g)" by (simp add: continuous_map_compose pathin_def) lemma pathin_subtopology: "pathin (subtopology X S) g \ pathin X g \ (\x \ {0..1}. g x \ S)" by (auto simp: pathin_def continuous_map_in_subtopology) lemma pathin_const: "pathin X (\x. a) \ a \ topspace X" by (simp add: pathin_def) lemma path_start_in_topspace: "pathin X g \ g 0 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_finish_in_topspace: "pathin X g \ g 1 \ topspace X" by (force simp: pathin_def continuous_map) lemma path_image_subset_topspace: "pathin X g \ g ` ({0..1}) \ topspace X" by (force simp: pathin_def continuous_map) definition path_connected_space :: "'a topology \ bool" where "path_connected_space X \ \x \ topspace X. \ y \ topspace X. \g. pathin X g \ g 0 = x \ g 1 = y" definition path_connectedin :: "'a topology \ 'a set \ bool" where "path_connectedin X S \ S \ topspace X \ path_connected_space(subtopology X S)" lemma path_connectedin_absolute [simp]: "path_connectedin (subtopology X S) S \ path_connectedin X S" by (simp add: path_connectedin_def subtopology_subtopology) lemma path_connectedin_subset_topspace: "path_connectedin X S \ S \ topspace X" by (simp add: path_connectedin_def) lemma path_connectedin_subtopology: "path_connectedin (subtopology X S) T \ path_connectedin X T \ T \ S" by (auto simp: path_connectedin_def subtopology_subtopology inf.absorb2) lemma path_connectedin: "path_connectedin X S \ S \ topspace X \ (\x \ S. \y \ S. \g. pathin X g \ g ` {0..1} \ S \ g 0 = x \ g 1 = y)" unfolding path_connectedin_def path_connected_space_def pathin_def continuous_map_in_subtopology by (intro conj_cong refl ball_cong) (simp_all add: inf.absorb_iff2) lemma path_connectedin_topspace: "path_connectedin X (topspace X) \ path_connected_space X" by (simp add: path_connectedin_def) lemma path_connected_imp_connected_space: assumes "path_connected_space X" shows "connected_space X" proof - have *: "\S. connectedin X S \ g 0 \ S \ g 1 \ S" if "pathin X g" for g proof (intro exI conjI) have "continuous_map (subtopology euclideanreal {0..1}) X g" using connectedin_absolute that by (simp add: pathin_def) then show "connectedin X (g ` {0..1})" by (rule connectedin_continuous_map_image) auto qed auto show ?thesis using assms by (auto intro: * simp add: path_connected_space_def connected_space_subconnected Ball_def) qed lemma path_connectedin_imp_connectedin: "path_connectedin X S \ connectedin X S" by (simp add: connectedin_def path_connected_imp_connected_space path_connectedin_def) lemma path_connected_space_topspace_empty: "topspace X = {} \ path_connected_space X" by (simp add: path_connected_space_def) lemma path_connectedin_empty [simp]: "path_connectedin X {}" by (simp add: path_connectedin) lemma path_connectedin_singleton [simp]: "path_connectedin X {a} \ a \ topspace X" proof show "path_connectedin X {a} \ a \ topspace X" by (simp add: path_connectedin) show "a \ topspace X \ path_connectedin X {a}" unfolding path_connectedin using pathin_const by fastforce qed lemma path_connectedin_continuous_map_image: assumes f: "continuous_map X Y f" and S: "path_connectedin X S" shows "path_connectedin Y (f ` S)" proof - have fX: "f ` (topspace X) \ topspace Y" by (metis f continuous_map_image_subset_topspace) show ?thesis unfolding path_connectedin proof (intro conjI ballI; clarify?) fix x assume "x \ S" show "f x \ topspace Y" by (meson S fX \x \ S\ image_subset_iff path_connectedin_subset_topspace set_mp) next fix x y assume "x \ S" and "y \ S" then obtain g where g: "pathin X g" "g ` {0..1} \ S" "g 0 = x" "g 1 = y" using S by (force simp: path_connectedin) show "\g. pathin Y g \ g ` {0..1} \ f ` S \ g 0 = f x \ g 1 = f y" proof (intro exI conjI) show "pathin Y (f \ g)" using \pathin X g\ f pathin_compose by auto qed (use g in auto) qed qed lemma path_connectedin_discrete_topology: "path_connectedin (discrete_topology U) S \ S \ U \ (\a. S \ {a})" apply safe using path_connectedin_subset_topspace apply fastforce apply (meson connectedin_discrete_topology path_connectedin_imp_connectedin) using subset_singletonD by fastforce lemma path_connected_space_discrete_topology: "path_connected_space (discrete_topology U) \ (\a. U \ {a})" by (metis path_connectedin_discrete_topology path_connectedin_topspace path_connected_space_topspace_empty subset_singletonD topspace_discrete_topology) lemma homeomorphic_path_connected_space_imp: "\path_connected_space X; X homeomorphic_space Y\ \ path_connected_space Y" unfolding homeomorphic_space_def homeomorphic_maps_def by (metis (no_types, opaque_lifting) continuous_map_closedin continuous_map_image_subset_topspace imageI order_class.order.antisym path_connectedin_continuous_map_image path_connectedin_topspace subsetI) lemma homeomorphic_path_connected_space: "X homeomorphic_space Y \ path_connected_space X \ path_connected_space Y" by (meson homeomorphic_path_connected_space_imp homeomorphic_space_sym) lemma homeomorphic_map_path_connectedness: assumes "homeomorphic_map X Y f" "U \ topspace X" shows "path_connectedin Y (f ` U) \ path_connectedin X U" unfolding path_connectedin_def proof (intro conj_cong homeomorphic_path_connected_space) show "(f ` U \ topspace Y) = (U \ topspace X)" using assms homeomorphic_imp_surjective_map by blast next assume "U \ topspace X" show "subtopology Y (f ` U) homeomorphic_space subtopology X U" using assms unfolding homeomorphic_eq_everything_map by (metis (no_types, opaque_lifting) assms homeomorphic_map_subtopologies homeomorphic_space homeomorphic_space_sym image_mono inf.absorb_iff2) qed lemma homeomorphic_map_path_connectedness_eq: "homeomorphic_map X Y f \ path_connectedin X U \ U \ topspace X \ path_connectedin Y (f ` U)" by (meson homeomorphic_map_path_connectedness path_connectedin_def) subsection\Connected components\ definition connected_component_of :: "'a topology \ 'a \ 'a \ bool" where "connected_component_of X x y \ \T. connectedin X T \ x \ T \ y \ T" abbreviation connected_component_of_set where "connected_component_of_set X x \ Collect (connected_component_of X x)" definition connected_components_of :: "'a topology \ ('a set) set" where "connected_components_of X \ connected_component_of_set X ` topspace X" lemma connected_component_in_topspace: "connected_component_of X x y \ x \ topspace X \ y \ topspace X" by (meson connected_component_of_def connectedin_subset_topspace in_mono) lemma connected_component_of_refl: "connected_component_of X x x \ x \ topspace X" by (meson connected_component_in_topspace connected_component_of_def connectedin_sing insertI1) lemma connected_component_of_sym: "connected_component_of X x y \ connected_component_of X y x" by (meson connected_component_of_def) lemma connected_component_of_trans: "\connected_component_of X x y; connected_component_of X y z\ \ connected_component_of X x z" unfolding connected_component_of_def using connectedin_Un by blast lemma connected_component_of_mono: "\connected_component_of (subtopology X S) x y; S \ T\ \ connected_component_of (subtopology X T) x y" by (metis connected_component_of_def connectedin_subtopology inf.absorb_iff2 subtopology_subtopology) lemma connected_component_of_set: "connected_component_of_set X x = {y. \T. connectedin X T \ x \ T \ y \ T}" by (meson connected_component_of_def) lemma connected_component_of_subset_topspace: "connected_component_of_set X x \ topspace X" using connected_component_in_topspace by force lemma connected_component_of_eq_empty: "connected_component_of_set X x = {} \ (x \ topspace X)" using connected_component_in_topspace connected_component_of_refl by fastforce lemma connected_space_iff_connected_component: "connected_space X \ (\x \ topspace X. \y \ topspace X. connected_component_of X x y)" by (simp add: connected_component_of_def connected_space_subconnected) lemma connected_space_imp_connected_component_of: "\connected_space X; a \ topspace X; b \ topspace X\ \ connected_component_of X a b" by (simp add: connected_space_iff_connected_component) lemma connected_space_connected_component_set: "connected_space X \ (\x \ topspace X. connected_component_of_set X x = topspace X)" using connected_component_of_subset_topspace connected_space_iff_connected_component by fastforce lemma connected_component_of_maximal: "\connectedin X S; x \ S\ \ S \ connected_component_of_set X x" by (meson Ball_Collect connected_component_of_def) lemma connected_component_of_equiv: "connected_component_of X x y \ x \ topspace X \ y \ topspace X \ connected_component_of X x = connected_component_of X y" apply (simp add: connected_component_in_topspace fun_eq_iff) by (meson connected_component_of_refl connected_component_of_sym connected_component_of_trans) lemma connected_component_of_disjoint: "disjnt (connected_component_of_set X x) (connected_component_of_set X y) \ ~(connected_component_of X x y)" using connected_component_of_equiv unfolding disjnt_iff by force lemma connected_component_of_eq: "connected_component_of X x = connected_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ connected_component_of X x y" by (metis Collect_empty_eq_bot connected_component_of_eq_empty connected_component_of_equiv) lemma connectedin_connected_component_of: "connectedin X (connected_component_of_set X x)" proof - have "connected_component_of_set X x = \ {T. connectedin X T \ x \ T}" by (auto simp: connected_component_of_def) then show ?thesis apply (rule ssubst) by (blast intro: connectedin_Union) qed lemma Union_connected_components_of: "\(connected_components_of X) = topspace X" unfolding connected_components_of_def apply (rule equalityI) apply (simp add: SUP_least connected_component_of_subset_topspace) using connected_component_of_refl by fastforce lemma connected_components_of_maximal: "\C \ connected_components_of X; connectedin X S; ~disjnt C S\ \ S \ C" unfolding connected_components_of_def disjnt_def apply clarify by (metis Int_emptyI connected_component_of_def connected_component_of_trans mem_Collect_eq) lemma pairwise_disjoint_connected_components_of: "pairwise disjnt (connected_components_of X)" unfolding connected_components_of_def pairwise_def apply clarify by (metis connected_component_of_disjoint connected_component_of_equiv) lemma complement_connected_components_of_Union: "C \ connected_components_of X \ topspace X - C = \ (connected_components_of X - {C})" apply (rule equalityI) using Union_connected_components_of apply fastforce by (metis Diff_cancel Diff_subset Union_connected_components_of cSup_singleton diff_Union_pairwise_disjoint equalityE insert_subsetI pairwise_disjoint_connected_components_of) lemma nonempty_connected_components_of: "C \ connected_components_of X \ C \ {}" unfolding connected_components_of_def by (metis (no_types, lifting) connected_component_of_eq_empty imageE) lemma connected_components_of_subset: "C \ connected_components_of X \ C \ topspace X" using Union_connected_components_of by fastforce lemma connectedin_connected_components_of: assumes "C \ connected_components_of X" shows "connectedin X C" proof - have "C \ connected_component_of_set X ` topspace X" using assms connected_components_of_def by blast then show ?thesis using connectedin_connected_component_of by fastforce qed lemma connected_component_in_connected_components_of: "connected_component_of_set X a \ connected_components_of X \ a \ topspace X" apply (rule iffI) using connected_component_of_eq_empty nonempty_connected_components_of apply fastforce by (simp add: connected_components_of_def) lemma connected_space_iff_components_eq: "connected_space X \ (\C \ connected_components_of X. \C' \ connected_components_of X. C = C')" apply (rule iffI) apply (force simp: connected_components_of_def connected_space_connected_component_set image_iff) by (metis connected_component_in_connected_components_of connected_component_of_refl connected_space_iff_connected_component mem_Collect_eq) lemma connected_components_of_eq_empty: "connected_components_of X = {} \ topspace X = {}" by (simp add: connected_components_of_def) lemma connected_components_of_empty_space: "topspace X = {} \ connected_components_of X = {}" by (simp add: connected_components_of_eq_empty) lemma connected_components_of_subset_sing: "connected_components_of X \ {S} \ connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (simp add: connected_components_of_empty_space connected_space_topspace_empty) next case False then show ?thesis by (metis (no_types, opaque_lifting) Union_connected_components_of ccpo_Sup_singleton connected_components_of_eq_empty connected_space_iff_components_eq insertI1 singletonD subsetI subset_singleton_iff) qed lemma connected_space_iff_components_subset_singleton: "connected_space X \ (\a. connected_components_of X \ {a})" by (simp add: connected_components_of_subset_sing) lemma connected_components_of_eq_singleton: "connected_components_of X = {S} \ connected_space X \ topspace X \ {} \ S = topspace X" by (metis ccpo_Sup_singleton connected_components_of_subset_sing insert_not_empty subset_singleton_iff) lemma connected_components_of_connected_space: "connected_space X \ connected_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: connected_components_of_eq_empty connected_components_of_eq_singleton) lemma exists_connected_component_of_superset: assumes "connectedin X S" and ne: "topspace X \ {}" shows "\C. C \ connected_components_of X \ S \ C" proof (cases "S = {}") case True then show ?thesis using ne connected_components_of_def by blast next case False then show ?thesis by (meson all_not_in_conv assms(1) connected_component_in_connected_components_of connected_component_of_maximal connectedin_subset_topspace in_mono) qed lemma closedin_connected_components_of: assumes "C \ connected_components_of X" shows "closedin X C" proof - obtain x where "x \ topspace X" and x: "C = connected_component_of_set X x" using assms by (auto simp: connected_components_of_def) have "connected_component_of_set X x \ topspace X" by (simp add: connected_component_of_subset_topspace) moreover have "X closure_of connected_component_of_set X x \ connected_component_of_set X x" proof (rule connected_component_of_maximal) show "connectedin X (X closure_of connected_component_of_set X x)" by (simp add: connectedin_closure_of connectedin_connected_component_of) show "x \ X closure_of connected_component_of_set X x" by (simp add: \x \ topspace X\ closure_of connected_component_of_refl) qed ultimately show ?thesis using closure_of_subset_eq x by auto qed lemma closedin_connected_component_of: "closedin X (connected_component_of_set X x)" by (metis closedin_connected_components_of closedin_empty connected_component_in_connected_components_of connected_component_of_eq_empty) lemma connected_component_of_eq_overlap: "connected_component_of_set X x = connected_component_of_set X y \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x \ connected_component_of_set X y = {})" using connected_component_of_equiv by fastforce lemma connected_component_of_nonoverlap: "connected_component_of_set X x \ connected_component_of_set X y = {} \ (x \ topspace X) \ (y \ topspace X) \ ~(connected_component_of_set X x = connected_component_of_set X y)" by (metis connected_component_of_eq_empty connected_component_of_eq_overlap inf.idem) lemma connected_component_of_overlap: "~(connected_component_of_set X x \ connected_component_of_set X y = {}) \ x \ topspace X \ y \ topspace X \ connected_component_of_set X x = connected_component_of_set X y" by (meson connected_component_of_nonoverlap) end \ No newline at end of file diff --git a/src/HOL/Analysis/Binary_Product_Measure.thy b/src/HOL/Analysis/Binary_Product_Measure.thy --- a/src/HOL/Analysis/Binary_Product_Measure.thy +++ b/src/HOL/Analysis/Binary_Product_Measure.thy @@ -1,1159 +1,1179 @@ (* Title: HOL/Analysis/Binary_Product_Measure.thy Author: Johannes Hölzl, TU München *) section \Binary Product Measure\ theory Binary_Product_Measure imports Nonnegative_Lebesgue_Integration begin lemma Pair_vimage_times[simp]: "Pair x -` (A \ B) = (if x \ A then B else {})" by auto lemma rev_Pair_vimage_times[simp]: "(\x. (x, y)) -` (A \ B) = (if y \ B then A else {})" by auto subsection "Binary products" definition\<^marker>\tag important\ pair_measure (infixr "\\<^sub>M" 80) where "A \\<^sub>M B = measure_of (space A \ space B) {a \ b | a b. a \ sets A \ b \ sets B} (\X. \\<^sup>+x. (\\<^sup>+y. indicator X (x,y) \B) \A)" lemma pair_measure_closed: "{a \ b | a b. a \ sets A \ b \ sets B} \ Pow (space A \ space B)" using sets.space_closed[of A] sets.space_closed[of B] by auto lemma space_pair_measure: "space (A \\<^sub>M B) = space A \ space B" unfolding pair_measure_def using pair_measure_closed[of A B] by (rule space_measure_of) lemma SIGMA_Collect_eq: "(SIGMA x:space M. {y\space N. P x y}) = {x\space (M \\<^sub>M N). P (fst x) (snd x)}" by (auto simp: space_pair_measure) lemma sets_pair_measure: "sets (A \\<^sub>M B) = sigma_sets (space A \ space B) {a \ b | a b. a \ sets A \ b \ sets B}" unfolding pair_measure_def using pair_measure_closed[of A B] by (rule sets_measure_of) lemma sets_pair_measure_cong[measurable_cong, cong]: "sets M1 = sets M1' \ sets M2 = sets M2' \ sets (M1 \\<^sub>M M2) = sets (M1' \\<^sub>M M2')" unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq) lemma pair_measureI[intro, simp, measurable]: "x \ sets A \ y \ sets B \ x \ y \ sets (A \\<^sub>M B)" by (auto simp: sets_pair_measure) lemma sets_Pair: "{x} \ sets M1 \ {y} \ sets M2 \ {(x, y)} \ sets (M1 \\<^sub>M M2)" using pair_measureI[of "{x}" M1 "{y}" M2] by simp lemma measurable_pair_measureI: assumes 1: "f \ space M \ space M1 \ space M2" assumes 2: "\A B. A \ sets M1 \ B \ sets M2 \ f -` (A \ B) \ space M \ sets M" shows "f \ measurable M (M1 \\<^sub>M M2)" unfolding pair_measure_def using 1 2 by (intro measurable_measure_of) (auto dest: sets.sets_into_space) lemma measurable_split_replace[measurable (raw)]: "(\x. f x (fst (g x)) (snd (g x))) \ measurable M N \ (\x. case_prod (f x) (g x)) \ measurable M N" unfolding split_beta' . lemma measurable_Pair[measurable (raw)]: assumes f: "f \ measurable M M1" and g: "g \ measurable M M2" shows "(\x. (f x, g x)) \ measurable M (M1 \\<^sub>M M2)" proof (rule measurable_pair_measureI) show "(\x. (f x, g x)) \ space M \ space M1 \ space M2" using f g by (auto simp: measurable_def) fix A B assume *: "A \ sets M1" "B \ sets M2" have "(\x. (f x, g x)) -` (A \ B) \ space M = (f -` A \ space M) \ (g -` B \ space M)" by auto also have "\ \ sets M" by (rule sets.Int) (auto intro!: measurable_sets * f g) finally show "(\x. (f x, g x)) -` (A \ B) \ space M \ sets M" . qed lemma measurable_fst[intro!, simp, measurable]: "fst \ measurable (M1 \\<^sub>M M2) M1" by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space Times_Int_Times measurable_def) lemma measurable_snd[intro!, simp, measurable]: "snd \ measurable (M1 \\<^sub>M M2) M2" by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space Times_Int_Times measurable_def) lemma measurable_Pair_compose_split[measurable_dest]: assumes f: "case_prod f \ measurable (M1 \\<^sub>M M2) N" assumes g: "g \ measurable M M1" and h: "h \ measurable M M2" shows "(\x. f (g x) (h x)) \ measurable M N" using measurable_compose[OF measurable_Pair f, OF g h] by simp lemma measurable_Pair1_compose[measurable_dest]: assumes f: "(\x. (f x, g x)) \ measurable M (M1 \\<^sub>M M2)" assumes [measurable]: "h \ measurable N M" shows "(\x. f (h x)) \ measurable N M1" using measurable_compose[OF f measurable_fst] by simp lemma measurable_Pair2_compose[measurable_dest]: assumes f: "(\x. (f x, g x)) \ measurable M (M1 \\<^sub>M M2)" assumes [measurable]: "h \ measurable N M" shows "(\x. g (h x)) \ measurable N M2" using measurable_compose[OF f measurable_snd] by simp lemma measurable_pair: assumes "(fst \ f) \ measurable M M1" "(snd \ f) \ measurable M M2" shows "f \ measurable M (M1 \\<^sub>M M2)" using measurable_Pair[OF assms] by simp lemma assumes f[measurable]: "f \ measurable M (N \\<^sub>M P)" shows measurable_fst': "(\x. fst (f x)) \ measurable M N" and measurable_snd': "(\x. snd (f x)) \ measurable M P" by simp_all lemma assumes f[measurable]: "f \ measurable M N" shows measurable_fst'': "(\x. f (fst x)) \ measurable (M \\<^sub>M P) N" and measurable_snd'': "(\x. f (snd x)) \ measurable (P \\<^sub>M M) N" by simp_all lemma sets_pair_in_sets: assumes "\a b. a \ sets A \ b \ sets B \ a \ b \ sets N" shows "sets (A \\<^sub>M B) \ sets N" unfolding sets_pair_measure by (intro sets.sigma_sets_subset') (auto intro!: assms) lemma sets_pair_eq_sets_fst_snd: "sets (A \\<^sub>M B) = sets (Sup {vimage_algebra (space A \ space B) fst A, vimage_algebra (space A \ space B) snd B})" (is "?P = sets (Sup {?fst, ?snd})") proof - { fix a b assume ab: "a \ sets A" "b \ sets B" then have "a \ b = (fst -` a \ (space A \ space B)) \ (snd -` b \ (space A \ space B))" by (auto dest: sets.sets_into_space) also have "\ \ sets (Sup {?fst, ?snd})" apply (rule sets.Int) apply (rule in_sets_Sup) apply auto [] apply (rule insertI1) apply (auto intro: ab in_vimage_algebra) [] apply (rule in_sets_Sup) apply auto [] apply (rule insertI2) apply (auto intro: ab in_vimage_algebra) done finally have "a \ b \ sets (Sup {?fst, ?snd})" . } moreover have "sets ?fst \ sets (A \\<^sub>M B)" by (rule sets_image_in_sets) (auto simp: space_pair_measure[symmetric]) moreover have "sets ?snd \ sets (A \\<^sub>M B)" by (rule sets_image_in_sets) (auto simp: space_pair_measure) ultimately show ?thesis apply (intro antisym[of "sets A" for A] sets_Sup_in_sets sets_pair_in_sets) apply simp apply simp apply simp apply (elim disjE) apply (simp add: space_pair_measure) apply (simp add: space_pair_measure) apply (auto simp add: space_pair_measure) done qed lemma measurable_pair_iff: "f \ measurable M (M1 \\<^sub>M M2) \ (fst \ f) \ measurable M M1 \ (snd \ f) \ measurable M M2" by (auto intro: measurable_pair[of f M M1 M2]) lemma measurable_split_conv: "(\(x, y). f x y) \ measurable A B \ (\x. f (fst x) (snd x)) \ measurable A B" by (intro arg_cong2[where f="(\)"]) auto lemma measurable_pair_swap': "(\(x,y). (y, x)) \ measurable (M1 \\<^sub>M M2) (M2 \\<^sub>M M1)" by (auto intro!: measurable_Pair simp: measurable_split_conv) lemma measurable_pair_swap: assumes f: "f \ measurable (M1 \\<^sub>M M2) M" shows "(\(x,y). f (y, x)) \ measurable (M2 \\<^sub>M M1) M" using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def) lemma measurable_pair_swap_iff: "f \ measurable (M2 \\<^sub>M M1) M \ (\(x,y). f (y,x)) \ measurable (M1 \\<^sub>M M2) M" by (auto dest: measurable_pair_swap) lemma measurable_Pair1': "x \ space M1 \ Pair x \ measurable M2 (M1 \\<^sub>M M2)" by simp lemma sets_Pair1[measurable (raw)]: assumes A: "A \ sets (M1 \\<^sub>M M2)" shows "Pair x -` A \ sets M2" proof - have "Pair x -` A = (if x \ space M1 then Pair x -` A \ space M2 else {})" using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure) also have "\ \ sets M2" using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: if_split_asm) finally show ?thesis . qed lemma measurable_Pair2': "y \ space M2 \ (\x. (x, y)) \ measurable M1 (M1 \\<^sub>M M2)" by (auto intro!: measurable_Pair) lemma sets_Pair2: assumes A: "A \ sets (M1 \\<^sub>M M2)" shows "(\x. (x, y)) -` A \ sets M1" proof - have "(\x. (x, y)) -` A = (if y \ space M2 then (\x. (x, y)) -` A \ space M1 else {})" using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure) also have "\ \ sets M1" using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: if_split_asm) finally show ?thesis . qed lemma measurable_Pair2: assumes f: "f \ measurable (M1 \\<^sub>M M2) M" and x: "x \ space M1" shows "(\y. f (x, y)) \ measurable M2 M" using measurable_comp[OF measurable_Pair1' f, OF x] by (simp add: comp_def) lemma measurable_Pair1: assumes f: "f \ measurable (M1 \\<^sub>M M2) M" and y: "y \ space M2" shows "(\x. f (x, y)) \ measurable M1 M" using measurable_comp[OF measurable_Pair2' f, OF y] by (simp add: comp_def) lemma Int_stable_pair_measure_generator: "Int_stable {a \ b | a b. a \ sets A \ b \ sets B}" unfolding Int_stable_def by safe (auto simp add: Times_Int_Times) lemma (in finite_measure) finite_measure_cut_measurable: assumes [measurable]: "Q \ sets (N \\<^sub>M M)" shows "(\x. emeasure M (Pair x -` Q)) \ borel_measurable N" (is "?s Q \ _") using Int_stable_pair_measure_generator pair_measure_closed assms unfolding sets_pair_measure proof (induct rule: sigma_sets_induct_disjoint) case (compl A) with sets.sets_into_space have "\x. emeasure M (Pair x -` ((space N \ space M) - A)) = (if x \ space N then emeasure M (space M) - ?s A x else 0)" unfolding sets_pair_measure[symmetric] by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1) with compl sets.top show ?case by (auto intro!: measurable_If simp: space_pair_measure) next case (union F) then have "\x. emeasure M (Pair x -` (\i. F i)) = (\i. ?s (F i) x)" by (simp add: suminf_emeasure disjoint_family_on_vimageI subset_eq vimage_UN sets_pair_measure[symmetric]) with union show ?case unfolding sets_pair_measure[symmetric] by simp qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If) lemma (in sigma_finite_measure) measurable_emeasure_Pair: assumes Q: "Q \ sets (N \\<^sub>M M)" shows "(\x. emeasure M (Pair x -` Q)) \ borel_measurable N" (is "?s Q \ _") proof - - from sigma_finite_disjoint guess F . note F = this + obtain F :: "nat \ 'a set" where F: + "range F \ sets M" + "\ (range F) = space M" + "\i. emeasure M (F i) \ \" + "disjoint_family F" by (blast intro: sigma_finite_disjoint) then have F_sets: "\i. F i \ sets M" by auto let ?C = "\x i. F i \ Pair x -` Q" { fix i have [simp]: "space N \ F i \ space N \ space M = space N \ F i" using F sets.sets_into_space by auto let ?R = "density M (indicator (F i))" have "finite_measure ?R" using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq) then have "(\x. emeasure ?R (Pair x -` (space N \ space ?R \ Q))) \ borel_measurable N" by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q) moreover have "\x. emeasure ?R (Pair x -` (space N \ space ?R \ Q)) = emeasure M (F i \ Pair x -` (space N \ space ?R \ Q))" using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1) moreover have "\x. F i \ Pair x -` (space N \ space ?R \ Q) = ?C x i" using sets.sets_into_space[OF Q] by (auto simp: space_pair_measure) ultimately have "(\x. emeasure M (?C x i)) \ borel_measurable N" by simp } moreover { fix x have "(\i. emeasure M (?C x i)) = emeasure M (\i. ?C x i)" proof (intro suminf_emeasure) show "range (?C x) \ sets M" using F \Q \ sets (N \\<^sub>M M)\ by (auto intro!: sets_Pair1) have "disjoint_family F" using F by auto show "disjoint_family (?C x)" by (rule disjoint_family_on_bisimulation[OF \disjoint_family F\]) auto qed also have "(\i. ?C x i) = Pair x -` Q" using F sets.sets_into_space[OF \Q \ sets (N \\<^sub>M M)\] by (auto simp: space_pair_measure) finally have "emeasure M (Pair x -` Q) = (\i. emeasure M (?C x i))" by simp } ultimately show ?thesis using \Q \ sets (N \\<^sub>M M)\ F_sets by auto qed lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]: assumes space: "\x. x \ space N \ A x \ space M" assumes A: "{x\space (N \\<^sub>M M). snd x \ A (fst x)} \ sets (N \\<^sub>M M)" shows "(\x. emeasure M (A x)) \ borel_measurable N" proof - from space have "\x. x \ space N \ Pair x -` {x \ space (N \\<^sub>M M). snd x \ A (fst x)} = A x" by (auto simp: space_pair_measure) with measurable_emeasure_Pair[OF A] show ?thesis by (auto cong: measurable_cong) qed lemma (in sigma_finite_measure) emeasure_pair_measure: assumes "X \ sets (N \\<^sub>M M)" shows "emeasure (N \\<^sub>M M) X = (\\<^sup>+ x. \\<^sup>+ y. indicator X (x, y) \M \N)" (is "_ = ?\ X") proof (rule emeasure_measure_of[OF pair_measure_def]) show "positive (sets (N \\<^sub>M M)) ?\" by (auto simp: positive_def) have eq[simp]: "\A x y. indicator A (x, y) = indicator (Pair x -` A) y" by (auto simp: indicator_def) show "countably_additive (sets (N \\<^sub>M M)) ?\" proof (rule countably_additiveI) fix F :: "nat \ ('b \ 'a) set" assume F: "range F \ sets (N \\<^sub>M M)" "disjoint_family F" from F have *: "\i. F i \ sets (N \\<^sub>M M)" by auto moreover have "\x. disjoint_family (\i. Pair x -` F i)" by (intro disjoint_family_on_bisimulation[OF F(2)]) auto moreover have "\x. range (\i. Pair x -` F i) \ sets M" using F by (auto simp: sets_Pair1) ultimately show "(\n. ?\ (F n)) = ?\ (\i. F i)" by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure intro!: nn_integral_cong nn_integral_indicator[symmetric]) qed show "{a \ b |a b. a \ sets N \ b \ sets M} \ Pow (space N \ space M)" using sets.space_closed[of N] sets.space_closed[of M] by auto qed fact lemma (in sigma_finite_measure) emeasure_pair_measure_alt: assumes X: "X \ sets (N \\<^sub>M M)" shows "emeasure (N \\<^sub>M M) X = (\\<^sup>+x. emeasure M (Pair x -` X) \N)" proof - have [simp]: "\x y. indicator X (x, y) = indicator (Pair x -` X) y" by (auto simp: indicator_def) show ?thesis using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure sets_Pair1) qed proposition (in sigma_finite_measure) emeasure_pair_measure_Times: assumes A: "A \ sets N" and B: "B \ sets M" shows "emeasure (N \\<^sub>M M) (A \ B) = emeasure N A * emeasure M B" proof - have "emeasure (N \\<^sub>M M) (A \ B) = (\\<^sup>+x. emeasure M B * indicator A x \N)" using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt) also have "\ = emeasure M B * emeasure N A" using A by (simp add: nn_integral_cmult_indicator) finally show ?thesis by (simp add: ac_simps) qed subsection \Binary products of \\\-finite emeasure spaces\ locale\<^marker>\tag unimportant\ pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2 for M1 :: "'a measure" and M2 :: "'b measure" lemma (in pair_sigma_finite) measurable_emeasure_Pair1: "Q \ sets (M1 \\<^sub>M M2) \ (\x. emeasure M2 (Pair x -` Q)) \ borel_measurable M1" using M2.measurable_emeasure_Pair . lemma (in pair_sigma_finite) measurable_emeasure_Pair2: assumes Q: "Q \ sets (M1 \\<^sub>M M2)" shows "(\y. emeasure M1 ((\x. (x, y)) -` Q)) \ borel_measurable M2" proof - have "(\(x, y). (y, x)) -` Q \ space (M2 \\<^sub>M M1) \ sets (M2 \\<^sub>M M1)" using Q measurable_pair_swap' by (auto intro: measurable_sets) note M1.measurable_emeasure_Pair[OF this] moreover have "\y. Pair y -` ((\(x, y). (y, x)) -` Q \ space (M2 \\<^sub>M M1)) = (\x. (x, y)) -` Q" using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure) ultimately show ?thesis by simp qed proposition (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: defines "E \ {A \ B | A B. A \ sets M1 \ B \ sets M2}" shows "\F::nat \ ('a \ 'b) set. range F \ E \ incseq F \ (\i. F i) = space M1 \ space M2 \ (\i. emeasure (M1 \\<^sub>M M2) (F i) \ \)" proof - - from M1.sigma_finite_incseq guess F1 . note F1 = this - from M2.sigma_finite_incseq guess F2 . note F2 = this + obtain F1 where F1: "range F1 \ sets M1" + "\ (range F1) = space M1" + "\i. emeasure M1 (F1 i) \ \" + "incseq F1" + by (rule M1.sigma_finite_incseq) blast + obtain F2 where F2: "range F2 \ sets M2" + "\ (range F2) = space M2" + "\i. emeasure M2 (F2 i) \ \" + "incseq F2" + by (rule M2.sigma_finite_incseq) blast from F1 F2 have space: "space M1 = (\i. F1 i)" "space M2 = (\i. F2 i)" by auto let ?F = "\i. F1 i \ F2 i" show ?thesis proof (intro exI[of _ ?F] conjI allI) show "range ?F \ E" using F1 F2 by (auto simp: E_def) (metis range_subsetD) next have "space M1 \ space M2 \ (\i. ?F i)" proof (intro subsetI) fix x assume "x \ space M1 \ space M2" then obtain i j where "fst x \ F1 i" "snd x \ F2 j" by (auto simp: space) then have "fst x \ F1 (max i j)" "snd x \ F2 (max j i)" using \incseq F1\ \incseq F2\ unfolding incseq_def by (force split: split_max)+ then have "(fst x, snd x) \ F1 (max i j) \ F2 (max i j)" by (intro SigmaI) (auto simp add: max.commute) then show "x \ (\i. ?F i)" by auto qed then show "(\i. ?F i) = space M1 \ space M2" using space by (auto simp: space) next fix i show "incseq (\i. F1 i \ F2 i)" using \incseq F1\ \incseq F2\ unfolding incseq_Suc_iff by auto next fix i from F1 F2 have "F1 i \ sets M1" "F2 i \ sets M2" by auto with F1 F2 show "emeasure (M1 \\<^sub>M M2) (F1 i \ F2 i) \ \" by (auto simp add: emeasure_pair_measure_Times ennreal_mult_eq_top_iff) qed qed sublocale\<^marker>\tag unimportant\ pair_sigma_finite \ P?: sigma_finite_measure "M1 \\<^sub>M M2" proof - from M1.sigma_finite_countable guess F1 .. - moreover from M2.sigma_finite_countable guess F2 .. - ultimately show + obtain F1 :: "'a set set" and F2 :: "'b set set" where + "countable F1 \ F1 \ sets M1 \ \ F1 = space M1 \ (\a\F1. emeasure M1 a \ \)" + "countable F2 \ F2 \ sets M2 \ \ F2 = space M2 \ (\a\F2. emeasure M2 a \ \)" + using M1.sigma_finite_countable M2.sigma_finite_countable by auto + then show "\A. countable A \ A \ sets (M1 \\<^sub>M M2) \ \A = space (M1 \\<^sub>M M2) \ (\a\A. emeasure (M1 \\<^sub>M M2) a \ \)" by (intro exI[of _ "(\(a, b). a \ b) ` (F1 \ F2)"] conjI) (auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq ennreal_mult_eq_top_iff) qed lemma sigma_finite_pair_measure: assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B" shows "sigma_finite_measure (A \\<^sub>M B)" proof - interpret A: sigma_finite_measure A by fact interpret B: sigma_finite_measure B by fact interpret AB: pair_sigma_finite A B .. show ?thesis .. qed lemma sets_pair_swap: assumes "A \ sets (M1 \\<^sub>M M2)" shows "(\(x, y). (y, x)) -` A \ space (M2 \\<^sub>M M1) \ sets (M2 \\<^sub>M M1)" using measurable_pair_swap' assms by (rule measurable_sets) lemma (in pair_sigma_finite) distr_pair_swap: "M1 \\<^sub>M M2 = distr (M2 \\<^sub>M M1) (M1 \\<^sub>M M2) (\(x, y). (y, x))" (is "?P = ?D") proof - - from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this let ?E = "{a \ b |a b. a \ sets M1 \ b \ sets M2}" + obtain F :: "nat \ ('a \ 'b) set" where F: "range F \ ?E" + "incseq F" "\ (range F) = space M1 \ space M2" "\i. emeasure (M1 \\<^sub>M M2) (F i) \ \" + using sigma_finite_up_in_pair_measure_generator by auto show ?thesis proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) show "?E \ Pow (space ?P)" using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure) show "sets ?P = sigma_sets (space ?P) ?E" by (simp add: sets_pair_measure space_pair_measure) then show "sets ?D = sigma_sets (space ?P) ?E" by simp - next - show "range F \ ?E" "(\i. F i) = space ?P" "\i. emeasure ?P (F i) \ \" - using F by (auto simp: space_pair_measure) + from F show "range F \ ?E" "(\i. F i) = space ?P" "\i. emeasure ?P (F i) \ \" + by (auto simp: space_pair_measure) next fix X assume "X \ ?E" then obtain A B where X[simp]: "X = A \ B" and A: "A \ sets M1" and B: "B \ sets M2" by auto have "(\(y, x). (x, y)) -` X \ space (M2 \\<^sub>M M1) = B \ A" using sets.sets_into_space[OF A] sets.sets_into_space[OF B] by (auto simp: space_pair_measure) with A B show "emeasure (M1 \\<^sub>M M2) X = emeasure ?D X" by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr measurable_pair_swap' ac_simps) qed qed lemma (in pair_sigma_finite) emeasure_pair_measure_alt2: assumes A: "A \ sets (M1 \\<^sub>M M2)" shows "emeasure (M1 \\<^sub>M M2) A = (\\<^sup>+y. emeasure M1 ((\x. (x, y)) -` A) \M2)" (is "_ = ?\ A") proof - have [simp]: "\y. (Pair y -` ((\(x, y). (y, x)) -` A \ space (M2 \\<^sub>M M1))) = (\x. (x, y)) -` A" using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) show ?thesis using A by (subst distr_pair_swap) (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap'] M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A]) qed lemma (in pair_sigma_finite) AE_pair: assumes "AE x in (M1 \\<^sub>M M2). Q x" shows "AE x in M1. (AE y in M2. Q (x, y))" proof - obtain N where N: "N \ sets (M1 \\<^sub>M M2)" "emeasure (M1 \\<^sub>M M2) N = 0" "{x\space (M1 \\<^sub>M M2). \ Q x} \ N" using assms unfolding eventually_ae_filter by auto show ?thesis proof (rule AE_I) from N measurable_emeasure_Pair1[OF \N \ sets (M1 \\<^sub>M M2)\] show "emeasure M1 {x\space M1. emeasure M2 (Pair x -` N) \ 0} = 0" by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff) show "{x \ space M1. emeasure M2 (Pair x -` N) \ 0} \ sets M1" by (intro borel_measurable_eq measurable_emeasure_Pair1 N sets.sets_Collect_neg N) simp { fix x assume "x \ space M1" "emeasure M2 (Pair x -` N) = 0" have "AE y in M2. Q (x, y)" proof (rule AE_I) show "emeasure M2 (Pair x -` N) = 0" by fact show "Pair x -` N \ sets M2" using N(1) by (rule sets_Pair1) show "{y \ space M2. \ Q (x, y)} \ Pair x -` N" using N \x \ space M1\ unfolding space_pair_measure by auto qed } then show "{x \ space M1. \ (AE y in M2. Q (x, y))} \ {x \ space M1. emeasure M2 (Pair x -` N) \ 0}" by auto qed qed lemma (in pair_sigma_finite) AE_pair_measure: assumes "{x\space (M1 \\<^sub>M M2). P x} \ sets (M1 \\<^sub>M M2)" assumes ae: "AE x in M1. AE y in M2. P (x, y)" shows "AE x in M1 \\<^sub>M M2. P x" proof (subst AE_iff_measurable[OF _ refl]) show "{x\space (M1 \\<^sub>M M2). \ P x} \ sets (M1 \\<^sub>M M2)" by (rule sets.sets_Collect) fact then have "emeasure (M1 \\<^sub>M M2) {x \ space (M1 \\<^sub>M M2). \ P x} = (\\<^sup>+ x. \\<^sup>+ y. indicator {x \ space (M1 \\<^sub>M M2). \ P x} (x, y) \M2 \M1)" by (simp add: M2.emeasure_pair_measure) also have "\ = (\\<^sup>+ x. \\<^sup>+ y. 0 \M2 \M1)" using ae apply (safe intro!: nn_integral_cong_AE) apply (intro AE_I2) apply (safe intro!: nn_integral_cong_AE) apply auto done finally show "emeasure (M1 \\<^sub>M M2) {x \ space (M1 \\<^sub>M M2). \ P x} = 0" by simp qed lemma (in pair_sigma_finite) AE_pair_iff: "{x\space (M1 \\<^sub>M M2). P (fst x) (snd x)} \ sets (M1 \\<^sub>M M2) \ (AE x in M1. AE y in M2. P x y) \ (AE x in (M1 \\<^sub>M M2). P (fst x) (snd x))" using AE_pair[of "\x. P (fst x) (snd x)"] AE_pair_measure[of "\x. P (fst x) (snd x)"] by auto lemma (in pair_sigma_finite) AE_commute: assumes P: "{x\space (M1 \\<^sub>M M2). P (fst x) (snd x)} \ sets (M1 \\<^sub>M M2)" shows "(AE x in M1. AE y in M2. P x y) \ (AE y in M2. AE x in M1. P x y)" proof - interpret Q: pair_sigma_finite M2 M1 .. have [simp]: "\x. (fst (case x of (x, y) \ (y, x))) = snd x" "\x. (snd (case x of (x, y) \ (y, x))) = fst x" by auto have "{x \ space (M2 \\<^sub>M M1). P (snd x) (fst x)} = (\(x, y). (y, x)) -` {x \ space (M1 \\<^sub>M M2). P (fst x) (snd x)} \ space (M2 \\<^sub>M M1)" by (auto simp: space_pair_measure) also have "\ \ sets (M2 \\<^sub>M M1)" by (intro sets_pair_swap P) finally show ?thesis apply (subst AE_pair_iff[OF P]) apply (subst distr_pair_swap) apply (subst AE_distr_iff[OF measurable_pair_swap' P]) apply (subst Q.AE_pair_iff) apply simp_all done qed subsection "Fubinis theorem" lemma measurable_compose_Pair1: "x \ space M1 \ g \ measurable (M1 \\<^sub>M M2) L \ (\y. g (x, y)) \ measurable M2 L" by simp lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst: assumes f: "f \ borel_measurable (M1 \\<^sub>M M)" shows "(\x. \\<^sup>+ y. f (x, y) \M) \ borel_measurable M1" using f proof induct case (cong u v) then have "\w x. w \ space M1 \ x \ space M \ u (w, x) = v (w, x)" by (auto simp: space_pair_measure) show ?case apply (subst measurable_cong) apply (rule nn_integral_cong) apply fact+ done next case (set Q) have [simp]: "\x y. indicator Q (x, y) = indicator (Pair x -` Q) y" by (auto simp: indicator_def) have "\x. x \ space M1 \ emeasure M (Pair x -` Q) = \\<^sup>+ y. indicator Q (x, y) \M" by (simp add: sets_Pair1[OF set]) from this measurable_emeasure_Pair[OF set] show ?case by (rule measurable_cong[THEN iffD1]) qed (simp_all add: nn_integral_add nn_integral_cmult measurable_compose_Pair1 nn_integral_monotone_convergence_SUP incseq_def le_fun_def image_comp cong: measurable_cong) lemma (in sigma_finite_measure) nn_integral_fst: assumes f: "f \ borel_measurable (M1 \\<^sub>M M)" shows "(\\<^sup>+ x. \\<^sup>+ y. f (x, y) \M \M1) = integral\<^sup>N (M1 \\<^sub>M M) f" (is "?I f = _") using f proof induct case (cong u v) then have "?I u = ?I v" by (intro nn_integral_cong) (auto simp: space_pair_measure) with cong show ?case by (simp cong: nn_integral_cong) qed (simp_all add: emeasure_pair_measure nn_integral_cmult nn_integral_add nn_integral_monotone_convergence_SUP measurable_compose_Pair1 borel_measurable_nn_integral_fst nn_integral_mono incseq_def le_fun_def image_comp cong: nn_integral_cong) lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]: "case_prod f \ borel_measurable (N \\<^sub>M M) \ (\x. \\<^sup>+ y. f x y \M) \ borel_measurable N" using borel_measurable_nn_integral_fst[of "case_prod f" N] by simp proposition (in pair_sigma_finite) nn_integral_snd: assumes f[measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = integral\<^sup>N (M1 \\<^sub>M M2) f" proof - note measurable_pair_swap[OF f] from M1.nn_integral_fst[OF this] have "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = (\\<^sup>+ (x, y). f (y, x) \(M2 \\<^sub>M M1))" by simp also have "(\\<^sup>+ (x, y). f (y, x) \(M2 \\<^sub>M M1)) = integral\<^sup>N (M1 \\<^sub>M M2) f" by (subst distr_pair_swap) (auto simp add: nn_integral_distr intro!: nn_integral_cong) finally show ?thesis . qed theorem (in pair_sigma_finite) Fubini: assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = (\\<^sup>+ x. (\\<^sup>+ y. f (x, y) \M2) \M1)" unfolding nn_integral_snd[OF assms] M2.nn_integral_fst[OF assms] .. theorem (in pair_sigma_finite) Fubini': assumes f: "case_prod f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+ y. (\\<^sup>+ x. f x y \M1) \M2) = (\\<^sup>+ x. (\\<^sup>+ y. f x y \M2) \M1)" using Fubini[OF f] by simp subsection \Products on counting spaces, densities and distributions\ proposition sigma_prod: assumes X_cover: "\E\A. countable E \ X = \E" and A: "A \ Pow X" assumes Y_cover: "\E\B. countable E \ Y = \E" and B: "B \ Pow Y" shows "sigma X A \\<^sub>M sigma Y B = sigma (X \ Y) {a \ b | a b. a \ A \ b \ B}" (is "?P = ?S") proof (rule measure_eqI) have [simp]: "snd \ X \ Y \ Y" "fst \ X \ Y \ X" by auto let ?XY = "{{fst -` a \ X \ Y | a. a \ A}, {snd -` b \ X \ Y | b. b \ B}}" have "sets ?P = sets (SUP xy\?XY. sigma (X \ Y) xy)" by (simp add: vimage_algebra_sigma sets_pair_eq_sets_fst_snd A B) also have "\ = sets (sigma (X \ Y) (\?XY))" by (intro Sup_sigma arg_cong[where f=sets]) auto also have "\ = sets ?S" proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) show "\?XY \ Pow (X \ Y)" "{a \ b |a b. a \ A \ b \ B} \ Pow (X \ Y)" using A B by auto next interpret XY: sigma_algebra "X \ Y" "sigma_sets (X \ Y) {a \ b |a b. a \ A \ b \ B}" using A B by (intro sigma_algebra_sigma_sets) auto fix Z assume "Z \ \?XY" then show "Z \ sigma_sets (X \ Y) {a \ b |a b. a \ A \ b \ B}" proof safe fix a assume "a \ A" from Y_cover obtain E where E: "E \ B" "countable E" and "Y = \E" by auto with \a \ A\ A have eq: "fst -` a \ X \ Y = (\e\E. a \ e)" by auto show "fst -` a \ X \ Y \ sigma_sets (X \ Y) {a \ b |a b. a \ A \ b \ B}" using \a \ A\ E unfolding eq by (auto intro!: XY.countable_UN') next fix b assume "b \ B" from X_cover obtain E where E: "E \ A" "countable E" and "X = \E" by auto with \b \ B\ B have eq: "snd -` b \ X \ Y = (\e\E. e \ b)" by auto show "snd -` b \ X \ Y \ sigma_sets (X \ Y) {a \ b |a b. a \ A \ b \ B}" using \b \ B\ E unfolding eq by (auto intro!: XY.countable_UN') qed next fix Z assume "Z \ {a \ b |a b. a \ A \ b \ B}" then obtain a b where "Z = a \ b" and ab: "a \ A" "b \ B" by auto then have Z: "Z = (fst -` a \ X \ Y) \ (snd -` b \ X \ Y)" using A B by auto interpret XY: sigma_algebra "X \ Y" "sigma_sets (X \ Y) (\?XY)" by (intro sigma_algebra_sigma_sets) auto show "Z \ sigma_sets (X \ Y) (\?XY)" unfolding Z by (rule XY.Int) (blast intro: ab)+ qed finally show "sets ?P = sets ?S" . next interpret finite_measure "sigma X A" for X A proof qed (simp add: emeasure_sigma) fix A assume "A \ sets ?P" then show "emeasure ?P A = emeasure ?S A" by (simp add: emeasure_pair_measure_alt emeasure_sigma) qed lemma sigma_sets_pair_measure_generator_finite: assumes "finite A" and "finite B" shows "sigma_sets (A \ B) { a \ b | a b. a \ A \ b \ B} = Pow (A \ B)" (is "sigma_sets ?prod ?sets = _") proof safe have fin: "finite (A \ B)" using assms by (rule finite_cartesian_product) fix x assume subset: "x \ A \ B" hence "finite x" using fin by (rule finite_subset) from this subset show "x \ sigma_sets ?prod ?sets" proof (induct x) case empty show ?case by (rule sigma_sets.Empty) next case (insert a x) hence "{a} \ sigma_sets ?prod ?sets" by auto moreover have "x \ sigma_sets ?prod ?sets" using insert by auto ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) qed next fix x a b assume "x \ sigma_sets ?prod ?sets" and "(a, b) \ x" from sigma_sets_into_sp[OF _ this(1)] this(2) show "a \ A" and "b \ B" by auto qed proposition sets_pair_eq: assumes Ea: "Ea \ Pow (space A)" "sets A = sigma_sets (space A) Ea" and Ca: "countable Ca" "Ca \ Ea" "\Ca = space A" and Eb: "Eb \ Pow (space B)" "sets B = sigma_sets (space B) Eb" and Cb: "countable Cb" "Cb \ Eb" "\Cb = space B" shows "sets (A \\<^sub>M B) = sets (sigma (space A \ space B) { a \ b | a b. a \ Ea \ b \ Eb })" (is "_ = sets (sigma ?\ ?E)") proof show "sets (sigma ?\ ?E) \ sets (A \\<^sub>M B)" using Ea(1) Eb(1) by (subst sigma_le_sets) (auto simp: Ea(2) Eb(2)) have "?E \ Pow ?\" using Ea(1) Eb(1) by auto then have E: "a \ Ea \ b \ Eb \ a \ b \ sets (sigma ?\ ?E)" for a b by auto have "sets (A \\<^sub>M B) \ sets (Sup {vimage_algebra ?\ fst A, vimage_algebra ?\ snd B})" unfolding sets_pair_eq_sets_fst_snd .. also have "vimage_algebra ?\ fst A = vimage_algebra ?\ fst (sigma (space A) Ea)" by (intro vimage_algebra_cong[OF refl refl]) (simp add: Ea) also have "\ = sigma ?\ {fst -` A \ ?\ |A. A \ Ea}" by (intro Ea vimage_algebra_sigma) auto also have "vimage_algebra ?\ snd B = vimage_algebra ?\ snd (sigma (space B) Eb)" by (intro vimage_algebra_cong[OF refl refl]) (simp add: Eb) also have "\ = sigma ?\ {snd -` A \ ?\ |A. A \ Eb}" by (intro Eb vimage_algebra_sigma) auto also have "{sigma ?\ {fst -` Aa \ ?\ |Aa. Aa \ Ea}, sigma ?\ {snd -` Aa \ ?\ |Aa. Aa \ Eb}} = sigma ?\ ` {{fst -` Aa \ ?\ |Aa. Aa \ Ea}, {snd -` Aa \ ?\ |Aa. Aa \ Eb}}" by auto also have "sets (SUP S\{{fst -` Aa \ ?\ |Aa. Aa \ Ea}, {snd -` Aa \ ?\ |Aa. Aa \ Eb}}. sigma ?\ S) = sets (sigma ?\ (\{{fst -` Aa \ ?\ |Aa. Aa \ Ea}, {snd -` Aa \ ?\ |Aa. Aa \ Eb}}))" using Ea(1) Eb(1) by (intro sets_Sup_sigma) auto also have "\ \ sets (sigma ?\ ?E)" proof (subst sigma_le_sets, safe intro!: space_in_measure_of) fix a assume "a \ Ea" then have "fst -` a \ ?\ = (\b\Cb. a \ b)" using Cb(3)[symmetric] Ea(1) by auto then show "fst -` a \ ?\ \ sets (sigma ?\ ?E)" using Cb \a \ Ea\ by (auto intro!: sets.countable_UN' E) next fix b assume "b \ Eb" then have "snd -` b \ ?\ = (\a\Ca. a \ b)" using Ca(3)[symmetric] Eb(1) by auto then show "snd -` b \ ?\ \ sets (sigma ?\ ?E)" using Ca \b \ Eb\ by (auto intro!: sets.countable_UN' E) qed finally show "sets (A \\<^sub>M B) \ sets (sigma ?\ ?E)" . qed proposition borel_prod: "(borel \\<^sub>M borel) = (borel :: ('a::second_countable_topology \ 'b::second_countable_topology) measure)" (is "?P = ?B") proof - have "?B = sigma UNIV {A \ B | A B. open A \ open B}" by (rule second_countable_borel_measurable[OF open_prod_generated]) also have "\ = ?P" unfolding borel_def by (subst sigma_prod) (auto intro!: exI[of _ "{UNIV}"]) finally show ?thesis .. qed proposition pair_measure_count_space: assumes A: "finite A" and B: "finite B" shows "count_space A \\<^sub>M count_space B = count_space (A \ B)" (is "?P = ?C") proof (rule measure_eqI) interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact interpret P: pair_sigma_finite "count_space A" "count_space B" .. show eq: "sets ?P = sets ?C" by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B) fix X assume X: "X \ sets ?P" with eq have X_subset: "X \ A \ B" by simp with A B have fin_Pair: "\x. finite (Pair x -` X)" by (intro finite_subset[OF _ B]) auto have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B) have card: "0 < card (Pair a -` X)" if "(a, b) \ X" for a b using card_gt_0_iff fin_Pair that by auto then have "emeasure ?P X = \\<^sup>+ x. emeasure (count_space B) (Pair x -` X) \count_space A" by (simp add: B.emeasure_pair_measure_alt X) also have "... = emeasure ?C X" apply (subst emeasure_count_space) using card X_subset A fin_Pair fin_X apply (auto simp add: nn_integral_count_space of_nat_sum[symmetric] card_SigmaI[symmetric] simp del: card_SigmaI intro!: arg_cong[where f=card]) done finally show "emeasure ?P X = emeasure ?C X" . qed lemma emeasure_prod_count_space: assumes A: "A \ sets (count_space UNIV \\<^sub>M M)" (is "A \ sets (?A \\<^sub>M ?B)") shows "emeasure (?A \\<^sub>M ?B) A = (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \?B \?A)" by (rule emeasure_measure_of[OF pair_measure_def]) (auto simp: countably_additive_def positive_def suminf_indicator A nn_integral_suminf[symmetric] dest: sets.sets_into_space) lemma emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \\<^sub>M count_space UNIV) {x} = 1" proof - have [simp]: "\a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ennreal)" by (auto split: split_indicator) show ?thesis by (cases x) (auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair) qed lemma emeasure_count_space_prod_eq: fixes A :: "('a \ 'b) set" assumes A: "A \ sets (count_space UNIV \\<^sub>M count_space UNIV)" (is "A \ sets (?A \\<^sub>M ?B)") shows "emeasure (?A \\<^sub>M ?B) A = emeasure (count_space UNIV) A" proof - { fix A :: "('a \ 'b) set" assume "countable A" then have "emeasure (?A \\<^sub>M ?B) (\a\A. {a}) = (\\<^sup>+a. emeasure (?A \\<^sub>M ?B) {a} \count_space A)" by (intro emeasure_UN_countable) (auto simp: sets_Pair disjoint_family_on_def) also have "\ = (\\<^sup>+a. indicator A a \count_space UNIV)" by (subst nn_integral_count_space_indicator) auto finally have "emeasure (?A \\<^sub>M ?B) A = emeasure (count_space UNIV) A" by simp } note * = this show ?thesis proof cases assume "finite A" then show ?thesis by (intro * countable_finite) next assume "infinite A" then obtain C where "countable C" and "infinite C" and "C \ A" by (auto dest: infinite_countable_subset') with A have "emeasure (?A \\<^sub>M ?B) C \ emeasure (?A \\<^sub>M ?B) A" by (intro emeasure_mono) auto also have "emeasure (?A \\<^sub>M ?B) C = emeasure (count_space UNIV) C" using \countable C\ by (rule *) finally show ?thesis using \infinite C\ \infinite A\ by (simp add: top_unique) qed qed lemma nn_integral_count_space_prod_eq: "nn_integral (count_space UNIV \\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f" (is "nn_integral ?P f = _") proof cases assume cntbl: "countable {x. f x \ 0}" have [simp]: "\x. card ({x} \ {x. f x \ 0}) = (indicator {x. f x \ 0} x::ennreal)" by (auto split: split_indicator) have [measurable]: "\y. (\x. indicator {y} x) \ borel_measurable ?P" by (rule measurable_discrete_difference[of "\x. 0" _ borel "{y}" "\x. indicator {y} x" for y]) (auto intro: sets_Pair) have "(\\<^sup>+x. f x \?P) = (\\<^sup>+x. \\<^sup>+x'. f x * indicator {x} x' \count_space {x. f x \ 0} \?P)" by (auto simp add: nn_integral_cmult nn_integral_indicator' intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+x. \\<^sup>+x'. f x' * indicator {x'} x \count_space {x. f x \ 0} \?P)" by (auto intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+x'. \\<^sup>+x. f x' * indicator {x'} x \?P \count_space {x. f x \ 0})" by (intro nn_integral_count_space_nn_integral cntbl) auto also have "\ = (\\<^sup>+x'. f x' \count_space {x. f x \ 0})" by (intro nn_integral_cong) (auto simp: nn_integral_cmult sets_Pair) finally show ?thesis by (auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) next { fix x assume "f x \ 0" then have "(\r\0. 0 < r \ f x = ennreal r) \ f x = \" by (cases "f x" rule: ennreal_cases) (auto simp: less_le) then have "\n. ennreal (1 / real (Suc n)) \ f x" by (auto elim!: nat_approx_posE intro!: less_imp_le) } note * = this assume cntbl: "uncountable {x. f x \ 0}" also have "{x. f x \ 0} = (\n. {x. 1/Suc n \ f x})" using * by auto finally obtain n where "infinite {x. 1/Suc n \ f x}" by (meson countableI_type countable_UN uncountable_infinite) then obtain C where C: "C \ {x. 1/Suc n \ f x}" and "countable C" "infinite C" by (metis infinite_countable_subset') have [measurable]: "C \ sets ?P" using sets.countable[OF _ \countable C\, of ?P] by (auto simp: sets_Pair) have "(\\<^sup>+x. ennreal (1/Suc n) * indicator C x \?P) \ nn_integral ?P f" using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) moreover have "(\\<^sup>+x. ennreal (1/Suc n) * indicator C x \?P) = \" using \infinite C\ by (simp add: nn_integral_cmult emeasure_count_space_prod_eq ennreal_mult_top) moreover have "(\\<^sup>+x. ennreal (1/Suc n) * indicator C x \count_space UNIV) \ nn_integral (count_space UNIV) f" using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric]) moreover have "(\\<^sup>+x. ennreal (1/Suc n) * indicator C x \count_space UNIV) = \" using \infinite C\ by (simp add: nn_integral_cmult ennreal_mult_top) ultimately show ?thesis by (simp add: top_unique) qed theorem pair_measure_density: assumes f: "f \ borel_measurable M1" assumes g: "g \ borel_measurable M2" assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)" shows "density M1 f \\<^sub>M density M2 g = density (M1 \\<^sub>M M2) (\(x,y). f x * g y)" (is "?L = ?R") proof (rule measure_eqI) interpret M2: sigma_finite_measure M2 by fact interpret D2: sigma_finite_measure "density M2 g" by fact fix A assume A: "A \ sets ?L" with f g have "(\\<^sup>+ x. f x * \\<^sup>+ y. g y * indicator A (x, y) \M2 \M1) = (\\<^sup>+ x. \\<^sup>+ y. f x * g y * indicator A (x, y) \M2 \M1)" by (intro nn_integral_cong_AE) (auto simp add: nn_integral_cmult[symmetric] ac_simps) with A f g show "emeasure ?L A = emeasure ?R A" by (simp add: D2.emeasure_pair_measure emeasure_density nn_integral_density M2.nn_integral_fst[symmetric] cong: nn_integral_cong) qed simp lemma sigma_finite_measure_distr: assumes "sigma_finite_measure (distr M N f)" and f: "f \ measurable M N" shows "sigma_finite_measure M" proof - interpret sigma_finite_measure "distr M N f" by fact - from sigma_finite_countable guess A .. note A = this + obtain A where A: "countable A" "A \ sets (distr M N f)" + "\ A = space (distr M N f)" "\a\A. emeasure (distr M N f) a \ \" + using sigma_finite_countable by auto show ?thesis proof show "\A. countable A \ A \ sets M \ \A = space M \ (\a\A. emeasure M a \ \)" using A f by (intro exI[of _ "(\a. f -` a \ space M) ` A"]) (auto simp: emeasure_distr set_eq_iff subset_eq intro: measurable_space) qed qed lemma pair_measure_distr: assumes f: "f \ measurable M S" and g: "g \ measurable N T" assumes "sigma_finite_measure (distr N T g)" shows "distr M S f \\<^sub>M distr N T g = distr (M \\<^sub>M N) (S \\<^sub>M T) (\(x, y). (f x, g y))" (is "?P = ?D") proof (rule measure_eqI) interpret T: sigma_finite_measure "distr N T g" by fact interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+ fix A assume A: "A \ sets ?P" with f g show "emeasure ?P A = emeasure ?D A" by (auto simp add: N.emeasure_pair_measure_alt space_pair_measure emeasure_distr T.emeasure_pair_measure_alt nn_integral_distr intro!: nn_integral_cong arg_cong[where f="emeasure N"]) qed simp lemma pair_measure_eqI: assumes "sigma_finite_measure M1" "sigma_finite_measure M2" assumes sets: "sets (M1 \\<^sub>M M2) = sets M" assumes emeasure: "\A B. A \ sets M1 \ B \ sets M2 \ emeasure M1 A * emeasure M2 B = emeasure M (A \ B)" shows "M1 \\<^sub>M M2 = M" proof - interpret M1: sigma_finite_measure M1 by fact interpret M2: sigma_finite_measure M2 by fact interpret pair_sigma_finite M1 M2 .. - from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this let ?E = "{a \ b |a b. a \ sets M1 \ b \ sets M2}" let ?P = "M1 \\<^sub>M M2" + obtain F :: "nat \ ('a \ 'b) set" where F: + "range F \ ?E" "incseq F" "\ (range F) = space M1 \ space M2" "\i. emeasure ?P (F i) \ \" + using sigma_finite_up_in_pair_measure_generator + by blast show ?thesis proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) show "?E \ Pow (space ?P)" using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure) show "sets ?P = sigma_sets (space ?P) ?E" by (simp add: sets_pair_measure space_pair_measure) then show "sets M = sigma_sets (space ?P) ?E" using sets[symmetric] by simp next show "range F \ ?E" "(\i. F i) = space ?P" "\i. emeasure ?P (F i) \ \" using F by (auto simp: space_pair_measure) next fix X assume "X \ ?E" then obtain A B where X[simp]: "X = A \ B" and A: "A \ sets M1" and B: "B \ sets M2" by auto then have "emeasure ?P X = emeasure M1 A * emeasure M2 B" by (simp add: M2.emeasure_pair_measure_Times) also have "\ = emeasure M (A \ B)" using A B emeasure by auto finally show "emeasure ?P X = emeasure M X" by simp qed qed lemma sets_pair_countable: assumes "countable S1" "countable S2" assumes M: "sets M = Pow S1" and N: "sets N = Pow S2" shows "sets (M \\<^sub>M N) = Pow (S1 \ S2)" proof auto fix x a b assume x: "x \ sets (M \\<^sub>M N)" "(a, b) \ x" from sets.sets_into_space[OF x(1)] x(2) sets_eq_imp_space_eq[of N "count_space S2"] sets_eq_imp_space_eq[of M "count_space S1"] M N show "a \ S1" "b \ S2" by (auto simp: space_pair_measure) next fix X assume X: "X \ S1 \ S2" then have "countable X" by (metis countable_subset \countable S1\ \countable S2\ countable_SIGMA) have "X = (\(a, b)\X. {a} \ {b})" by auto also have "\ \ sets (M \\<^sub>M N)" using X by (safe intro!: sets.countable_UN' \countable X\ subsetI pair_measureI) (auto simp: M N) finally show "X \ sets (M \\<^sub>M N)" . qed lemma pair_measure_countable: assumes "countable S1" "countable S2" shows "count_space S1 \\<^sub>M count_space S2 = count_space (S1 \ S2)" proof (rule pair_measure_eqI) show "sigma_finite_measure (count_space S1)" "sigma_finite_measure (count_space S2)" using assms by (auto intro!: sigma_finite_measure_count_space_countable) show "sets (count_space S1 \\<^sub>M count_space S2) = sets (count_space (S1 \ S2))" by (subst sets_pair_countable[OF assms]) auto next fix A B assume "A \ sets (count_space S1)" "B \ sets (count_space S2)" then show "emeasure (count_space S1) A * emeasure (count_space S2) B = emeasure (count_space (S1 \ S2)) (A \ B)" by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff ennreal_mult_top ennreal_top_mult) qed proposition nn_integral_fst_count_space: "(\\<^sup>+ x. \\<^sup>+ y. f (x, y) \count_space UNIV \count_space UNIV) = integral\<^sup>N (count_space UNIV) f" (is "?lhs = ?rhs") proof(cases) assume *: "countable {xy. f xy \ 0}" let ?A = "fst ` {xy. f xy \ 0}" let ?B = "snd ` {xy. f xy \ 0}" from * have [simp]: "countable ?A" "countable ?B" by(rule countable_image)+ have "?lhs = (\\<^sup>+ x. \\<^sup>+ y. f (x, y) \count_space UNIV \count_space ?A)" by(rule nn_integral_count_space_eq) (auto simp add: nn_integral_0_iff_AE AE_count_space not_le intro: rev_image_eqI) also have "\ = (\\<^sup>+ x. \\<^sup>+ y. f (x, y) \count_space ?B \count_space ?A)" by(intro nn_integral_count_space_eq nn_integral_cong)(auto intro: rev_image_eqI) also have "\ = (\\<^sup>+ xy. f xy \count_space (?A \ ?B))" by(subst sigma_finite_measure.nn_integral_fst) (simp_all add: sigma_finite_measure_count_space_countable pair_measure_countable) also have "\ = ?rhs" by(rule nn_integral_count_space_eq)(auto intro: rev_image_eqI) finally show ?thesis . next { fix xy assume "f xy \ 0" then have "(\r\0. 0 < r \ f xy = ennreal r) \ f xy = \" by (cases "f xy" rule: ennreal_cases) (auto simp: less_le) then have "\n. ennreal (1 / real (Suc n)) \ f xy" by (auto elim!: nat_approx_posE intro!: less_imp_le) } note * = this assume cntbl: "uncountable {xy. f xy \ 0}" also have "{xy. f xy \ 0} = (\n. {xy. 1/Suc n \ f xy})" using * by auto finally obtain n where "infinite {xy. 1/Suc n \ f xy}" by (meson countableI_type countable_UN uncountable_infinite) then obtain C where C: "C \ {xy. 1/Suc n \ f xy}" and "countable C" "infinite C" by (metis infinite_countable_subset') have "\ = (\\<^sup>+ xy. ennreal (1 / Suc n) * indicator C xy \count_space UNIV)" using \infinite C\ by(simp add: nn_integral_cmult ennreal_mult_top) also have "\ \ ?rhs" using C by(intro nn_integral_mono)(auto split: split_indicator) finally have "?rhs = \" by (simp add: top_unique) moreover have "?lhs = \" proof(cases "finite (fst ` C)") case True then obtain x C' where x: "x \ fst ` C" and C': "C' = fst -` {x} \ C" and "infinite C'" using \infinite C\ by(auto elim!: inf_img_fin_domE') from x C C' have **: "C' \ {xy. 1 / Suc n \ f xy}" by auto from C' \infinite C'\ have "infinite (snd ` C')" by(auto dest!: finite_imageD simp add: inj_on_def) then have "\ = (\\<^sup>+ y. ennreal (1 / Suc n) * indicator (snd ` C') y \count_space UNIV)" by(simp add: nn_integral_cmult ennreal_mult_top) also have "\ = (\\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \count_space UNIV)" by(rule nn_integral_cong)(force split: split_indicator intro: rev_image_eqI simp add: C') also have "\ = (\\<^sup>+ x'. (\\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \count_space UNIV) * indicator {x} x' \count_space UNIV)" by(simp add: one_ereal_def[symmetric]) also have "\ \ (\\<^sup>+ x. \\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \count_space UNIV \count_space UNIV)" by(rule nn_integral_mono)(simp split: split_indicator) also have "\ \ ?lhs" using ** by(intro nn_integral_mono)(auto split: split_indicator) finally show ?thesis by (simp add: top_unique) next case False define C' where "C' = fst ` C" have "\ = \\<^sup>+ x. ennreal (1 / Suc n) * indicator C' x \count_space UNIV" using C'_def False by(simp add: nn_integral_cmult ennreal_mult_top) also have "\ = \\<^sup>+ x. \\<^sup>+ y. ennreal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \ C} y \count_space UNIV \count_space UNIV" by(auto simp add: one_ereal_def[symmetric] max_def intro: nn_integral_cong) also have "\ \ \\<^sup>+ x. \\<^sup>+ y. ennreal (1 / Suc n) * indicator C (x, y) \count_space UNIV \count_space UNIV" by(intro nn_integral_mono)(auto simp add: C'_def split: split_indicator intro: someI) also have "\ \ ?lhs" using C by(intro nn_integral_mono)(auto split: split_indicator) finally show ?thesis by (simp add: top_unique) qed ultimately show ?thesis by simp qed proposition nn_integral_snd_count_space: "(\\<^sup>+ y. \\<^sup>+ x. f (x, y) \count_space UNIV \count_space UNIV) = integral\<^sup>N (count_space UNIV) f" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ y. \\<^sup>+ x. (\(y, x). f (x, y)) (y, x) \count_space UNIV \count_space UNIV)" by(simp) also have "\ = \\<^sup>+ yx. (\(y, x). f (x, y)) yx \count_space UNIV" by(rule nn_integral_fst_count_space) also have "\ = \\<^sup>+ xy. f xy \count_space ((\(x, y). (y, x)) ` UNIV)" by(subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric]) (simp_all add: inj_on_def split_def) also have "\ = ?rhs" by(rule nn_integral_count_space_eq) auto finally show ?thesis . qed lemma measurable_pair_measure_countable1: assumes "countable A" and [measurable]: "\x. x \ A \ (\y. f (x, y)) \ measurable N K" shows "f \ measurable (count_space A \\<^sub>M N) K" using _ _ assms(1) by(rule measurable_compose_countable'[where f="\a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all subsection \Product of Borel spaces\ theorem borel_Times: fixes A :: "'a::topological_space set" and B :: "'b::topological_space set" assumes A: "A \ sets borel" and B: "B \ sets borel" shows "A \ B \ sets borel" proof - have "A \ B = (A\UNIV) \ (UNIV \ B)" by auto moreover { have "A \ sigma_sets UNIV {S. open S}" using A by (simp add: sets_borel) then have "A\UNIV \ sets borel" proof (induct A) case (Basic S) then show ?case by (auto intro!: borel_open open_Times) next case (Compl A) moreover have *: "(UNIV - A) \ UNIV = UNIV - (A \ UNIV)" by auto ultimately show ?case unfolding * by auto next case (Union A) moreover have *: "(\(A ` UNIV)) \ UNIV = \((\i. A i \ UNIV) ` UNIV)" by auto ultimately show ?case unfolding * by auto qed simp } moreover { have "B \ sigma_sets UNIV {S. open S}" using B by (simp add: sets_borel) then have "UNIV\B \ sets borel" proof (induct B) case (Basic S) then show ?case by (auto intro!: borel_open open_Times) next case (Compl B) moreover have *: "UNIV \ (UNIV - B) = UNIV - (UNIV \ B)" by auto ultimately show ?case unfolding * by auto next case (Union B) moreover have *: "UNIV \ (\(B ` UNIV)) = \((\i. UNIV \ B i) ` UNIV)" by auto ultimately show ?case unfolding * by auto qed simp } ultimately show ?thesis by auto qed lemma finite_measure_pair_measure: assumes "finite_measure M" "finite_measure N" shows "finite_measure (N \\<^sub>M M)" proof (rule finite_measureI) interpret M: finite_measure M by fact interpret N: finite_measure N by fact show "emeasure (N \\<^sub>M M) (space (N \\<^sub>M M)) \ \" by (auto simp: space_pair_measure M.emeasure_pair_measure_Times ennreal_mult_eq_top_iff) qed end diff --git a/src/HOL/Analysis/Bochner_Integration.thy b/src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy +++ b/src/HOL/Analysis/Bochner_Integration.thy @@ -1,3428 +1,3432 @@ (* Title: HOL/Analysis/Bochner_Integration.thy Author: Johannes Hölzl, TU München *) section \Bochner Integration for Vector-Valued Functions\ theory Bochner_Integration imports Finite_Product_Measure begin text \ In the following development of the Bochner integral we use second countable topologies instead of separable spaces. A second countable topology is also separable. \ proposition borel_measurable_implies_sequence_metric: fixes f :: "'a \ 'b :: {metric_space, second_countable_topology}" assumes [measurable]: "f \ borel_measurable M" shows "\F. (\i. simple_function M (F i)) \ (\x\space M. (\i. F i x) \ f x) \ (\i. \x\space M. dist (F i x) z \ 2 * dist (f x) z)" proof - obtain D :: "'b set" where "countable D" and D: "\X. open X \ X \ {} \ \d\D. d \ X" by (erule countable_dense_setE) define e where "e = from_nat_into D" { fix n x obtain d where "d \ D" and d: "d \ ball x (1 / Suc n)" using D[of "ball x (1 / Suc n)"] by auto from \d \ D\ D[of UNIV] \countable D\ obtain i where "d = e i" unfolding e_def by (auto dest: from_nat_into_surj) with d have "\i. dist x (e i) < 1 / Suc n" by auto } note e = this define A where [abs_def]: "A m n = {x\space M. dist (f x) (e n) < 1 / (Suc m) \ 1 / (Suc m) \ dist (f x) z}" for m n define B where [abs_def]: "B m = disjointed (A m)" for m define m where [abs_def]: "m N x = Max {m. m \ N \ x \ (\n\N. B m n)}" for N x define F where [abs_def]: "F N x = (if (\m\N. x \ (\n\N. B m n)) \ (\n\N. x \ B (m N x) n) then e (LEAST n. x \ B (m N x) n) else z)" for N x have B_imp_A[intro, simp]: "\x m n. x \ B m n \ x \ A m n" using disjointed_subset[of "A m" for m] unfolding B_def by auto { fix m have "\n. A m n \ sets M" by (auto simp: A_def) then have "\n. B m n \ sets M" using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) } note this[measurable] { fix N i x assume "\m\N. x \ (\n\N. B m n)" then have "m N x \ {m::nat. m \ N \ x \ (\n\N. B m n)}" unfolding m_def by (intro Max_in) auto then have "m N x \ N" "\n\N. x \ B (m N x) n" by auto } note m = this { fix j N i x assume "j \ N" "i \ N" "x \ B j i" then have "j \ m N x" unfolding m_def by (intro Max_ge) auto } note m_upper = this show ?thesis unfolding simple_function_def proof (safe intro!: exI[of _ F]) have [measurable]: "\i. F i \ borel_measurable M" unfolding F_def m_def by measurable show "\x i. F i -` {x} \ space M \ sets M" by measurable { fix i { fix n x assume "x \ B (m i x) n" then have "(LEAST n. x \ B (m i x) n) \ n" by (intro Least_le) also assume "n \ i" finally have "(LEAST n. x \ B (m i x) n) \ i" . } then have "F i ` space M \ {z} \ e ` {.. i}" by (auto simp: F_def) then show "finite (F i ` space M)" by (rule finite_subset) auto } { fix N i n x assume "i \ N" "n \ N" "x \ B i n" then have 1: "\m\N. x \ (\n\N. B m n)" by auto from m[OF this] obtain n where n: "m N x \ N" "n \ N" "x \ B (m N x) n" by auto moreover define L where "L = (LEAST n. x \ B (m N x) n)" have "dist (f x) (e L) < 1 / Suc (m N x)" proof - have "x \ B (m N x) L" using n(3) unfolding L_def by (rule LeastI) then have "x \ A (m N x) L" by auto then show ?thesis unfolding A_def by simp qed ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)" by (auto simp add: F_def L_def) } note * = this fix x assume "x \ space M" show "(\i. F i x) \ f x" proof cases assume "f x = z" then have "\i n. x \ A i n" unfolding A_def by auto then have "\i. F i x = z" by (auto simp: F_def) then show ?thesis using \f x = z\ by auto next assume "f x \ z" show ?thesis proof (rule tendstoI) fix e :: real assume "0 < e" with \f x \ z\ obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" by (metis dist_nz order_less_trans neq_iff nat_approx_posE) with \x\space M\ \f x \ z\ have "x \ (\i. B n i)" unfolding A_def B_def UN_disjointed_eq using e by auto then obtain i where i: "x \ B n i" by auto show "eventually (\i. dist (F i x) (f x) < e) sequentially" using eventually_ge_at_top[of "max n i"] proof eventually_elim fix j assume j: "max n i \ j" with i have "dist (f x) (F j x) < 1 / Suc (m j x)" by (intro *[OF _ _ i]) auto also have "\ \ 1 / Suc n" using j m_upper[OF _ _ i] by (auto simp: field_simps) also note \1 / Suc n < e\ finally show "dist (F j x) (f x) < e" by (simp add: less_imp_le dist_commute) qed qed qed fix i { fix n m assume "x \ A n m" then have "dist (e m) (f x) + dist (f x) z \ 2 * dist (f x) z" unfolding A_def by (auto simp: dist_commute) also have "dist (e m) z \ dist (e m) (f x) + dist (f x) z" by (rule dist_triangle) finally (xtrans) have "dist (e m) z \ 2 * dist (f x) z" . } then show "dist (F i x) z \ 2 * dist (f x) z" unfolding F_def apply auto apply (rule LeastI2) apply auto done qed qed lemma fixes f :: "'a \ 'b::semiring_1" assumes "finite A" shows sum_mult_indicator[simp]: "(\x \ A. f x * indicator (B x) (g x)) = (\x\{x\A. g x \ B x}. f x)" and sum_indicator_mult[simp]: "(\x \ A. indicator (B x) (g x) * f x) = (\x\{x\A. g x \ B x}. f x)" unfolding indicator_def using assms by (auto intro!: sum.mono_neutral_cong_right split: if_split_asm) lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]: fixes P :: "('a \ real) \ bool" assumes u: "u \ borel_measurable M" "\x. 0 \ u x" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ (\x. x \ space M \ (\i. U i x) \ u x) \ P u" shows "P u" proof - have "(\x. ennreal (u x)) \ borel_measurable M" using u by auto from borel_measurable_implies_simple_function_sequence'[OF this] obtain U where U: "\i. simple_function M (U i)" "incseq U" "\i x. U i x < top" and sup: "\x. (SUP i. U i x) = ennreal (u x)" by blast define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x then have U'_sf[measurable]: "\i. simple_function M (U' i)" using U by (auto intro!: simple_function_compose1[where g=enn2real]) show "P u" proof (rule seq) show U': "U' i \ borel_measurable M" "\x. 0 \ U' i x" for i using U by (auto intro: borel_measurable_simple_function intro!: borel_measurable_enn2real borel_measurable_times simp: U'_def zero_le_mult_iff) show "incseq U'" using U(2,3) by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono) fix x assume x: "x \ space M" have "(\i. U i x) \ (SUP i. U i x)" using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def) moreover have "(\i. U i x) = (\i. ennreal (U' i x))" using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute) moreover have "(SUP i. U i x) = ennreal (u x)" using sup u(2) by (simp add: max_def) ultimately show "(\i. U' i x) \ u x" using u U' by simp next fix i have "U' i ` space M \ enn2real ` (U i ` space M)" "finite (U i ` space M)" unfolding U'_def using U(1) by (auto dest: simple_functionD) then have fin: "finite (U' i ` space M)" by (metis finite_subset finite_imageI) moreover have "\z. {y. U' i z = y \ y \ U' i ` space M \ z \ space M} = (if z \ space M then {U' i z} else {})" by auto ultimately have U': "(\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z) = U' i" by (simp add: U'_def fun_eq_iff) have "\x. x \ U' i ` space M \ 0 \ x" by (auto simp: U'_def) with fin have "P (\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z)" proof induct case empty from set[of "{}"] show ?case by (simp add: indicator_def[abs_def]) next case (insert x F) from insert.prems have nonneg: "x \ 0" "\y. y \ F \ y \ 0" by simp_all hence *: "P (\xa. x * indicat_real {x' \ space M. U' i x' = x} xa)" by (intro mult set) auto have "P (\z. x * indicat_real {x' \ space M. U' i x' = x} z + (\y\F. y * indicat_real {x \ space M. U' i x = y} z))" using insert(1-3) by (intro add * sum_nonneg mult_nonneg_nonneg) (auto simp: nonneg indicator_def of_bool_def sum_nonneg_eq_0_iff) thus ?case using insert.hyps by (subst sum.insert) auto qed with U' show "P (U' i)" by simp qed qed lemma scaleR_cong_right: fixes x :: "'a :: real_vector" shows "(x \ 0 \ r = p) \ r *\<^sub>R x = p *\<^sub>R x" by (cases "x = 0") auto inductive simple_bochner_integrable :: "'a measure \ ('a \ 'b::real_vector) \ bool" for M f where "simple_function M f \ emeasure M {y\space M. f y \ 0} \ \ \ simple_bochner_integrable M f" lemma simple_bochner_integrable_compose2: assumes p_0: "p 0 0 = 0" shows "simple_bochner_integrable M f \ simple_bochner_integrable M g \ simple_bochner_integrable M (\x. p (f x) (g x))" proof (safe intro!: simple_bochner_integrable.intros elim!: simple_bochner_integrable.cases del: notI) assume sf: "simple_function M f" "simple_function M g" then show "simple_function M (\x. p (f x) (g x))" by (rule simple_function_compose2) from sf have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function) assume fin: "emeasure M {y \ space M. f y \ 0} \ \" "emeasure M {y \ space M. g y \ 0} \ \" have "emeasure M {x\space M. p (f x) (g x) \ 0} \ emeasure M ({x\space M. f x \ 0} \ {x\space M. g x \ 0})" by (intro emeasure_mono) (auto simp: p_0) also have "\ \ emeasure M {x\space M. f x \ 0} + emeasure M {x\space M. g x \ 0}" by (intro emeasure_subadditive) auto finally show "emeasure M {y \ space M. p (f y) (g y) \ 0} \ \" using fin by (auto simp: top_unique) qed lemma simple_function_finite_support: assumes f: "simple_function M f" and fin: "(\\<^sup>+x. f x \M) < \" and nn: "\x. 0 \ f x" shows "emeasure M {x\space M. f x \ 0} \ \" proof cases from f have meas[measurable]: "f \ borel_measurable M" by (rule borel_measurable_simple_function) assume non_empty: "\x\space M. f x \ 0" define m where "m = Min (f`space M - {0})" have "m \ f`space M - {0}" unfolding m_def using f non_empty by (intro Min_in) (auto simp: simple_function_def) then have m: "0 < m" using nn by (auto simp: less_le) from m have "m * emeasure M {x\space M. 0 \ f x} = (\\<^sup>+x. m * indicator {x\space M. 0 \ f x} x \M)" using f by (intro nn_integral_cmult_indicator[symmetric]) auto also have "\ \ (\\<^sup>+x. f x \M)" using AE_space proof (intro nn_integral_mono_AE, eventually_elim) fix x assume "x \ space M" with nn show "m * indicator {x \ space M. 0 \ f x} x \ f x" using f by (auto split: split_indicator simp: simple_function_def m_def) qed also note \\ < \\ finally show ?thesis using m by (auto simp: ennreal_mult_less_top) next assume "\ (\x\space M. f x \ 0)" with nn have *: "{x\space M. f x \ 0} = {}" by auto show ?thesis unfolding * by simp qed lemma simple_bochner_integrableI_bounded: assumes f: "simple_function M f" and fin: "(\\<^sup>+x. norm (f x) \M) < \" shows "simple_bochner_integrable M f" proof have "emeasure M {y \ space M. ennreal (norm (f y)) \ 0} \ \" proof (rule simple_function_finite_support) show "simple_function M (\x. ennreal (norm (f x)))" using f by (rule simple_function_compose1) show "(\\<^sup>+ y. ennreal (norm (f y)) \M) < \" by fact qed simp then show "emeasure M {y \ space M. f y \ 0} \ \" by simp qed fact definition\<^marker>\tag important\ simple_bochner_integral :: "'a measure \ ('a \ 'b::real_vector) \ 'b" where "simple_bochner_integral M f = (\y\f`space M. measure M {x\space M. f x = y} *\<^sub>R y)" proposition simple_bochner_integral_partition: assumes f: "simple_bochner_integrable M f" and g: "simple_function M g" assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" assumes v: "\x. x \ space M \ f x = v (g x)" shows "simple_bochner_integral M f = (\y\g ` space M. measure M {x\space M. g x = y} *\<^sub>R v y)" (is "_ = ?r") proof - from f g have [simp]: "finite (f`space M)" "finite (g`space M)" by (auto simp: simple_function_def elim: simple_bochner_integrable.cases) from f have [measurable]: "f \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) from g have [measurable]: "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) { fix y assume "y \ space M" then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" by (auto cong: sub simp: v[symmetric]) } note eq = this have "simple_bochner_integral M f = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} else 0) *\<^sub>R y)" unfolding simple_bochner_integral_def proof (safe intro!: sum.cong scaleR_cong_right) fix y assume y: "y \ space M" "f y \ 0" have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = {z. \x\space M. f y = f x \ z = g x}" by auto have eq:"{x \ space M. f x = f y} = (\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i})" by (auto simp: eq_commute cong: sub rev_conj_cong) have "finite (g`space M)" by simp then have "finite {z. \x\space M. f y = f x \ z = g x}" by (rule rev_finite_subset) auto moreover { fix x assume "x \ space M" "f x = f y" then have "x \ space M" "f x \ 0" using y by auto then have "emeasure M {y \ space M. g y = g x} \ emeasure M {y \ space M. f y \ 0}" by (auto intro!: emeasure_mono cong: sub) then have "emeasure M {xa \ space M. g xa = g x} < \" using f by (auto simp: simple_bochner_integrable.simps less_top) } ultimately show "measure M {x \ space M. f x = f y} = (\z\g ` space M. if \x\space M. f y = f x \ z = g x then measure M {x \ space M. g x = z} else 0)" apply (simp add: sum.If_cases eq) apply (subst measure_finite_Union[symmetric]) apply (auto simp: disjoint_family_on_def less_top) done qed also have "\ = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} *\<^sub>R y else 0))" by (auto intro!: sum.cong simp: scaleR_sum_left) also have "\ = ?r" by (subst sum.swap) (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq) finally show "simple_bochner_integral M f = ?r" . qed lemma simple_bochner_integral_add: assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f x + g x) = simple_bochner_integral M f + simple_bochner_integral M g" proof - from f g have "simple_bochner_integral M (\x. f x + g x) = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R (fst y + snd y))" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) moreover from f g have "simple_bochner_integral M f = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R fst y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) moreover from f g have "simple_bochner_integral M g = (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R snd y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) ultimately show ?thesis by (simp add: sum.distrib[symmetric] scaleR_add_right) qed lemma simple_bochner_integral_linear: assumes "linear f" assumes g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f (g x)) = f (simple_bochner_integral M g)" proof - interpret linear f by fact from g have "simple_bochner_integral M (\x. f (g x)) = (\y\g ` space M. measure M {x \ space M. g x = y} *\<^sub>R f y)" by (intro simple_bochner_integral_partition) (auto simp: simple_bochner_integrable_compose2[where p="\x y. f x"] elim: simple_bochner_integrable.cases) also have "\ = f (simple_bochner_integral M g)" by (simp add: simple_bochner_integral_def sum scale) finally show ?thesis . qed lemma simple_bochner_integral_minus: assumes f: "simple_bochner_integrable M f" shows "simple_bochner_integral M (\x. - f x) = - simple_bochner_integral M f" proof - from linear_uminus f show ?thesis by (rule simple_bochner_integral_linear) qed lemma simple_bochner_integral_diff: assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" shows "simple_bochner_integral M (\x. f x - g x) = simple_bochner_integral M f - simple_bochner_integral M g" unfolding diff_conv_add_uminus using f g by (subst simple_bochner_integral_add) (auto simp: simple_bochner_integral_minus simple_bochner_integrable_compose2[where p="\x y. - y"]) lemma simple_bochner_integral_norm_bound: assumes f: "simple_bochner_integrable M f" shows "norm (simple_bochner_integral M f) \ simple_bochner_integral M (\x. norm (f x))" proof - have "norm (simple_bochner_integral M f) \ (\y\f ` space M. norm (measure M {x \ space M. f x = y} *\<^sub>R y))" unfolding simple_bochner_integral_def by (rule norm_sum) also have "\ = (\y\f ` space M. measure M {x \ space M. f x = y} *\<^sub>R norm y)" by simp also have "\ = simple_bochner_integral M (\x. norm (f x))" using f by (intro simple_bochner_integral_partition[symmetric]) (auto intro: f simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) finally show ?thesis . qed lemma simple_bochner_integral_nonneg[simp]: fixes f :: "'a \ real" shows "(\x. 0 \ f x) \ 0 \ simple_bochner_integral M f" by (force simp add: simple_bochner_integral_def intro: sum_nonneg) lemma simple_bochner_integral_eq_nn_integral: assumes f: "simple_bochner_integrable M f" "\x. 0 \ f x" shows "simple_bochner_integral M f = (\\<^sup>+x. f x \M)" proof - { fix x y z have "(x \ 0 \ y = z) \ ennreal x * y = ennreal x * z" by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) } note ennreal_cong_mult = this have [measurable]: "f \ borel_measurable M" using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) { fix y assume y: "y \ space M" "f y \ 0" have "ennreal (measure M {x \ space M. f x = f y}) = emeasure M {x \ space M. f x = f y}" proof (rule emeasure_eq_ennreal_measure[symmetric]) have "emeasure M {x \ space M. f x = f y} \ emeasure M {x \ space M. f x \ 0}" using y by (intro emeasure_mono) auto with f show "emeasure M {x \ space M. f x = f y} \ top" by (auto simp: simple_bochner_integrable.simps top_unique) qed moreover have "{x \ space M. f x = f y} = (\x. ennreal (f x)) -` {ennreal (f y)} \ space M" using f by auto ultimately have "ennreal (measure M {x \ space M. f x = f y}) = emeasure M ((\x. ennreal (f x)) -` {ennreal (f y)} \ space M)" by simp } with f have "simple_bochner_integral M f = (\\<^sup>Sx. f x \M)" unfolding simple_integral_def by (subst simple_bochner_integral_partition[OF f(1), where g="\x. ennreal (f x)" and v=enn2real]) (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases intro!: sum.cong ennreal_cong_mult simp: ac_simps ennreal_mult simp flip: sum_ennreal) also have "\ = (\\<^sup>+x. f x \M)" using f by (intro nn_integral_eq_simple_integral[symmetric]) (auto simp: simple_function_compose1 simple_bochner_integrable.simps) finally show ?thesis . qed lemma simple_bochner_integral_bounded: fixes f :: "'a \ 'b::{real_normed_vector, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t" shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \ (\\<^sup>+ x. norm (f x - s x) \M) + (\\<^sup>+ x. norm (f x - t x) \M)" (is "ennreal (norm (?s - ?t)) \ ?S + ?T") proof - have [measurable]: "s \ borel_measurable M" "t \ borel_measurable M" using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\x. s x - t x))" using s t by (subst simple_bochner_integral_diff) auto also have "\ \ simple_bochner_integral M (\x. norm (s x - t x))" using simple_bochner_integrable_compose2[of "(-)" M "s" "t"] s t by (auto intro!: simple_bochner_integral_norm_bound) also have "\ = (\\<^sup>+x. norm (s x - t x) \M)" using simple_bochner_integrable_compose2[of "\x y. norm (x - y)" M "s" "t"] s t by (auto intro!: simple_bochner_integral_eq_nn_integral) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis (erased, opaque_lifting) add_diff_cancel_left add_diff_eq diff_add_eq order_trans norm_minus_commute norm_triangle_ineq4 order_refl) also have "\ = ?S + ?T" by (rule nn_integral_add) auto finally show ?thesis . qed inductive has_bochner_integral :: "'a measure \ ('a \ 'b) \ 'b::{real_normed_vector, second_countable_topology} \ bool" for M f x where "f \ borel_measurable M \ (\i. simple_bochner_integrable M (s i)) \ (\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0 \ (\i. simple_bochner_integral M (s i)) \ x \ has_bochner_integral M f x" lemma has_bochner_integral_cong: assumes "M = N" "\x. x \ space N \ f x = g x" "x = y" shows "has_bochner_integral M f x \ has_bochner_integral N g y" unfolding has_bochner_integral.simps assms(1,3) using assms(2) by (simp cong: measurable_cong_simp nn_integral_cong_simp) lemma has_bochner_integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. f x = g x) \ has_bochner_integral M f x \ has_bochner_integral M g x" unfolding has_bochner_integral.simps by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\x. x \ 0"] nn_integral_cong_AE) auto lemma borel_measurable_has_bochner_integral: "has_bochner_integral M f x \ f \ borel_measurable M" by (rule has_bochner_integral.cases) lemma borel_measurable_has_bochner_integral'[measurable_dest]: "has_bochner_integral M f x \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" using borel_measurable_has_bochner_integral[measurable] by measurable lemma has_bochner_integral_simple_bochner_integrable: "simple_bochner_integrable M f \ has_bochner_integral M f (simple_bochner_integral M f)" by (rule has_bochner_integral.intros[where s="\_. f"]) (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases simp: zero_ennreal_def[symmetric]) lemma has_bochner_integral_real_indicator: assumes [measurable]: "A \ sets M" and A: "emeasure M A < \" shows "has_bochner_integral M (indicator A) (measure M A)" proof - have sbi: "simple_bochner_integrable M (indicator A::'a \ real)" proof have "{y \ space M. (indicator A y::real) \ 0} = A" using sets.sets_into_space[OF \A\sets M\] by (auto split: split_indicator) then show "emeasure M {y \ space M. (indicator A y::real) \ 0} \ \" using A by auto qed (rule simple_function_indicator assms)+ moreover have "simple_bochner_integral M (indicator A) = measure M A" using simple_bochner_integral_eq_nn_integral[OF sbi] A by (simp add: ennreal_indicator emeasure_eq_ennreal_measure) ultimately show ?thesis by (metis has_bochner_integral_simple_bochner_integrable) qed lemma has_bochner_integral_add[intro]: "has_bochner_integral M f x \ has_bochner_integral M g y \ has_bochner_integral M (\x. f x + g x) (x + y)" proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) fix sf sg assume f_sf: "(\i. \\<^sup>+ x. norm (f x - sf i x) \M) \ 0" assume g_sg: "(\i. \\<^sup>+ x. norm (g x - sg i x) \M) \ 0" assume sf: "\i. simple_bochner_integrable M (sf i)" and sg: "\i. simple_bochner_integrable M (sg i)" then have [measurable]: "\i. sf i \ borel_measurable M" "\i. sg i \ borel_measurable M" by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) assume [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" show "\i. simple_bochner_integrable M (\x. sf i x + sg i x)" using sf sg by (simp add: simple_bochner_integrable_compose2) show "(\i. \\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \M) \ 0" (is "?f \ 0") proof (rule tendsto_sandwich) show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by auto show "eventually (\i. ?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) \M) + \\<^sup>+ x. (norm (g x - sg i x)) \M) sequentially" (is "eventually (\i. ?f i \ ?g i) sequentially") proof (intro always_eventually allI) fix i have "?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \M)" by (auto intro!: nn_integral_mono norm_diff_triangle_ineq simp flip: ennreal_plus) also have "\ = ?g i" by (intro nn_integral_add) auto finally show "?f i \ ?g i" . qed show "?g \ 0" using tendsto_add[OF f_sf g_sg] by simp qed qed (auto simp: simple_bochner_integral_add tendsto_add) lemma has_bochner_integral_bounded_linear: assumes "bounded_linear T" shows "has_bochner_integral M f x \ has_bochner_integral M (\x. T (f x)) (T x)" proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) interpret T: bounded_linear T by fact have [measurable]: "T \ borel_measurable borel" by (intro borel_measurable_continuous_onI T.continuous_on continuous_on_id) assume [measurable]: "f \ borel_measurable M" then show "(\x. T (f x)) \ borel_measurable M" by auto fix s assume f_s: "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" assume s: "\i. simple_bochner_integrable M (s i)" then show "\i. simple_bochner_integrable M (\x. T (s i x))" by (auto intro: simple_bochner_integrable_compose2 T.zero) have [measurable]: "\i. s i \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) obtain K where K: "K > 0" "\x i. norm (T (f x) - T (s i x)) \ norm (f x - s i x) * K" using T.pos_bounded by (auto simp: T.diff[symmetric]) show "(\i. \\<^sup>+ x. norm (T (f x) - T (s i x)) \M) \ 0" (is "?f \ 0") proof (rule tendsto_sandwich) show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by auto show "eventually (\i. ?f i \ K * (\\<^sup>+ x. norm (f x - s i x) \M)) sequentially" (is "eventually (\i. ?f i \ ?g i) sequentially") proof (intro always_eventually allI) fix i have "?f i \ (\\<^sup>+ x. ennreal K * norm (f x - s i x) \M)" using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric]) also have "\ = ?g i" using K by (intro nn_integral_cmult) auto finally show "?f i \ ?g i" . qed show "?g \ 0" using ennreal_tendsto_cmult[OF _ f_s] by simp qed assume "(\i. simple_bochner_integral M (s i)) \ x" with s show "(\i. simple_bochner_integral M (\x. T (s i x))) \ T x" by (auto intro!: T.tendsto simp: simple_bochner_integral_linear T.linear_axioms) qed lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\x. 0) 0" by (auto intro!: has_bochner_integral.intros[where s="\_ _. 0"] simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps simple_bochner_integral_def image_constant_conv) lemma has_bochner_integral_scaleR_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x *\<^sub>R c) (x *\<^sub>R c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left]) lemma has_bochner_integral_scaleR_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c *\<^sub>R f x) (c *\<^sub>R x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right]) lemma has_bochner_integral_mult_left[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x * c) (x * c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left]) lemma has_bochner_integral_mult_right[intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c * f x) (c * x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right]) lemmas has_bochner_integral_divide = has_bochner_integral_bounded_linear[OF bounded_linear_divide] lemma has_bochner_integral_divide_zero[intro]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x / c) (x / c)" using has_bochner_integral_divide by (cases "c = 0") auto lemma has_bochner_integral_inner_left[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x \ c) (x \ c)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left]) lemma has_bochner_integral_inner_right[intro]: "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c \ f x) (c \ x)" by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right]) lemmas has_bochner_integral_minus = has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] lemmas has_bochner_integral_Re = has_bochner_integral_bounded_linear[OF bounded_linear_Re] lemmas has_bochner_integral_Im = has_bochner_integral_bounded_linear[OF bounded_linear_Im] lemmas has_bochner_integral_cnj = has_bochner_integral_bounded_linear[OF bounded_linear_cnj] lemmas has_bochner_integral_of_real = has_bochner_integral_bounded_linear[OF bounded_linear_of_real] lemmas has_bochner_integral_fst = has_bochner_integral_bounded_linear[OF bounded_linear_fst] lemmas has_bochner_integral_snd = has_bochner_integral_bounded_linear[OF bounded_linear_snd] lemma has_bochner_integral_indicator: "A \ sets M \ emeasure M A < \ \ has_bochner_integral M (\x. indicator A x *\<^sub>R c) (measure M A *\<^sub>R c)" by (intro has_bochner_integral_scaleR_left has_bochner_integral_real_indicator) lemma has_bochner_integral_diff: "has_bochner_integral M f x \ has_bochner_integral M g y \ has_bochner_integral M (\x. f x - g x) (x - y)" unfolding diff_conv_add_uminus by (intro has_bochner_integral_add has_bochner_integral_minus) lemma has_bochner_integral_sum: "(\i. i \ I \ has_bochner_integral M (f i) (x i)) \ has_bochner_integral M (\x. \i\I. f i x) (\i\I. x i)" by (induct I rule: infinite_finite_induct) auto proposition has_bochner_integral_implies_finite_norm: "has_bochner_integral M f x \ (\\<^sup>+x. norm (f x) \M) < \" proof (elim has_bochner_integral.cases) fix s v assume [measurable]: "f \ borel_measurable M" and s: "\i. simple_bochner_integrable M (s i)" and lim_0: "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" from order_tendstoD[OF lim_0, of "\"] obtain i where f_s_fin: "(\\<^sup>+ x. ennreal (norm (f x - s i x)) \M) < \" by (auto simp: eventually_sequentially) have [measurable]: "\i. s i \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) define m where "m = (if space M = {} then 0 else Max ((\x. norm (s i x))`space M))" have "finite (s i ` space M)" using s by (auto simp: simple_function_def simple_bochner_integrable.simps) then have "finite (norm ` s i ` space M)" by (rule finite_imageI) then have "\x. x \ space M \ norm (s i x) \ m" "0 \ m" by (auto simp: m_def image_comp comp_def Max_ge_iff) then have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ennreal m * indicator {x\space M. s i x \ 0} x \M)" by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:) also have "\ < \" using s by (subst nn_integral_cmult_indicator) (auto simp: \0 \ m\ simple_bochner_integrable.simps ennreal_mult_less_top less_top) finally have s_fin: "(\\<^sup>+x. norm (s i x) \M) < \" . have "(\\<^sup>+ x. norm (f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis add.commute norm_triangle_sub) also have "\ = (\\<^sup>+x. norm (f x - s i x) \M) + (\\<^sup>+x. norm (s i x) \M)" by (rule nn_integral_add) auto also have "\ < \" using s_fin f_s_fin by auto finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed proposition has_bochner_integral_norm_bound: assumes i: "has_bochner_integral M f x" shows "norm x \ (\\<^sup>+x. norm (f x) \M)" using assms proof fix s assume x: "(\i. simple_bochner_integral M (s i)) \ x" (is "?s \ x") and s[simp]: "\i. simple_bochner_integrable M (s i)" and lim: "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" and f[measurable]: "f \ borel_measurable M" have [measurable]: "\i. s i \ borel_measurable M" using s by (auto simp: simple_bochner_integrable.simps intro: borel_measurable_simple_function) show "norm x \ (\\<^sup>+x. norm (f x) \M)" proof (rule LIMSEQ_le) show "(\i. ennreal (norm (?s i))) \ norm x" using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros) show "\N. \n\N. norm (?s n) \ (\\<^sup>+x. norm (f x - s n x) \M) + (\\<^sup>+x. norm (f x) \M)" (is "\N. \n\N. _ \ ?t n") proof (intro exI allI impI) fix n have "ennreal (norm (?s n)) \ simple_bochner_integral M (\x. norm (s n x))" by (auto intro!: simple_bochner_integral_norm_bound) also have "\ = (\\<^sup>+x. norm (s n x) \M)" by (intro simple_bochner_integral_eq_nn_integral) (auto intro: s simple_bochner_integrable_compose2) also have "\ \ (\\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \M)" by (auto intro!: nn_integral_mono simp flip: ennreal_plus) (metis add.commute norm_minus_commute norm_triangle_sub) also have "\ = ?t n" by (rule nn_integral_add) auto finally show "norm (?s n) \ ?t n" . qed have "?t \ 0 + (\\<^sup>+ x. ennreal (norm (f x)) \M)" using has_bochner_integral_implies_finite_norm[OF i] by (intro tendsto_add tendsto_const lim) then show "?t \ \\<^sup>+ x. ennreal (norm (f x)) \M" by simp qed qed lemma has_bochner_integral_eq: "has_bochner_integral M f x \ has_bochner_integral M f y \ x = y" proof (elim has_bochner_integral.cases) assume f[measurable]: "f \ borel_measurable M" fix s t assume "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") assume "(\i. \\<^sup>+ x. norm (f x - t i x) \M) \ 0" (is "?T \ 0") assume s: "\i. simple_bochner_integrable M (s i)" assume t: "\i. simple_bochner_integrable M (t i)" have [measurable]: "\i. s i \ borel_measurable M" "\i. t i \ borel_measurable M" using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) let ?s = "\i. simple_bochner_integral M (s i)" let ?t = "\i. simple_bochner_integral M (t i)" assume "?s \ x" "?t \ y" then have "(\i. norm (?s i - ?t i)) \ norm (x - y)" by (intro tendsto_intros) moreover have "(\i. ennreal (norm (?s i - ?t i))) \ ennreal 0" proof (rule tendsto_sandwich) show "eventually (\i. 0 \ ennreal (norm (?s i - ?t i))) sequentially" "(\_. 0) \ ennreal 0" by auto show "eventually (\i. norm (?s i - ?t i) \ ?S i + ?T i) sequentially" by (intro always_eventually allI simple_bochner_integral_bounded s t f) show "(\i. ?S i + ?T i) \ ennreal 0" using tendsto_add[OF \?S \ 0\ \?T \ 0\] by simp qed then have "(\i. norm (?s i - ?t i)) \ 0" by (simp flip: ennreal_0) ultimately have "norm (x - y) = 0" by (rule LIMSEQ_unique) then show "x = y" by simp qed lemma has_bochner_integralI_AE: assumes f: "has_bochner_integral M f x" and g: "g \ borel_measurable M" and ae: "AE x in M. f x = g x" shows "has_bochner_integral M g x" using f proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) fix s assume "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" also have "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) = (\i. \\<^sup>+ x. ennreal (norm (g x - s i x)) \M)" using ae by (intro ext nn_integral_cong_AE, eventually_elim) simp finally show "(\i. \\<^sup>+ x. ennreal (norm (g x - s i x)) \M) \ 0" . qed (auto intro: g) lemma has_bochner_integral_eq_AE: assumes f: "has_bochner_integral M f x" and g: "has_bochner_integral M g y" and ae: "AE x in M. f x = g x" shows "x = y" proof - from assms have "has_bochner_integral M g x" by (auto intro: has_bochner_integralI_AE) from this g show "x = y" by (rule has_bochner_integral_eq) qed lemma simple_bochner_integrable_restrict_space: fixes f :: "_ \ 'b::real_normed_vector" assumes \: "\ \ space M \ sets M" shows "simple_bochner_integrable (restrict_space M \) f \ simple_bochner_integrable M (\x. indicator \ x *\<^sub>R f x)" by (simp add: simple_bochner_integrable.simps space_restrict_space simple_function_restrict_space[OF \] emeasure_restrict_space[OF \] Collect_restrict indicator_eq_0_iff conj_left_commute) lemma simple_bochner_integral_restrict_space: fixes f :: "_ \ 'b::real_normed_vector" assumes \: "\ \ space M \ sets M" assumes f: "simple_bochner_integrable (restrict_space M \) f" shows "simple_bochner_integral (restrict_space M \) f = simple_bochner_integral M (\x. indicator \ x *\<^sub>R f x)" proof - have "finite ((\x. indicator \ x *\<^sub>R f x)`space M)" using f simple_bochner_integrable_restrict_space[OF \, of f] by (simp add: simple_bochner_integrable.simps simple_function_def) then show ?thesis by (auto simp: space_restrict_space measure_restrict_space[OF \(1)] le_infI2 simple_bochner_integral_def Collect_restrict split: split_indicator split_indicator_asm intro!: sum.mono_neutral_cong_left arg_cong2[where f=measure]) qed context notes [[inductive_internals]] begin inductive integrable for M f where "has_bochner_integral M f x \ integrable M f" end definition\<^marker>\tag important\ lebesgue_integral ("integral\<^sup>L") where "integral\<^sup>L M f = (if \x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)" syntax "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" ("\((2 _./ _)/ \_)" [60,61] 110) translations "\ x. f \M" == "CONST lebesgue_integral M (\x. f)" syntax "_ascii_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60) translations "LINT x|M. f" == "CONST lebesgue_integral M (\x. f)" lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \ integral\<^sup>L M f = x" by (metis the_equality has_bochner_integral_eq lebesgue_integral_def) lemma has_bochner_integral_integrable: "integrable M f \ has_bochner_integral M f (integral\<^sup>L M f)" by (auto simp: has_bochner_integral_integral_eq integrable.simps) lemma has_bochner_integral_iff: "has_bochner_integral M f x \ integrable M f \ integral\<^sup>L M f = x" by (metis has_bochner_integral_integrable has_bochner_integral_integral_eq integrable.intros) lemma simple_bochner_integrable_eq_integral: "simple_bochner_integrable M f \ simple_bochner_integral M f = integral\<^sup>L M f" using has_bochner_integral_simple_bochner_integrable[of M f] by (simp add: has_bochner_integral_integral_eq) lemma not_integrable_integral_eq: "\ integrable M f \ integral\<^sup>L M f = 0" unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The]) lemma integral_eq_cases: "integrable M f \ integrable N g \ (integrable M f \ integrable N g \ integral\<^sup>L M f = integral\<^sup>L N g) \ integral\<^sup>L M f = integral\<^sup>L N g" by (metis not_integrable_integral_eq) lemma borel_measurable_integrable[measurable_dest]: "integrable M f \ f \ borel_measurable M" by (auto elim: integrable.cases has_bochner_integral.cases) lemma borel_measurable_integrable'[measurable_dest]: "integrable M f \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" using borel_measurable_integrable[measurable] by measurable lemma integrable_cong: "M = N \ (\x. x \ space N \ f x = g x) \ integrable M f \ integrable N g" by (simp cong: has_bochner_integral_cong add: integrable.simps) lemma integrable_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ integrable M f \ integrable M g" unfolding integrable.simps by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext) lemma integrable_cong_AE_imp: "integrable M g \ f \ borel_measurable M \ (AE x in M. g x = f x) \ integrable M f" using integrable_cong_AE[of f M g] by (auto simp: eq_commute) lemma integral_cong: "M = N \ (\x. x \ space N \ f x = g x) \ integral\<^sup>L M f = integral\<^sup>L N g" by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def) lemma integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ integral\<^sup>L M f = integral\<^sup>L M g" unfolding lebesgue_integral_def by (rule arg_cong[where x="has_bochner_integral M f"]) (intro has_bochner_integral_cong_AE ext) lemma integrable_add[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x + g x)" by (auto simp: integrable.simps) lemma integrable_zero[simp, intro]: "integrable M (\x. 0)" by (metis has_bochner_integral_zero integrable.simps) lemma integrable_sum[simp, intro]: "(\i. i \ I \ integrable M (f i)) \ integrable M (\x. \i\I. f i x)" by (metis has_bochner_integral_sum integrable.simps) lemma integrable_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ integrable M (\x. indicator A x *\<^sub>R c)" by (metis has_bochner_integral_indicator integrable.simps) lemma integrable_real_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ integrable M (indicator A :: 'a \ real)" by (metis has_bochner_integral_real_indicator integrable.simps) lemma integrable_diff[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x - g x)" by (auto simp: integrable.simps intro: has_bochner_integral_diff) lemma integrable_bounded_linear: "bounded_linear T \ integrable M f \ integrable M (\x. T (f x))" by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear) lemma integrable_scaleR_left[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. f x *\<^sub>R c)" unfolding integrable.simps by fastforce lemma integrable_scaleR_right[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. c *\<^sub>R f x)" unfolding integrable.simps by fastforce lemma integrable_mult_left[simp, intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x * c)" unfolding integrable.simps by fastforce lemma integrable_mult_right[simp, intro]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. c * f x)" unfolding integrable.simps by fastforce lemma integrable_divide_zero[simp, intro]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x / c)" unfolding integrable.simps by fastforce lemma integrable_inner_left[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. f x \ c)" unfolding integrable.simps by fastforce lemma integrable_inner_right[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. c \ f x)" unfolding integrable.simps by fastforce lemmas integrable_minus[simp, intro] = integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] lemmas integrable_divide[simp, intro] = integrable_bounded_linear[OF bounded_linear_divide] lemmas integrable_Re[simp, intro] = integrable_bounded_linear[OF bounded_linear_Re] lemmas integrable_Im[simp, intro] = integrable_bounded_linear[OF bounded_linear_Im] lemmas integrable_cnj[simp, intro] = integrable_bounded_linear[OF bounded_linear_cnj] lemmas integrable_of_real[simp, intro] = integrable_bounded_linear[OF bounded_linear_of_real] lemmas integrable_fst[simp, intro] = integrable_bounded_linear[OF bounded_linear_fst] lemmas integrable_snd[simp, intro] = integrable_bounded_linear[OF bounded_linear_snd] lemma integral_zero[simp]: "integral\<^sup>L M (\x. 0) = 0" by (intro has_bochner_integral_integral_eq has_bochner_integral_zero) lemma integral_add[simp]: "integrable M f \ integrable M g \ integral\<^sup>L M (\x. f x + g x) = integral\<^sup>L M f + integral\<^sup>L M g" by (intro has_bochner_integral_integral_eq has_bochner_integral_add has_bochner_integral_integrable) lemma integral_diff[simp]: "integrable M f \ integrable M g \ integral\<^sup>L M (\x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g" by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable) lemma integral_sum: "(\i. i \ I \ integrable M (f i)) \ integral\<^sup>L M (\x. \i\I. f i x) = (\i\I. integral\<^sup>L M (f i))" by (intro has_bochner_integral_integral_eq has_bochner_integral_sum has_bochner_integral_integrable) lemma integral_sum'[simp]: "(\i. i \ I =simp=> integrable M (f i)) \ integral\<^sup>L M (\x. \i\I. f i x) = (\i\I. integral\<^sup>L M (f i))" unfolding simp_implies_def by (rule integral_sum) lemma integral_bounded_linear: "bounded_linear T \ integrable M f \ integral\<^sup>L M (\x. T (f x)) = T (integral\<^sup>L M f)" by (metis has_bochner_integral_bounded_linear has_bochner_integral_integrable has_bochner_integral_integral_eq) lemma integral_bounded_linear': assumes T: "bounded_linear T" and T': "bounded_linear T'" assumes *: "\ (\x. T x = 0) \ (\x. T' (T x) = x)" shows "integral\<^sup>L M (\x. T (f x)) = T (integral\<^sup>L M f)" proof cases assume "(\x. T x = 0)" then show ?thesis by simp next assume **: "\ (\x. T x = 0)" show ?thesis proof cases assume "integrable M f" with T show ?thesis by (rule integral_bounded_linear) next assume not: "\ integrable M f" moreover have "\ integrable M (\x. T (f x))" proof assume "integrable M (\x. T (f x))" from integrable_bounded_linear[OF T' this] not *[OF **] show False by auto qed ultimately show ?thesis using T by (simp add: not_integrable_integral_eq linear_simps) qed qed lemma integral_scaleR_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x *\<^sub>R c \M) = integral\<^sup>L M f *\<^sub>R c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left) lemma integral_scaleR_right[simp]: "(\ x. c *\<^sub>R f x \M) = c *\<^sub>R integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_scaleR_right bounded_linear_scaleR_right[of "1 / c"]]) simp lemma integral_mult_left[simp]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ (\ x. f x * c \M) = integral\<^sup>L M f * c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_left) lemma integral_mult_right[simp]: fixes c :: "_::{real_normed_algebra,second_countable_topology}" shows "(c \ 0 \ integrable M f) \ (\ x. c * f x \M) = c * integral\<^sup>L M f" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_right) lemma integral_mult_left_zero[simp]: fixes c :: "_::{real_normed_field,second_countable_topology}" shows "(\ x. f x * c \M) = integral\<^sup>L M f * c" by (rule integral_bounded_linear'[OF bounded_linear_mult_left bounded_linear_mult_left[of "1 / c"]]) simp lemma integral_mult_right_zero[simp]: fixes c :: "_::{real_normed_field,second_countable_topology}" shows "(\ x. c * f x \M) = c * integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_mult_right bounded_linear_mult_right[of "1 / c"]]) simp lemma integral_inner_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x \ c \M) = integral\<^sup>L M f \ c" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left) lemma integral_inner_right[simp]: "(c \ 0 \ integrable M f) \ (\ x. c \ f x \M) = c \ integral\<^sup>L M f" by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right) lemma integral_divide_zero[simp]: fixes c :: "_::{real_normed_field, field, second_countable_topology}" shows "integral\<^sup>L M (\x. f x / c) = integral\<^sup>L M f / c" by (rule integral_bounded_linear'[OF bounded_linear_divide bounded_linear_mult_left[of c]]) simp lemma integral_minus[simp]: "integral\<^sup>L M (\x. - f x) = - integral\<^sup>L M f" by (rule integral_bounded_linear'[OF bounded_linear_minus[OF bounded_linear_ident] bounded_linear_minus[OF bounded_linear_ident]]) simp lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)" by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp lemma integral_cnj[simp]: "integral\<^sup>L M (\x. cnj (f x)) = cnj (integral\<^sup>L M f)" by (rule integral_bounded_linear'[OF bounded_linear_cnj bounded_linear_cnj]) simp lemmas integral_divide[simp] = integral_bounded_linear[OF bounded_linear_divide] lemmas integral_Re[simp] = integral_bounded_linear[OF bounded_linear_Re] lemmas integral_Im[simp] = integral_bounded_linear[OF bounded_linear_Im] lemmas integral_of_real[simp] = integral_bounded_linear[OF bounded_linear_of_real] lemmas integral_fst[simp] = integral_bounded_linear[OF bounded_linear_fst] lemmas integral_snd[simp] = integral_bounded_linear[OF bounded_linear_snd] lemma integral_norm_bound_ennreal: "integrable M f \ norm (integral\<^sup>L M f) \ (\\<^sup>+x. norm (f x) \M)" by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound) lemma integrableI_sequence: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes s: "\i. simple_bochner_integrable M (s i)" assumes lim: "(\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") shows "integrable M f" proof - let ?s = "\n. simple_bochner_integral M (s n)" have "\x. ?s \ x" unfolding convergent_eq_Cauchy proof (rule metric_CauchyI) fix e :: real assume "0 < e" then have "0 < ennreal (e / 2)" by auto from order_tendstoD(2)[OF lim this] obtain M where M: "\n. M \ n \ ?S n < e / 2" by (auto simp: eventually_sequentially) show "\M. \m\M. \n\M. dist (?s m) (?s n) < e" proof (intro exI allI impI) fix m n assume m: "M \ m" and n: "M \ n" have "?S n \ \" using M[OF n] by auto have "norm (?s n - ?s m) \ ?S n + ?S m" by (intro simple_bochner_integral_bounded s f) also have "\ < ennreal (e / 2) + e / 2" by (intro add_strict_mono M n m) also have "\ = e" using \0 by (simp flip: ennreal_plus) finally show "dist (?s n) (?s m) < e" using \0 by (simp add: dist_norm ennreal_less_iff) qed qed then obtain x where "?s \ x" .. show ?thesis by (rule, rule) fact+ qed proposition nn_integral_dominated_convergence_norm: fixes u' :: "_ \ _::{real_normed_vector, second_countable_topology}" assumes [measurable]: "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" and bound: "\j. AE x in M. norm (u j x) \ w x" and w: "(\\<^sup>+x. w x \M) < \" and u': "AE x in M. (\i. u i x) \ u' x" shows "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ 0" proof - have "AE x in M. \j. norm (u j x) \ w x" unfolding AE_all_countable by rule fact with u' have bnd: "AE x in M. \j. norm (u' x - u j x) \ 2 * w x" proof (eventually_elim, intro allI) fix i x assume "(\i. u i x) \ u' x" "\j. norm (u j x) \ w x" "\j. norm (u j x) \ w x" then have "norm (u' x) \ w x" "norm (u i x) \ w x" by (auto intro: LIMSEQ_le_const2 tendsto_norm) then have "norm (u' x) + norm (u i x) \ 2 * w x" by simp also have "norm (u' x - u i x) \ norm (u' x) + norm (u i x)" by (rule norm_triangle_ineq4) finally (xtrans) show "norm (u' x - u i x) \ 2 * w x" . qed have w_nonneg: "AE x in M. 0 \ w x" using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero]) have "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ (\\<^sup>+x. 0 \M)" proof (rule nn_integral_dominated_convergence) show "(\\<^sup>+x. 2 * w x \M) < \" by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult ) show "AE x in M. (\i. ennreal (norm (u' x - u i x))) \ 0" using u' proof eventually_elim fix x assume "(\i. u i x) \ u' x" from tendsto_diff[OF tendsto_const[of "u' x"] this] show "(\i. ennreal (norm (u' x - u i x))) \ 0" by (simp add: tendsto_norm_zero_iff flip: ennreal_0) qed qed (insert bnd w_nonneg, auto) then show ?thesis by simp qed proposition integrableI_bounded: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" and fin: "(\\<^sup>+x. norm (f x) \M) < \" shows "integrable M f" proof - from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where s: "\i. simple_function M (s i)" and pointwise: "\x. x \ space M \ (\i. s i x) \ f x" and bound: "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" by simp metis show ?thesis proof (rule integrableI_sequence) { fix i have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ennreal (2 * norm (f x)) \M)" by (intro nn_integral_mono) (simp add: bound) also have "\ = 2 * (\\<^sup>+x. ennreal (norm (f x)) \M)" by (simp add: ennreal_mult nn_integral_cmult) also have "\ < top" using fin by (simp add: ennreal_mult_less_top) finally have "(\\<^sup>+x. norm (s i x) \M) < \" by simp } note fin_s = this show "\i. simple_bochner_integrable M (s i)" by (rule simple_bochner_integrableI_bounded) fact+ show "(\i. \\<^sup>+ x. ennreal (norm (f x - s i x)) \M) \ 0" proof (rule nn_integral_dominated_convergence_norm) show "\j. AE x in M. norm (s j x) \ 2 * norm (f x)" using bound by auto show "\i. s i \ borel_measurable M" "(\x. 2 * norm (f x)) \ borel_measurable M" using s by (auto intro: borel_measurable_simple_function) show "(\\<^sup>+ x. ennreal (2 * norm (f x)) \M) < \" using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top) show "AE x in M. (\i. s i x) \ f x" using pointwise by auto qed fact qed fact qed lemma integrableI_bounded_set: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "A \ sets M" "f \ borel_measurable M" assumes finite: "emeasure M A < \" and bnd: "AE x in M. x \ A \ norm (f x) \ B" and null: "AE x in M. x \ A \ f x = 0" shows "integrable M f" proof (rule integrableI_bounded) { fix x :: 'b have "norm x \ B \ 0 \ B" using norm_ge_zero[of x] by arith } with bnd null have "(\\<^sup>+ x. ennreal (norm (f x)) \M) \ (\\<^sup>+ x. ennreal (max 0 B) * indicator A x \M)" by (intro nn_integral_mono_AE) (auto split: split_indicator split_max) also have "\ < \" using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top) finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed simp lemma integrableI_bounded_set_indicator: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "A \ sets M \ f \ borel_measurable M \ emeasure M A < \ \ (AE x in M. x \ A \ norm (f x) \ B) \ integrable M (\x. indicator A x *\<^sub>R f x)" by (rule integrableI_bounded_set[where A=A]) auto lemma integrableI_nonneg: fixes f :: "'a \ real" assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "(\\<^sup>+x. f x \M) < \" shows "integrable M f" proof - have "(\\<^sup>+x. norm (f x) \M) = (\\<^sup>+x. f x \M)" using assms by (intro nn_integral_cong_AE) auto then show ?thesis using assms by (intro integrableI_bounded) auto qed lemma integrable_iff_bounded: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "integrable M f \ f \ borel_measurable M \ (\\<^sup>+x. norm (f x) \M) < \" using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f] unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto lemma (in finite_measure) square_integrable_imp_integrable: fixes f :: "'a \ 'b :: {second_countable_topology, banach, real_normed_div_algebra}" assumes [measurable]: "f \ borel_measurable M" assumes "integrable M (\x. f x ^ 2)" shows "integrable M f" proof - have less_top: "emeasure M (space M) < top" using finite_emeasure_space by (meson top.not_eq_extremum) have "nn_integral M (\x. norm (f x)) ^ 2 \ nn_integral M (\x. norm (f x) ^ 2) * emeasure M (space M)" using Cauchy_Schwarz_nn_integral[of "\x. norm (f x)" M "\_. 1"] by (simp add: ennreal_power) also have "\ < \" using assms(2) less_top by (subst (asm) integrable_iff_bounded) (auto simp: norm_power ennreal_mult_less_top) finally have "nn_integral M (\x. norm (f x)) < \" by (simp add: power_less_top_ennreal) thus ?thesis by (simp add: integrable_iff_bounded) qed lemma integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" shows "integrable M f \ g \ borel_measurable M \ (AE x in M. norm (g x) \ norm (f x)) \ integrable M g" unfolding integrable_iff_bounded proof safe assume "f \ borel_measurable M" "g \ borel_measurable M" assume "AE x in M. norm (g x) \ norm (f x)" then have "(\\<^sup>+ x. ennreal (norm (g x)) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono_AE) auto also assume "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" . qed lemma integrable_mult_indicator: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "A \ sets M \ integrable M f \ integrable M (\x. indicator A x *\<^sub>R f x)" by (rule integrable_bound[of M f]) (auto split: split_indicator) lemma integrable_real_mult_indicator: fixes f :: "'a \ real" shows "A \ sets M \ integrable M f \ integrable M (\x. f x * indicator A x)" using integrable_mult_indicator[of A M f] by (simp add: mult_ac) lemma integrable_abs[simp, intro]: fixes f :: "'a \ real" assumes [measurable]: "integrable M f" shows "integrable M (\x. \f x\)" using assms by (rule integrable_bound) auto lemma integrable_norm[simp, intro]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "integrable M f" shows "integrable M (\x. norm (f x))" using assms by (rule integrable_bound) auto lemma integrable_norm_cancel: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "integrable M (\x. norm (f x))" "f \ borel_measurable M" shows "integrable M f" using assms by (rule integrable_bound) auto lemma integrable_norm_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "f \ borel_measurable M \ integrable M (\x. norm (f x)) \ integrable M f" by (auto intro: integrable_norm_cancel) lemma integrable_abs_cancel: fixes f :: "'a \ real" assumes [measurable]: "integrable M (\x. \f x\)" "f \ borel_measurable M" shows "integrable M f" using assms by (rule integrable_bound) auto lemma integrable_abs_iff: fixes f :: "'a \ real" shows "f \ borel_measurable M \ integrable M (\x. \f x\) \ integrable M f" by (auto intro: integrable_abs_cancel) lemma integrable_max[simp, intro]: fixes f :: "'a \ real" assumes fg[measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. max (f x) (g x))" using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] by (rule integrable_bound) auto lemma integrable_min[simp, intro]: fixes f :: "'a \ real" assumes fg[measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. min (f x) (g x))" using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] by (rule integrable_bound) auto lemma integral_minus_iff[simp]: "integrable M (\x. - f x ::'a::{banach, second_countable_topology}) \ integrable M f" unfolding integrable_iff_bounded by (auto) lemma integrable_indicator_iff: "integrable M (indicator A::_ \ real) \ A \ space M \ sets M \ emeasure M (A \ space M) < \" by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator' cong: conj_cong) lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \ space M)" proof cases assume *: "A \ space M \ sets M \ emeasure M (A \ space M) < \" have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \ space M))" by (intro integral_cong) (auto split: split_indicator) also have "\ = measure M (A \ space M)" using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto finally show ?thesis . next assume *: "\ (A \ space M \ sets M \ emeasure M (A \ space M) < \)" have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \ space M) :: _ \ real)" by (intro integral_cong) (auto split: split_indicator) also have "\ = 0" using * by (subst not_integrable_integral_eq) (auto simp: integrable_indicator_iff) also have "\ = measure M (A \ space M)" using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique) finally show ?thesis . qed lemma integrable_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "integrable M f \ integrable M g" unfolding integrable_iff_bounded proof (rule conj_cong) { assume "f \ borel_measurable M" then have "g \ borel_measurable M" by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } moreover { assume "g \ borel_measurable M" then have "f \ borel_measurable M" by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) } ultimately show "f \ borel_measurable M \ g \ borel_measurable M" .. next have "AE x in M. x \ X" by (rule AE_discrete_difference) fact+ then have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. norm (g x) \M)" by (intro nn_integral_cong_AE) (auto simp: eq) then show "(\\<^sup>+ x. norm (f x) \M) < \ \ (\\<^sup>+ x. norm (g x) \M) < \" by simp qed lemma integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "integral\<^sup>L M f = integral\<^sup>L M g" proof (rule integral_eq_cases) show eq: "integrable M f \ integrable M g" by (rule integrable_discrete_difference[where X=X]) fact+ assume f: "integrable M f" show "integral\<^sup>L M f = integral\<^sup>L M g" proof (rule integral_cong_AE) show "f \ borel_measurable M" "g \ borel_measurable M" using f eq by (auto intro: borel_measurable_integrable) have "AE x in M. x \ X" by (rule AE_discrete_difference) fact+ with AE_space show "AE x in M. f x = g x" by eventually_elim fact qed qed lemma has_bochner_integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "has_bochner_integral M f x \ has_bochner_integral M g x" using integrable_discrete_difference[of X M f g, OF assms] using integral_discrete_difference[of X M f g, OF assms] by (metis has_bochner_integral_iff) lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" assumes "f \ borel_measurable M" "\i. s i \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. (\i. s i x) \ f x" assumes bound: "\i. AE x in M. norm (s i x) \ w x" shows integrable_dominated_convergence: "integrable M f" and integrable_dominated_convergence2: "\i. integrable M (s i)" and integral_dominated_convergence: "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" proof - have w_nonneg: "AE x in M. 0 \ w x" using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans) then have "(\\<^sup>+x. w x \M) = (\\<^sup>+x. norm (w x) \M)" by (intro nn_integral_cong_AE) auto with \integrable M w\ have w: "w \ borel_measurable M" "(\\<^sup>+x. w x \M) < \" unfolding integrable_iff_bounded by auto show int_s: "\i. integrable M (s i)" unfolding integrable_iff_bounded proof fix i have "(\\<^sup>+ x. ennreal (norm (s i x)) \M) \ (\\<^sup>+x. w x \M)" using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto with w show "(\\<^sup>+ x. ennreal (norm (s i x)) \M) < \" by auto qed fact have all_bound: "AE x in M. \i. norm (s i x) \ w x" using bound unfolding AE_all_countable by auto show int_f: "integrable M f" unfolding integrable_iff_bounded proof have "(\\<^sup>+ x. ennreal (norm (f x)) \M) \ (\\<^sup>+x. w x \M)" using all_bound lim w_nonneg proof (intro nn_integral_mono_AE, eventually_elim) fix x assume "\i. norm (s i x) \ w x" "(\i. s i x) \ f x" "0 \ w x" then show "ennreal (norm (f x)) \ ennreal (w x)" by (intro LIMSEQ_le_const2[where X="\i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros) qed with w show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" by auto qed fact have "(\n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \ ennreal 0" (is "?d \ ennreal 0") proof (rule tendsto_sandwich) show "eventually (\n. ennreal 0 \ ?d n) sequentially" "(\_. ennreal 0) \ ennreal 0" by auto show "eventually (\n. ?d n \ (\\<^sup>+x. norm (s n x - f x) \M)) sequentially" proof (intro always_eventually allI) fix n have "?d n = norm (integral\<^sup>L M (\x. s n x - f x))" using int_f int_s by simp also have "\ \ (\\<^sup>+x. norm (s n x - f x) \M)" by (intro int_f int_s integrable_diff integral_norm_bound_ennreal) finally show "?d n \ (\\<^sup>+x. norm (s n x - f x) \M)" . qed show "(\n. \\<^sup>+x. norm (s n x - f x) \M) \ ennreal 0" unfolding ennreal_0 apply (subst norm_minus_commute) proof (rule nn_integral_dominated_convergence_norm[where w=w]) show "\n. s n \ borel_measurable M" using int_s unfolding integrable_iff_bounded by auto qed fact+ qed then have "(\n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \ 0" by (simp add: tendsto_norm_zero_iff del: ennreal_0) from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]] show "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" by simp qed context fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) at_top" assumes bound: "\\<^sub>F i in at_top. AE x in M. norm (s i x) \ w x" begin lemma integral_dominated_convergence_at_top: "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume X: "filterlim X at_top sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) at_top" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma integrable_dominated_convergence_at_top: "integrable M f" proof - from bound obtain N where w: "\n. N \ n \ AE x in M. norm (s n x) \ w x" by (auto simp: eventually_at_top_linorder) show ?thesis proof (rule integrable_dominated_convergence) show "AE x in M. norm (s (N + i) x) \ w x" for i :: nat by (intro w) auto show "AE x in M. (\i. s (N + real i) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) at_top" then show "(\n. s (N + n) x) \ f x" by (rule filterlim_compose) (auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially) qed qed fact+ qed end lemma integrable_mult_left_iff [simp]: fixes f :: "'a \ real" shows "integrable M (\x. c * f x) \ c = 0 \ integrable M f" using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\x. c * f x"] by (cases "c = 0") auto lemma integrable_mult_right_iff [simp]: fixes f :: "'a \ real" shows "integrable M (\x. f x * c) \ c = 0 \ integrable M f" using integrable_mult_left_iff [of M c f] by (simp add: mult.commute) lemma integrableI_nn_integral_finite: assumes [measurable]: "f \ borel_measurable M" and nonneg: "AE x in M. 0 \ f x" and finite: "(\\<^sup>+x. f x \M) = ennreal x" shows "integrable M f" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (f x)) \M) = (\\<^sup>+ x. ennreal (f x) \M)" using nonneg by (intro nn_integral_cong_AE) auto with finite show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" by auto qed simp lemma integral_nonneg_AE: fixes f :: "'a \ real" assumes nonneg: "AE x in M. 0 \ f x" shows "0 \ integral\<^sup>L M f" proof cases assume f: "integrable M f" then have [measurable]: "f \ M \\<^sub>M borel" by auto have "(\x. max 0 (f x)) \ M \\<^sub>M borel" "\x. 0 \ max 0 (f x)" "integrable M (\x. max 0 (f x))" using f by auto from this have "0 \ integral\<^sup>L M (\x. max 0 (f x))" proof (induction rule: borel_measurable_induct_real) case (add f g) then have "integrable M f" "integrable M g" by (auto intro!: integrable_bound[OF add.prems]) with add show ?case by (simp add: nn_integral_add) next case (seq U) show ?case proof (rule LIMSEQ_le_const) have U_le: "x \ space M \ U i x \ max 0 (f x)" for x i using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) with seq nonneg show "(\i. integral\<^sup>L M (U i)) \ LINT x|M. max 0 (f x)" by (intro integral_dominated_convergence) (simp_all, fastforce) have "integrable M (U i)" for i using seq.prems by (rule integrable_bound) (insert U_le seq, auto) with seq show "\N. \n\N. 0 \ integral\<^sup>L M (U n)" by auto qed qed (auto) also have "\ = integral\<^sup>L M f" using nonneg by (auto intro!: integral_cong_AE) finally show ?thesis . qed (simp add: not_integrable_integral_eq) lemma integral_nonneg[simp]: fixes f :: "'a \ real" shows "(\x. x \ space M \ 0 \ f x) \ 0 \ integral\<^sup>L M f" by (intro integral_nonneg_AE) auto proposition nn_integral_eq_integral: assumes f: "integrable M f" assumes nonneg: "AE x in M. 0 \ f x" shows "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" proof - { fix f :: "'a \ real" assume f: "f \ borel_measurable M" "\x. 0 \ f x" "integrable M f" then have "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" proof (induct rule: borel_measurable_induct_real) case (set A) then show ?case by (simp add: integrable_indicator_iff ennreal_indicator emeasure_eq_ennreal_measure) next case (mult f c) then show ?case by (auto simp add: nn_integral_cmult ennreal_mult integral_nonneg_AE) next case (add g f) then have "integrable M f" "integrable M g" by (auto intro!: integrable_bound[OF add.prems]) with add show ?case by (simp add: nn_integral_add integral_nonneg_AE) next case (seq U) show ?case proof (rule LIMSEQ_unique) have U_le_f: "x \ space M \ U i x \ f x" for x i using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def) have int_U: "\i. integrable M (U i)" using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto from U_le_f seq have "(\i. integral\<^sup>L M (U i)) \ integral\<^sup>L M f" by (intro integral_dominated_convergence) auto then show "(\i. ennreal (integral\<^sup>L M (U i))) \ ennreal (integral\<^sup>L M f)" using seq f int_U by (simp add: f integral_nonneg_AE) have "(\i. \\<^sup>+ x. U i x \M) \ \\<^sup>+ x. f x \M" using seq U_le_f f by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded) then show "(\i. \x. U i x \M) \ \\<^sup>+x. f x \M" using seq int_U by simp qed qed } from this[of "\x. max 0 (f x)"] assms have "(\\<^sup>+ x. max 0 (f x) \M) = integral\<^sup>L M (\x. max 0 (f x))" by simp also have "\ = integral\<^sup>L M f" using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE) also have "(\\<^sup>+ x. max 0 (f x) \M) = (\\<^sup>+ x. f x \M)" using assms by (auto intro!: nn_integral_cong_AE simp: max_def) finally show ?thesis . qed lemma nn_integral_eq_integrable: assumes f: "f \ M \\<^sub>M borel" "AE x in M. 0 \ f x" and "0 \ x" shows "(\\<^sup>+x. f x \M) = ennreal x \ (integrable M f \ integral\<^sup>L M f = x)" proof (safe intro!: nn_integral_eq_integral assms) assume *: "(\\<^sup>+x. f x \M) = ennreal x" with integrableI_nn_integral_finite[OF f this] nn_integral_eq_integral[of M f, OF _ f(2)] show "integrable M f" "integral\<^sup>L M f = x" by (simp_all add: * assms integral_nonneg_AE) qed lemma fixes f :: "_ \ _ \ 'a :: {banach, second_countable_topology}" assumes integrable[measurable]: "\i. integrable M (f i)" and summable: "AE x in M. summable (\i. norm (f i x))" and sums: "summable (\i. (\x. norm (f i x) \M))" shows integrable_suminf: "integrable M (\x. (\i. f i x))" (is "integrable M ?S") and sums_integral: "(\i. integral\<^sup>L M (f i)) sums (\x. (\i. f i x) \M)" (is "?f sums ?x") and integral_suminf: "(\x. (\i. f i x) \M) = (\i. integral\<^sup>L M (f i))" and summable_integral: "summable (\i. integral\<^sup>L M (f i))" proof - have 1: "integrable M (\x. \i. norm (f i x))" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) = (\\<^sup>+ x. (\i. ennreal (norm (f i x))) \M)" apply (intro nn_integral_cong_AE) using summable apply eventually_elim apply (simp add: suminf_nonneg ennreal_suminf_neq_top) done also have "\ = (\i. \\<^sup>+ x. norm (f i x) \M)" by (intro nn_integral_suminf) auto also have "\ = (\i. ennreal (\x. norm (f i x) \M))" by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto finally show "(\\<^sup>+ x. ennreal (norm (\i. norm (f i x))) \M) < \" by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE) qed simp have 2: "AE x in M. (\n. \i (\i. f i x)" using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel) have 3: "\j. AE x in M. norm (\i (\i. norm (f i x))" using summable proof eventually_elim fix j x assume [simp]: "summable (\i. norm (f i x))" have "norm (\i (\i \ (\i. norm (f i x))" using sum_le_suminf[of "\i. norm (f i x)"] unfolding sums_iff by auto finally show "norm (\i (\i. norm (f i x))" by simp qed note ibl = integrable_dominated_convergence[OF _ _ 1 2 3] note int = integral_dominated_convergence[OF _ _ 1 2 3] show "integrable M ?S" by (rule ibl) measurable show "?f sums ?x" unfolding sums_def using int by (simp add: integrable) then show "?x = suminf ?f" "summable ?f" unfolding sums_iff by auto qed proposition integral_norm_bound [simp]: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" shows "norm (integral\<^sup>L M f) \ (\x. norm (f x) \M)" proof (cases "integrable M f") case True then show ?thesis using nn_integral_eq_integral[of M "\x. norm (f x)"] integral_norm_bound_ennreal[of M f] by (simp add: integral_nonneg_AE) next case False then have "norm (integral\<^sup>L M f) = 0" by (simp add: not_integrable_integral_eq) moreover have "(\x. norm (f x) \M) \ 0" by auto ultimately show ?thesis by simp qed proposition integral_abs_bound [simp]: fixes f :: "'a \ real" shows "abs (\x. f x \M) \ (\x. \f x\ \M)" using integral_norm_bound[of M f] by auto lemma integral_eq_nn_integral: assumes [measurable]: "f \ borel_measurable M" assumes nonneg: "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = enn2real (\\<^sup>+ x. ennreal (f x) \M)" proof cases assume *: "(\\<^sup>+ x. ennreal (f x) \M) = \" also have "(\\<^sup>+ x. ennreal (f x) \M) = (\\<^sup>+ x. ennreal (norm (f x)) \M)" using nonneg by (intro nn_integral_cong_AE) auto finally have "\ integrable M f" by (auto simp: integrable_iff_bounded) then show ?thesis by (simp add: * not_integrable_integral_eq) next assume "(\\<^sup>+ x. ennreal (f x) \M) \ \" then have "integrable M f" by (cases "\\<^sup>+ x. ennreal (f x) \M" rule: ennreal_cases) (auto intro!: integrableI_nn_integral_finite assms) from nn_integral_eq_integral[OF this] nonneg show ?thesis by (simp add: integral_nonneg_AE) qed lemma enn2real_nn_integral_eq_integral: assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \ g x" and fin: "(\\<^sup>+x. f x \M) < top" and [measurable]: "g \ M \\<^sub>M borel" shows "enn2real (\\<^sup>+x. f x \M) = (\x. g x \M)" proof - have "ennreal (enn2real (\\<^sup>+x. f x \M)) = (\\<^sup>+x. f x \M)" using fin by (intro ennreal_enn2real) auto also have "\ = (\\<^sup>+x. g x \M)" using eq by (rule nn_integral_cong_AE) also have "\ = (\x. g x \M)" proof (rule nn_integral_eq_integral) show "integrable M g" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (g x)) \M) = (\\<^sup>+ x. f x \M)" using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2) also note fin finally show "(\\<^sup>+ x. ennreal (norm (g x)) \M) < \" by simp qed simp qed fact finally show ?thesis using nn by (simp add: integral_nonneg_AE) qed lemma has_bochner_integral_nn_integral: assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "0 \ x" assumes "(\\<^sup>+x. f x \M) = ennreal x" shows "has_bochner_integral M f x" unfolding has_bochner_integral_iff using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite) lemma integrableI_simple_bochner_integrable: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "simple_bochner_integrable M f \ integrable M f" by (intro integrableI_sequence[where s="\_. f"] borel_measurable_simple_function) (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps) proposition integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "integrable M f" assumes base: "\A c. A \ sets M \ emeasure M A < \ \ P (\x. indicator A x *\<^sub>R c)" assumes add: "\f g. integrable M f \ P f \ integrable M g \ P g \ P (\x. f x + g x)" assumes lim: "\f s. (\i. integrable M (s i)) \ (\i. P (s i)) \ (\x. x \ space M \ (\i. s i x) \ f x) \ (\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)) \ integrable M f \ P f" shows "P f" proof - from \integrable M f\ have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" unfolding integrable_iff_bounded by auto from borel_measurable_implies_sequence_metric[OF f(1)] obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) \ f x" "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" unfolding norm_conv_dist by metis { fix f A have [simp]: "P (\x. 0)" using base[of "{}" undefined] by simp have "(\i::'b. i \ A \ integrable M (f i::'a \ 'b)) \ (\i. i \ A \ P (f i)) \ P (\x. \i\A. f i x)" by (induct A rule: infinite_finite_induct) (auto intro!: add) } note sum = this define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z then have s'_eq_s: "\i x. x \ space M \ s' i x = s i x" by simp have sf[measurable]: "\i. simple_function M (s' i)" unfolding s'_def using s(1) by (intro simple_function_compose2[where h="(*\<^sub>R)"] simple_function_indicator) auto { fix i have "\z. {y. s' i z = y \ y \ s' i ` space M \ y \ 0 \ z \ space M} = (if z \ space M \ s' i z \ 0 then {s' i z} else {})" by (auto simp add: s'_def split: split_indicator) then have "\z. s' i = (\z. \y\s' i`space M - {0}. indicator {x\space M. s' i x = y} z *\<^sub>R y)" using sf by (auto simp: fun_eq_iff simple_function_def s'_def) } note s'_eq = this show "P f" proof (rule lim) fix i have "(\\<^sup>+x. norm (s' i x) \M) \ (\\<^sup>+x. ennreal (2 * norm (f x)) \M)" using s by (intro nn_integral_mono) (auto simp: s'_eq_s) also have "\ < \" using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult) finally have sbi: "simple_bochner_integrable M (s' i)" using sf by (intro simple_bochner_integrableI_bounded) auto then show "integrable M (s' i)" by (rule integrableI_simple_bochner_integrable) { fix x assume"x \ space M" "s' i x \ 0" then have "emeasure M {y \ space M. s' i y = s' i x} \ emeasure M {y \ space M. s' i y \ 0}" by (intro emeasure_mono) auto also have "\ < \" using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top) finally have "emeasure M {y \ space M. s' i y = s' i x} \ \" by simp } then show "P (s' i)" by (subst s'_eq) (auto intro!: sum base simp: less_top) fix x assume "x \ space M" with s show "(\i. s' i x) \ f x" by (simp add: s'_eq_s) show "norm (s' i x) \ 2 * norm (f x)" using \x \ space M\ s by (simp add: s'_eq_s) qed fact qed lemma integral_eq_zero_AE: "(AE x in M. f x = 0) \ integral\<^sup>L M f = 0" using integral_cong_AE[of f M "\_. 0"] by (cases "integrable M f") (simp_all add: not_integrable_integral_eq) lemma integral_nonneg_eq_0_iff_AE: fixes f :: "_ \ real" assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = 0 \ (AE x in M. f x = 0)" proof assume "integral\<^sup>L M f = 0" then have "integral\<^sup>N M f = 0" using nn_integral_eq_integral[OF f nonneg] by simp then have "AE x in M. ennreal (f x) \ 0" by (simp add: nn_integral_0_iff_AE) with nonneg show "AE x in M. f x = 0" by auto qed (auto simp add: integral_eq_zero_AE) lemma integral_mono_AE: fixes f :: "'a \ real" assumes "integrable M f" "integrable M g" "AE x in M. f x \ g x" shows "integral\<^sup>L M f \ integral\<^sup>L M g" proof - have "0 \ integral\<^sup>L M (\x. g x - f x)" using assms by (intro integral_nonneg_AE integrable_diff assms) auto also have "\ = integral\<^sup>L M g - integral\<^sup>L M f" by (intro integral_diff assms) finally show ?thesis by simp qed lemma integral_mono: fixes f :: "'a \ real" shows "integrable M f \ integrable M g \ (\x. x \ space M \ f x \ g x) \ integral\<^sup>L M f \ integral\<^sup>L M g" by (intro integral_mono_AE) auto lemma integral_norm_bound_integral: fixes f :: "'a::euclidean_space \ 'b::{banach,second_countable_topology}" assumes "integrable M f" "integrable M g" "\x. x \ space M \ norm(f x) \ g x" shows "norm (\x. f x \M) \ (\x. g x \M)" proof - have "norm (\x. f x \M) \ (\x. norm (f x) \M)" by (rule integral_norm_bound) also have "... \ (\x. g x \M)" using assms integrable_norm integral_mono by blast finally show ?thesis . qed lemma integral_abs_bound_integral: fixes f :: "'a::euclidean_space \ real" assumes "integrable M f" "integrable M g" "\x. x \ space M \ \f x\ \ g x" shows "\\x. f x \M\ \ (\x. g x \M)" by (metis integral_norm_bound_integral assms real_norm_def) text \The next two statements are useful to bound Lebesgue integrals, as they avoid one integrability assumption. The price to pay is that the upper function has to be nonnegative, but this is often true and easy to check in computations.\ lemma integral_mono_AE': fixes f::"_ \ real" assumes "integrable M f" "AE x in M. g x \ f x" "AE x in M. 0 \ f x" shows "(\x. g x \M) \ (\x. f x \M)" proof (cases "integrable M g") case True show ?thesis by (rule integral_mono_AE, auto simp add: assms True) next case False then have "(\x. g x \M) = 0" by (simp add: not_integrable_integral_eq) also have "... \ (\x. f x \M)" by (simp add: integral_nonneg_AE[OF assms(3)]) finally show ?thesis by simp qed lemma integral_mono': fixes f::"_ \ real" assumes "integrable M f" "\x. x \ space M \ g x \ f x" "\x. x \ space M \ 0 \ f x" shows "(\x. g x \M) \ (\x. f x \M)" by (rule integral_mono_AE', insert assms, auto) lemma (in finite_measure) integrable_measure: assumes I: "disjoint_family_on X I" "countable I" shows "integrable (count_space I) (\i. measure M (X i))" proof - have "(\\<^sup>+i. measure M (X i) \count_space I) = (\\<^sup>+i. measure M (if X i \ sets M then X i else {}) \count_space I)" by (auto intro!: nn_integral_cong measure_notin_sets) also have "\ = measure M (\i\I. if X i \ sets M then X i else {})" using I unfolding emeasure_eq_measure[symmetric] by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def) finally show ?thesis by (auto intro!: integrableI_bounded) qed lemma integrableI_real_bounded: assumes f: "f \ borel_measurable M" and ae: "AE x in M. 0 \ f x" and fin: "integral\<^sup>N M f < \" shows "integrable M f" proof (rule integrableI_bounded) have "(\\<^sup>+ x. ennreal (norm (f x)) \M) = \\<^sup>+ x. ennreal (f x) \M" using ae by (auto intro: nn_integral_cong_AE) also note fin finally show "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" . qed fact lemma nn_integral_nonneg_infinite: fixes f::"'a \ real" assumes "f \ borel_measurable M" "\ integrable M f" "AE x in M. f x \ 0" shows "(\\<^sup>+x. f x \M) = \" using assms integrableI_real_bounded less_top by auto lemma integral_real_bounded: assumes "0 \ r" "integral\<^sup>N M f \ ennreal r" shows "integral\<^sup>L M f \ r" proof cases assume [simp]: "integrable M f" have "integral\<^sup>L M (\x. max 0 (f x)) = integral\<^sup>N M (\x. max 0 (f x))" by (intro nn_integral_eq_integral[symmetric]) auto also have "\ = integral\<^sup>N M f" by (intro nn_integral_cong) (simp add: max_def ennreal_neg) also have "\ \ r" by fact finally have "integral\<^sup>L M (\x. max 0 (f x)) \ r" using \0 \ r\ by simp moreover have "integral\<^sup>L M f \ integral\<^sup>L M (\x. max 0 (f x))" by (rule integral_mono_AE) auto ultimately show ?thesis by simp next assume "\ integrable M f" then show ?thesis using \0 \ r\ by (simp add: not_integrable_integral_eq) qed lemma integrable_MIN: fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MIN i\I. f i x)" by (induct rule: finite_ne_induct) simp+ lemma integrable_MAX: fixes f:: "_ \ _ \ real" shows "\ finite I; I \ {}; \i. i \ I \ integrable M (f i) \ \ integrable M (\x. MAX i\I. f i x)" by (induct rule: finite_ne_induct) simp+ theorem integral_Markov_inequality: assumes [measurable]: "integrable M u" and "AE x in M. 0 \ u x" "0 < (c::real)" shows "(emeasure M) {x\space M. u x \ c} \ (1/c) * (\x. u x \M)" proof - have "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\\<^sup>+ x. u x \M)" by (rule nn_integral_mono_AE, auto simp add: \c>0\ less_eq_real_def) also have "... = (\x. u x \M)" by (rule nn_integral_eq_integral, auto simp add: assms) finally have *: "(\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M) \ (\x. u x \M)" by simp have "{x \ space M. u x \ c} = {x \ space M. ennreal(1/c) * u x \ 1} \ (space M)" using \c>0\ by (auto simp: ennreal_mult'[symmetric]) then have "emeasure M {x \ space M. u x \ c} = emeasure M ({x \ space M. ennreal(1/c) * u x \ 1})" by simp also have "... \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M)" by (rule nn_integral_Markov_inequality) (auto simp add: assms) also have "... \ ennreal(1/c) * (\x. u x \M)" by (rule mult_left_mono) (use * \c > 0\ in auto) finally show ?thesis using \0 by (simp add: ennreal_mult'[symmetric]) qed theorem integral_Markov_inequality_measure: assumes [measurable]: "integrable M u" and "A \ sets M" and "AE x in M. 0 \ u x" "0 < (c::real)" shows "measure M {x\space M. u x \ c} \ (\x. u x \M) / c" proof - have le: "emeasure M {x\space M. u x \ c} \ ennreal ((1/c) * (\x. u x \M))" by (rule integral_Markov_inequality) (use assms in auto) also have "\ < top" by auto finally have "ennreal (measure M {x\space M. u x \ c}) = emeasure M {x\space M. u x \ c}" by (intro emeasure_eq_ennreal_measure [symmetric]) auto also note le finally show ?thesis by (subst (asm) ennreal_le_iff) (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms) qed theorem%important (in finite_measure) second_moment_method: assumes [measurable]: "f \ M \\<^sub>M borel" assumes "integrable M (\x. f x ^ 2)" defines "\ \ lebesgue_integral M f" assumes "a > 0" shows "measure M {x\space M. \f x\ \ a} \ lebesgue_integral M (\x. f x ^ 2) / a\<^sup>2" proof - have integrable: "integrable M f" using assms by (blast dest: square_integrable_imp_integrable) have "{x\space M. \f x\ \ a} = {x\space M. f x ^ 2 \ a\<^sup>2}" using \a > 0\ by (simp flip: abs_le_square_iff) hence "measure M {x\space M. \f x\ \ a} = measure M {x\space M. f x ^ 2 \ a\<^sup>2}" by simp also have "\ \ lebesgue_integral M (\x. f x ^ 2) / a\<^sup>2" using assms by (intro integral_Markov_inequality_measure) auto finally show ?thesis . qed lemma integral_ineq_eq_0_then_AE: fixes f::"_ \ real" assumes "AE x in M. f x \ g x" "integrable M f" "integrable M g" "(\x. f x \M) = (\x. g x \M)" shows "AE x in M. f x = g x" proof - define h where "h = (\x. g x - f x)" have "AE x in M. h x = 0" apply (subst integral_nonneg_eq_0_iff_AE[symmetric]) unfolding h_def using assms by auto then show ?thesis unfolding h_def by auto qed lemma not_AE_zero_int_E: fixes f::"'a \ real" assumes "AE x in M. f x \ 0" "(\x. f x \M) > 0" and [measurable]: "f \ borel_measurable M" shows "\A e. A \ sets M \ e>0 \ emeasure M A > 0 \ (\x \ A. f x \ e)" proof (rule not_AE_zero_E, auto simp add: assms) assume *: "AE x in M. f x = 0" have "(\x. f x \M) = (\x. 0 \M)" by (rule integral_cong_AE, auto simp add: *) then have "(\x. f x \M) = 0" by simp then show False using assms(2) by simp qed proposition tendsto_L1_int: fixes u :: "_ \ _ \ 'b::{banach, second_countable_topology}" assumes [measurable]: "\n. integrable M (u n)" "integrable M f" and "((\n. (\\<^sup>+x. norm(u n x - f x) \M)) \ 0) F" shows "((\n. (\x. u n x \M)) \ (\x. f x \M)) F" proof - have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ (0::ennreal)) F" proof (rule tendsto_sandwich[of "\_. 0", where ?h = "\n. (\\<^sup>+x. norm(u n x - f x) \M)"], auto simp add: assms) { fix n have "(\x. u n x \M) - (\x. f x \M) = (\x. u n x - f x \M)" apply (rule Bochner_Integration.integral_diff[symmetric]) using assms by auto then have "norm((\x. u n x \M) - (\x. f x \M)) = norm (\x. u n x - f x \M)" by auto also have "... \ (\x. norm(u n x - f x) \M)" by (rule integral_norm_bound) finally have "ennreal(norm((\x. u n x \M) - (\x. f x \M))) \ (\x. norm(u n x - f x) \M)" by simp also have "... = (\\<^sup>+x. norm(u n x - f x) \M)" apply (rule nn_integral_eq_integral[symmetric]) using assms by auto finally have "norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)" by simp } then show "eventually (\n. norm((\x. u n x \M) - (\x. f x \M)) \ (\\<^sup>+x. norm(u n x - f x) \M)) F" by auto qed then have "((\n. norm((\x. u n x \M) - (\x. f x \M))) \ 0) F" by (simp flip: ennreal_0) then have "((\n. ((\x. u n x \M) - (\x. f x \M))) \ 0) F" using tendsto_norm_zero_iff by blast then show ?thesis using Lim_null by auto qed text \The next lemma asserts that, if a sequence of functions converges in \L\<^sup>1\, then it admits a subsequence that converges almost everywhere.\ proposition tendsto_L1_AE_subseq: fixes u :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "\n. integrable M (u n)" and "(\n. (\x. norm(u n x) \M)) \ 0" shows "\r::nat\nat. strict_mono r \ (AE x in M. (\n. u (r n) x) \ 0)" proof - { fix k have "eventually (\n. (\x. norm(u n x) \M) < (1/2)^k) sequentially" using order_tendstoD(2)[OF assms(2)] by auto with eventually_elim2[OF eventually_gt_at_top[of k] this] have "\n>k. (\x. norm(u n x) \M) < (1/2)^k" by (metis eventually_False_sequentially) } then have "\r. \n. True \ (r (Suc n) > r n \ (\x. norm(u (r (Suc n)) x) \M) < (1/2)^(r n))" by (intro dependent_nat_choice, auto) then obtain r0 where r0: "strict_mono r0" "\n. (\x. norm(u (r0 (Suc n)) x) \M) < (1/2)^(r0 n)" by (auto simp: strict_mono_Suc_iff) define r where "r = (\n. r0(n+1))" have "strict_mono r" unfolding r_def using r0(1) by (simp add: strict_mono_Suc_iff) have I: "(\\<^sup>+x. norm(u (r n) x) \M) < ennreal((1/2)^n)" for n proof - have "r0 n \ n" using \strict_mono r0\ by (simp add: seq_suble) have "(1/2::real)^(r0 n) \ (1/2)^n" by (rule power_decreasing[OF \r0 n \ n\], auto) then have "(\x. norm(u (r0 (Suc n)) x) \M) < (1/2)^n" using r0(2) less_le_trans by blast then have "(\x. norm(u (r n) x) \M) < (1/2)^n" unfolding r_def by auto moreover have "(\\<^sup>+x. norm(u (r n) x) \M) = (\x. norm(u (r n) x) \M)" by (rule nn_integral_eq_integral, auto simp add: integrable_norm[OF assms(1)[of "r n"]]) ultimately show ?thesis by (auto intro: ennreal_lessI) qed have "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0" proof (rule AE_upper_bound_inf_ennreal) fix e::real assume "e > 0" define A where "A = (\n. {x \ space M. norm(u (r n) x) \ e})" have A_meas [measurable]: "\n. A n \ sets M" unfolding A_def by auto have A_bound: "emeasure M (A n) < (1/e) * ennreal((1/2)^n)" for n proof - have *: "indicator (A n) x \ (1/e) * ennreal(norm(u (r n) x))" for x apply (cases "x \ A n") unfolding A_def using \0 < e\ by (auto simp: ennreal_mult[symmetric]) have "emeasure M (A n) = (\\<^sup>+x. indicator (A n) x \M)" by auto also have "... \ (\\<^sup>+x. (1/e) * ennreal(norm(u (r n) x)) \M)" apply (rule nn_integral_mono) using * by auto also have "... = (1/e) * (\\<^sup>+x. norm(u (r n) x) \M)" apply (rule nn_integral_cmult) using \e > 0\ by auto also have "... < (1/e) * ennreal((1/2)^n)" using I[of n] \e > 0\ by (intro ennreal_mult_strict_left_mono) auto finally show ?thesis by simp qed have A_fin: "emeasure M (A n) < \" for n using \e > 0\ A_bound[of n] by (auto simp add: ennreal_mult less_top[symmetric]) have A_sum: "summable (\n. measure M (A n))" proof (rule summable_comparison_test'[of "\n. (1/e) * (1/2)^n" 0]) have "summable (\n. (1/(2::real))^n)" by (simp add: summable_geometric) then show "summable (\n. (1/e) * (1/2)^n)" using summable_mult by blast fix n::nat assume "n \ 0" have "norm(measure M (A n)) = measure M (A n)" by simp also have "... = enn2real(emeasure M (A n))" unfolding measure_def by simp also have "... < enn2real((1/e) * (1/2)^n)" using A_bound[of n] \emeasure M (A n) < \\ \0 < e\ by (auto simp: emeasure_eq_ennreal_measure ennreal_mult[symmetric] ennreal_less_iff) also have "... = (1/e) * (1/2)^n" using \0 by auto finally show "norm(measure M (A n)) \ (1/e) * (1/2)^n" by simp qed have "AE x in M. eventually (\n. x \ space M - A n) sequentially" by (rule borel_cantelli_AE1[OF A_meas A_fin A_sum]) moreover { fix x assume "eventually (\n. x \ space M - A n) sequentially" moreover have "norm(u (r n) x) \ ennreal e" if "x \ space M - A n" for n using that unfolding A_def by (auto intro: ennreal_leI) ultimately have "eventually (\n. norm(u (r n) x) \ ennreal e) sequentially" by (simp add: eventually_mono) then have "limsup (\n. ennreal (norm(u (r n) x))) \ e" by (simp add: Limsup_bounded) } ultimately show "AE x in M. limsup (\n. ennreal (norm(u (r n) x))) \ 0 + ennreal e" by auto qed moreover { fix x assume "limsup (\n. ennreal (norm(u (r n) x))) \ 0" moreover then have "liminf (\n. ennreal (norm(u (r n) x))) \ 0" by (rule order_trans[rotated]) (auto intro: Liminf_le_Limsup) ultimately have "(\n. ennreal (norm(u (r n) x))) \ 0" using tendsto_Limsup[of sequentially "\n. ennreal (norm(u (r n) x))"] by auto then have "(\n. norm(u (r n) x)) \ 0" by (simp flip: ennreal_0) then have "(\n. u (r n) x) \ 0" by (simp add: tendsto_norm_zero_iff) } ultimately have "AE x in M. (\n. u (r n) x) \ 0" by auto then show ?thesis using \strict_mono r\ by auto qed subsection \Restricted measure spaces\ lemma integrable_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \[simp]: "\ \ space M \ sets M" shows "integrable (restrict_space M \) f \ integrable M (\x. indicator \ x *\<^sub>R f x)" unfolding integrable_iff_bounded borel_measurable_restrict_space_iff[OF \] nn_integral_restrict_space[OF \] by (simp add: ac_simps ennreal_indicator ennreal_mult) lemma integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes \[simp]: "\ \ space M \ sets M" shows "integral\<^sup>L (restrict_space M \) f = integral\<^sup>L M (\x. indicator \ x *\<^sub>R f x)" proof (rule integral_eq_cases) assume "integrable (restrict_space M \) f" then show ?thesis proof induct case (base A c) then show ?case by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff emeasure_restrict_space Int_absorb1 measure_restrict_space) next case (add g f) then show ?case by (simp add: scaleR_add_right integrable_restrict_space) next case (lim f s) show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ integral\<^sup>L (restrict_space M \) f" using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) simp_all show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ (\ x. indicator \ x *\<^sub>R f x \M)" unfolding lim using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (indicator \ x *\<^sub>R f x)"]) (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR split: split_indicator) qed qed qed (simp add: integrable_restrict_space) lemma integral_empty: assumes "space M = {}" shows "integral\<^sup>L M f = 0" proof - have "(\ x. f x \M) = (\ x. 0 \M)" by(rule integral_cong)(simp_all add: assms) thus ?thesis by simp qed subsection \Measure spaces with an associated density\ lemma integrable_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" and nn: "AE x in M. 0 \ g x" shows "integrable (density M g) f \ integrable M (\x. g x *\<^sub>R f x)" unfolding integrable_iff_bounded using nn apply (simp add: nn_integral_density less_top[symmetric]) apply (intro arg_cong2[where f="(=)"] refl nn_integral_cong_AE) apply (auto simp: ennreal_mult) done lemma integral_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" assumes f: "f \ borel_measurable M" and g[measurable]: "g \ borel_measurable M" "AE x in M. 0 \ g x" shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\x. g x *\<^sub>R f x)" proof (rule integral_eq_cases) assume "integrable (density M g) f" then show ?thesis proof induct case (base A c) then have [measurable]: "A \ sets M" by auto have int: "integrable M (\x. g x * indicator A x)" using g base integrable_density[of "indicator A :: 'a \ real" M g] by simp then have "integral\<^sup>L M (\x. g x * indicator A x) = (\\<^sup>+ x. ennreal (g x * indicator A x) \M)" using g by (subst nn_integral_eq_integral) auto also have "\ = (\\<^sup>+ x. ennreal (g x) * indicator A x \M)" by (intro nn_integral_cong) (auto split: split_indicator) also have "\ = emeasure (density M g) A" by (rule emeasure_density[symmetric]) auto also have "\ = ennreal (measure (density M g) A)" using base by (auto intro: emeasure_eq_ennreal_measure) also have "\ = integral\<^sup>L (density M g) (indicator A)" using base by simp finally show ?case using base g apply (simp add: int integral_nonneg_AE) apply (subst (asm) ennreal_inj) apply (auto intro!: integral_nonneg_AE) done next case (add f h) then have [measurable]: "f \ borel_measurable M" "h \ borel_measurable M" by (auto dest!: borel_measurable_integrable) from add g show ?case by (simp add: scaleR_add_right integrable_density) next case (lim f s) have [measurable]: "f \ borel_measurable M" "\i. s i \ borel_measurable M" using lim(1,5)[THEN borel_measurable_integrable] by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L M (\x. g x *\<^sub>R f x)" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (g x *\<^sub>R f x))" by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto show "AE x in M. (\i. g x *\<^sub>R s i x) \ g x *\<^sub>R f x" using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M]) show "\i. AE x in M. norm (g x *\<^sub>R s i x) \ 2 * norm (g x *\<^sub>R f x)" using lim(4) g by (auto intro!: AE_I2[of M] mult_left_mono simp: field_simps) qed auto show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L (density M g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) qed qed qed (simp add: f g integrable_density) lemma fixes g :: "'a \ real" assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "g \ borel_measurable M" shows integral_real_density: "integral\<^sup>L (density M f) g = (\ x. f x * g x \M)" and integrable_real_density: "integrable (density M f) g \ integrable M (\x. f x * g x)" using assms integral_density[of g M f] integrable_density[of g M f] by auto lemma has_bochner_integral_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" shows "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. 0 \ g x) \ has_bochner_integral M (\x. g x *\<^sub>R f x) x \ has_bochner_integral (density M g) f x" by (simp add: has_bochner_integral_iff integrable_density integral_density) subsection \Distributions\ lemma integrable_distr_eq: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "g \ measurable M N" "f \ borel_measurable N" shows "integrable (distr M N g) f \ integrable M (\x. f (g x))" unfolding integrable_iff_bounded by (simp_all add: nn_integral_distr) lemma integrable_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "T \ measurable M M' \ integrable (distr M M' T) f \ integrable M (\x. f (T x))" by (subst integrable_distr_eq[symmetric, where g=T]) (auto dest: borel_measurable_integrable) lemma integral_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g[measurable]: "g \ measurable M N" and f: "f \ borel_measurable N" shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\x. f (g x))" proof (rule integral_eq_cases) assume "integrable (distr M N g) f" then show ?thesis proof induct case (base A c) then have [measurable]: "A \ sets N" by auto from base have int: "integrable (distr M N g) (\a. indicator A a *\<^sub>R c)" by (intro integrable_indicator) have "integral\<^sup>L (distr M N g) (\a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c" using base by auto also have "\ = measure M (g -` A \ space M) *\<^sub>R c" by (subst measure_distr) auto also have "\ = integral\<^sup>L M (\a. indicator (g -` A \ space M) a *\<^sub>R c)" using base by (auto simp: emeasure_distr) also have "\ = integral\<^sup>L M (\a. indicator A (g a) *\<^sub>R c)" using int base by (intro integral_cong_AE) (auto simp: emeasure_distr split: split_indicator) finally show ?case . next case (add f h) then have [measurable]: "f \ borel_measurable N" "h \ borel_measurable N" by (auto dest!: borel_measurable_integrable) from add g show ?case by (simp add: scaleR_add_right integrable_distr_eq) next case (lim f s) have [measurable]: "f \ borel_measurable N" "\i. s i \ borel_measurable N" using lim(1,5)[THEN borel_measurable_integrable] by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L M (\x. f (g x))" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (f (g x)))" using lim by (auto simp: integrable_distr_eq) show "AE x in M. (\i. s i (g x)) \ f (g x)" using lim(3) g[THEN measurable_space] by auto show "\i. AE x in M. norm (s i (g x)) \ 2 * norm (f (g x))" using lim(4) g[THEN measurable_space] by auto qed auto show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L (distr M N g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) qed qed qed (simp add: f g integrable_distr_eq) lemma has_bochner_integral_distr: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "f \ borel_measurable N \ g \ measurable M N \ has_bochner_integral M (\x. f (g x)) x \ has_bochner_integral (distr M N g) f x" by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr) subsection \Lebesgue integration on \<^const>\count_space\\ lemma integrable_count_space: fixes f :: "'a \ 'b::{banach,second_countable_topology}" shows "finite X \ integrable (count_space X) f" by (auto simp: nn_integral_count_space integrable_iff_bounded) lemma measure_count_space[simp]: "B \ A \ finite B \ measure (count_space A) B = card B" unfolding measure_def by (subst emeasure_count_space ) auto lemma lebesgue_integral_count_space_finite_support: assumes f: "finite {a\A. f a \ 0}" shows "(\x. f x \count_space A) = (\a | a \ A \ f a \ 0. f a)" proof - have eq: "\x. x \ A \ (\a | x = a \ a \ A \ f a \ 0. f a) = (\x\{x}. f x)" by (intro sum.mono_neutral_cong_left) auto have "(\x. f x \count_space A) = (\x. (\a | a \ A \ f a \ 0. indicator {a} x *\<^sub>R f a) \count_space A)" by (intro integral_cong refl) (simp add: f eq) also have "\ = (\a | a \ A \ f a \ 0. measure (count_space A) {a} *\<^sub>R f a)" by (subst integral_sum) (auto intro!: sum.cong) finally show ?thesis by auto qed lemma lebesgue_integral_count_space_finite: "finite A \ (\x. f x \count_space A) = (\a\A. f a)" by (subst lebesgue_integral_count_space_finite_support) (auto intro!: sum.mono_neutral_cong_left) lemma integrable_count_space_nat_iff: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ summable (\x. norm (f x))" by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top intro: summable_suminf_not_top) lemma sums_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" assumes *: "integrable (count_space UNIV) f" shows "f sums (integral\<^sup>L (count_space UNIV) f)" proof - let ?f = "\n i. indicator {n} i *\<^sub>R f i" have f': "\n i. ?f n i = indicator {n} i *\<^sub>R f n" by (auto simp: fun_eq_iff split: split_indicator) have "(\i. \n. ?f i n \count_space UNIV) sums \ n. (\i. ?f i n) \count_space UNIV" proof (rule sums_integral) show "\i. integrable (count_space UNIV) (?f i)" using * by (intro integrable_mult_indicator) auto show "AE n in count_space UNIV. summable (\i. norm (?f i n))" using summable_finite[of "{n}" "\i. norm (?f i n)" for n] by simp show "summable (\i. \ n. norm (?f i n) \count_space UNIV)" using * by (subst f') (simp add: integrable_count_space_nat_iff) qed also have "(\ n. (\i. ?f i n) \count_space UNIV) = (\n. f n \count_space UNIV)" using suminf_finite[of "{n}" "\i. ?f i n" for n] by (auto intro!: integral_cong) also have "(\i. \n. ?f i n \count_space UNIV) = f" by (subst f') simp finally show ?thesis . qed lemma integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "integrable (count_space UNIV) f \ integral\<^sup>L (count_space UNIV) f = (\x. f x)" using sums_integral_count_space_nat by (rule sums_unique) lemma integrable_bij_count_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g: "bij_betw g A B" shows "integrable (count_space A) (\x. f (g x)) \ integrable (count_space B) f" unfolding integrable_iff_bounded by (subst nn_integral_bij_count_space[OF g]) auto lemma integral_bij_count_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes g: "bij_betw g A B" shows "integral\<^sup>L (count_space A) (\x. f (g x)) = integral\<^sup>L (count_space B) f" using g[THEN bij_betw_imp_funcset] apply (subst distr_bij_count_space[OF g, symmetric]) apply (intro integral_distr[symmetric]) apply auto done lemma has_bochner_integral_count_space_nat: fixes f :: "nat \ _::{banach,second_countable_topology}" shows "has_bochner_integral (count_space UNIV) f x \ f sums x" unfolding has_bochner_integral_iff by (auto intro!: sums_integral_count_space_nat) subsection \Point measure\ lemma lebesgue_integral_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" shows "finite A \ (\a. a \ A \ 0 \ f a) \ integral\<^sup>L (point_measure A f) g = (\a\A. f a *\<^sub>R g a)" by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def) proposition integrable_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" and f :: "'a \ real" shows "finite A \ integrable (point_measure A f) g" unfolding point_measure_def apply (subst density_cong[where f'="\x. ennreal (max 0 (f x))"]) apply (auto split: split_max simp: ennreal_neg) apply (subst integrable_density) apply (auto simp: AE_count_space integrable_count_space) done subsection \Lebesgue integration on \<^const>\null_measure\\ lemma has_bochner_integral_null_measure_iff[iff]: "has_bochner_integral (null_measure M) f 0 \ f \ borel_measurable M" by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def] intro!: exI[of _ "\n x. 0"] simple_bochner_integrable.intros) lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \ f \ borel_measurable M" by (auto simp add: integrable.simps) lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0" by (cases "integrable (null_measure M) f") (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq) subsection \Legacy lemmas for the real-valued Lebesgue integral\ theorem real_lebesgue_integral_def: assumes f[measurable]: "integrable M f" shows "integral\<^sup>L M f = enn2real (\\<^sup>+x. f x \M) - enn2real (\\<^sup>+x. ennreal (- f x) \M)" proof - have "integral\<^sup>L M f = integral\<^sup>L M (\x. max 0 (f x) - max 0 (- f x))" by (auto intro!: arg_cong[where f="integral\<^sup>L M"]) also have "\ = integral\<^sup>L M (\x. max 0 (f x)) - integral\<^sup>L M (\x. max 0 (- f x))" by (intro integral_diff integrable_max integrable_minus integrable_zero f) also have "integral\<^sup>L M (\x. max 0 (f x)) = enn2real (\\<^sup>+x. ennreal (f x) \M)" by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) also have "integral\<^sup>L M (\x. max 0 (- f x)) = enn2real (\\<^sup>+x. ennreal (- f x) \M)" by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg) finally show ?thesis . qed theorem real_integrable_def: "integrable M f \ f \ borel_measurable M \ (\\<^sup>+ x. ennreal (f x) \M) \ \ \ (\\<^sup>+ x. ennreal (- f x) \M) \ \" unfolding integrable_iff_bounded proof (safe del: notI) assume *: "(\\<^sup>+ x. ennreal (norm (f x)) \M) < \" have "(\\<^sup>+ x. ennreal (f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono) auto also note * finally show "(\\<^sup>+ x. ennreal (f x) \M) \ \" by simp have "(\\<^sup>+ x. ennreal (- f x) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) \M)" by (intro nn_integral_mono) auto also note * finally show "(\\<^sup>+ x. ennreal (- f x) \M) \ \" by simp next assume [measurable]: "f \ borel_measurable M" assume fin: "(\\<^sup>+ x. ennreal (f x) \M) \ \" "(\\<^sup>+ x. ennreal (- f x) \M) \ \" have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. ennreal (f x) + ennreal (- f x) \M)" by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg) also have"\ = (\\<^sup>+ x. ennreal (f x) \M) + (\\<^sup>+ x. ennreal (- f x) \M)" by (intro nn_integral_add) auto also have "\ < \" using fin by (auto simp: less_top) finally show "(\\<^sup>+ x. norm (f x) \M) < \" . qed lemma integrableD[dest]: assumes "integrable M f" shows "f \ borel_measurable M" "(\\<^sup>+ x. ennreal (f x) \M) \ \" "(\\<^sup>+ x. ennreal (- f x) \M) \ \" using assms unfolding real_integrable_def by auto lemma integrableE: assumes "integrable M f" obtains r q where "0 \ r" "0 \ q" "(\\<^sup>+x. ennreal (f x)\M) = ennreal r" "(\\<^sup>+x. ennreal (-f x)\M) = ennreal q" "f \ borel_measurable M" "integral\<^sup>L M f = r - q" using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms] by (cases rule: ennreal2_cases[of "(\\<^sup>+x. ennreal (-f x)\M)" "(\\<^sup>+x. ennreal (f x)\M)"]) auto lemma integral_monotone_convergence_nonneg: fixes f :: "nat \ 'a \ real" assumes i: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" and pos: "\i. AE x in M. 0 \ f i x" and lim: "AE x in M. (\i. f i x) \ u x" and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows "integrable M u" and "integral\<^sup>L M u = x" proof - have nn: "AE x in M. \i. 0 \ f i x" using pos unfolding AE_all_countable by auto with lim have u_nn: "AE x in M. 0 \ u x" by eventually_elim (auto intro: LIMSEQ_le_const) have [simp]: "0 \ x" by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos) have "(\\<^sup>+ x. ennreal (u x) \M) = (SUP n. (\\<^sup>+ x. ennreal (f n x) \M))" proof (subst nn_integral_monotone_convergence_SUP_AE[symmetric]) fix i from mono nn show "AE x in M. ennreal (f i x) \ ennreal (f (Suc i) x)" by eventually_elim (auto simp: mono_def) show "(\x. ennreal (f i x)) \ borel_measurable M" using i by auto next show "(\\<^sup>+ x. ennreal (u x) \M) = \\<^sup>+ x. (SUP i. ennreal (f i x)) \M" apply (rule nn_integral_cong_AE) using lim mono nn u_nn apply eventually_elim apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def) done qed also have "\ = ennreal x" using mono i nn unfolding nn_integral_eq_integral[OF i pos] by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim) finally have "(\\<^sup>+ x. ennreal (u x) \M) = ennreal x" . moreover have "(\\<^sup>+ x. ennreal (- u x) \M) = 0" using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg) ultimately show "integrable M u" "integral\<^sup>L M u = x" by (auto simp: real_integrable_def real_lebesgue_integral_def u) qed lemma fixes f :: "nat \ 'a \ real" assumes f: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" and lim: "AE x in M. (\i. f i x) \ u x" and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows integrable_monotone_convergence: "integrable M u" and integral_monotone_convergence: "integral\<^sup>L M u = x" and has_bochner_integral_monotone_convergence: "has_bochner_integral M u x" proof - have 1: "\i. integrable M (\x. f i x - f 0 x)" using f by auto have 2: "AE x in M. mono (\n. f n x - f 0 x)" using mono by (auto simp: mono_def le_fun_def) have 3: "\n. AE x in M. 0 \ f n x - f 0 x" using mono by (auto simp: field_simps mono_def le_fun_def) have 4: "AE x in M. (\i. f i x - f 0 x) \ u x - f 0 x" using lim by (auto intro!: tendsto_diff) have 5: "(\i. (\x. f i x - f 0 x \M)) \ x - integral\<^sup>L M (f 0)" using f ilim by (auto intro!: tendsto_diff) have 6: "(\x. u x - f 0 x) \ borel_measurable M" using f[of 0] u by auto note diff = integral_monotone_convergence_nonneg[OF 1 2 3 4 5 6] have "integrable M (\x. (u x - f 0 x) + f 0 x)" using diff(1) f by (rule integrable_add) with diff(2) f show "integrable M u" "integral\<^sup>L M u = x" by auto then show "has_bochner_integral M u x" by (metis has_bochner_integral_integrable) qed lemma integral_norm_eq_0_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "integrable M f" shows "(\x. norm (f x) \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" proof - have "(\\<^sup>+x. norm (f x) \M) = (\x. norm (f x) \M)" using f by (intro nn_integral_eq_integral integrable_norm) auto then have "(\x. norm (f x) \M) = 0 \ (\\<^sup>+x. norm (f x) \M) = 0" by simp also have "\ \ emeasure M {x\space M. ennreal (norm (f x)) \ 0} = 0" by (intro nn_integral_0_iff) auto finally show ?thesis by simp qed lemma integral_0_iff: fixes f :: "'a \ real" shows "integrable M f \ (\x. \f x\ \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" using integral_norm_eq_0_iff[of M f] by simp lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\x. a)" using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric]) lemma lebesgue_integral_const[simp]: fixes a :: "'a :: {banach, second_countable_topology}" shows "(\x. a \M) = measure M (space M) *\<^sub>R a" proof - { assume "emeasure M (space M) = \" "a \ 0" then have ?thesis by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) } moreover { assume "a = 0" then have ?thesis by simp } moreover { assume "emeasure M (space M) \ \" interpret finite_measure M proof qed fact have "(\x. a \M) = (\x. indicator (space M) x *\<^sub>R a \M)" by (intro integral_cong) auto also have "\ = measure M (space M) *\<^sub>R a" by (simp add: less_top[symmetric]) finally have ?thesis . } ultimately show ?thesis by blast qed lemma (in finite_measure) integrable_const_bound: fixes f :: "'a \ 'b::{banach,second_countable_topology}" shows "AE x in M. norm (f x) \ B \ f \ borel_measurable M \ integrable M f" apply (rule integrable_bound[OF integrable_const[of B], of f]) apply assumption apply (cases "0 \ B") apply auto done lemma (in finite_measure) integral_bounded_eq_bound_then_AE: assumes "AE x in M. f x \ (c::real)" "integrable M f" "(\x. f x \M) = c * measure M (space M)" shows "AE x in M. f x = c" apply (rule integral_ineq_eq_0_then_AE) using assms by auto lemma integral_indicator_finite_real: fixes f :: "'a \ real" assumes [simp]: "finite A" assumes [measurable]: "\a. a \ A \ {a} \ sets M" assumes finite: "\a. a \ A \ emeasure M {a} < \" shows "(\x. f x * indicator A x \M) = (\a\A. f a * measure M {a})" proof - have "(\x. f x * indicator A x \M) = (\x. (\a\A. f a * indicator {a} x) \M)" proof (intro integral_cong refl) fix x show "f x * indicator A x = (\a\A. f a * indicator {a} x)" by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong) qed also have "\ = (\a\A. f a * measure M {a})" using finite by (subst integral_sum) (auto) finally show ?thesis . qed lemma (in finite_measure) ennreal_integral_real: assumes [measurable]: "f \ borel_measurable M" assumes ae: "AE x in M. f x \ ennreal B" "0 \ B" shows "ennreal (\x. enn2real (f x) \M) = (\\<^sup>+x. f x \M)" proof (subst nn_integral_eq_integral[symmetric]) show "integrable M (\x. enn2real (f x))" using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI) show "(\\<^sup>+ x. ennreal (enn2real (f x)) \M) = integral\<^sup>N M f" using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top]) qed auto lemma (in finite_measure) integral_less_AE: fixes X Y :: "'a \ real" assumes int: "integrable M X" "integrable M Y" assumes A: "(emeasure M) A \ 0" "A \ sets M" "AE x in M. x \ A \ X x \ Y x" assumes gt: "AE x in M. X x \ Y x" shows "integral\<^sup>L M X < integral\<^sup>L M Y" proof - have "integral\<^sup>L M X \ integral\<^sup>L M Y" using gt int by (intro integral_mono_AE) auto moreover have "integral\<^sup>L M X \ integral\<^sup>L M Y" proof assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y" have "integral\<^sup>L M (\x. \Y x - X x\) = integral\<^sup>L M (\x. Y x - X x)" using gt int by (intro integral_cong_AE) auto also have "\ = 0" using eq int by simp finally have "(emeasure M) {x \ space M. Y x - X x \ 0} = 0" using int by (simp add: integral_0_iff) moreover have "(\\<^sup>+x. indicator A x \M) \ (\\<^sup>+x. indicator {x \ space M. Y x - X x \ 0} x \M)" using A by (intro nn_integral_mono_AE) auto then have "(emeasure M) A \ (emeasure M) {x \ space M. Y x - X x \ 0}" using int A by (simp add: integrable_def) ultimately have "emeasure M A = 0" by simp with \(emeasure M) A \ 0\ show False by auto qed ultimately show ?thesis by auto qed lemma (in finite_measure) integral_less_AE_space: fixes X Y :: "'a \ real" assumes int: "integrable M X" "integrable M Y" assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \ 0" shows "integral\<^sup>L M X < integral\<^sup>L M Y" using gt by (intro integral_less_AE[OF int, where A="space M"]) auto lemma tendsto_integral_at_top: fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f" shows "((\y. \ x. indicator {.. y} x *\<^sub>R f x \M) \ \ x. f x \M) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" show "(\n. \x. indicator {..X n} x *\<^sub>R f x \M) \ integral\<^sup>L M f" proof (rule integral_dominated_convergence) show "integrable M (\x. norm (f x))" by (rule integrable_norm) fact show "AE x in M. (\n. indicator {..X n} x *\<^sub>R f x) \ f x" proof fix x from \filterlim X at_top sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_top_ge[where c=x] by auto then show "(\n. indicator {..X n} x *\<^sub>R f x) \ f x" by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {..X n} x *\<^sub>R f x) \ norm (f x)" by (auto split: split_indicator) qed auto qed lemma fixes f :: "real \ real" assumes M: "sets M = sets borel" assumes nonneg: "AE x in M. 0 \ f x" assumes borel: "f \ borel_measurable borel" assumes int: "\y. integrable M (\x. f x * indicator {.. y} x)" assumes conv: "((\y. \ x. f x * indicator {.. y} x \M) \ x) at_top" shows has_bochner_integral_monotone_convergence_at_top: "has_bochner_integral M f x" and integrable_monotone_convergence_at_top: "integrable M f" and integral_monotone_convergence_at_top:"integral\<^sup>L M f = x" proof - from nonneg have "AE x in M. mono (\n::nat. f x * indicator {..real n} x)" by (auto split: split_indicator intro!: monoI) { fix x have "eventually (\n. f x * indicator {..real n} x = f x) sequentially" by (rule eventually_sequentiallyI[of "nat \x\"]) (auto split: split_indicator simp: nat_le_iff ceiling_le_iff) } from filterlim_cong[OF refl refl this] have "AE x in M. (\i. f x * indicator {..real i} x) \ f x" by simp have "(\i. \ x. f x * indicator {..real i} x \M) \ x" using conv filterlim_real_sequentially by (rule filterlim_compose) have M_measure[simp]: "borel_measurable M = borel_measurable borel" using M by (simp add: sets_eq_imp_space_eq measurable_def) have "f \ borel_measurable M" using borel by simp show "has_bochner_integral M f x" by (rule has_bochner_integral_monotone_convergence) fact+ then show "integrable M f" "integral\<^sup>L M f = x" by (auto simp: _has_bochner_integral_iff) qed subsection \Product measure\ lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes [measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "Measurable.pred N (\x. integrable M (f x))" proof - have [simp]: "\x. x \ space N \ integrable M (f x) \ (\\<^sup>+y. norm (f x y) \M) < \" unfolding integrable_iff_bounded by simp show ?thesis by (simp cong: measurable_cong) qed lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]: "(\x. x \ space N \ A x \ space M) \ {x \ space (N \\<^sub>M M). snd x \ A (fst x)} \ sets (N \\<^sub>M M) \ (\x. measure M (A x)) \ borel_measurable N" unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto proposition (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]: fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f[measurable]: "case_prod f \ borel_measurable (N \\<^sub>M M)" shows "(\x. \y. f x y \M) \ borel_measurable N" proof - - from borel_measurable_implies_sequence_metric[OF f, of 0] guess s .. - then have s: "\i. simple_function (N \\<^sub>M M) (s i)" + from borel_measurable_implies_sequence_metric[OF f, of 0] + obtain s where s: "\i. simple_function (N \\<^sub>M M) (s i)" + and "\x\space (N \\<^sub>M M). (\i. s i x) \ (case x of (x, y) \ f x y)" + and "\i. \x\space (N \\<^sub>M M). dist (s i x) 0 \ 2 * dist (case x of (x, xa) \ f x xa) 0" + by auto + then have *: "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) \ f x y" "\i x y. x \ space N \ y \ space M \ norm (s i (x, y)) \ 2 * norm (f x y)" by (auto simp: space_pair_measure) have [measurable]: "\i. s i \ borel_measurable (N \\<^sub>M M)" by (rule borel_measurable_simple_function) fact have "\i. s i \ measurable (N \\<^sub>M M) (count_space UNIV)" by (rule measurable_simple_function) fact define f' where [abs_def]: "f' i x = (if integrable M (f x) then simple_bochner_integral M (\y. s i (x, y)) else 0)" for i x { fix i x assume "x \ space N" then have "simple_bochner_integral M (\y. s i (x, y)) = (\z\s i ` (space N \ space M). measure M {y \ space M. s i (x, y) = z} *\<^sub>R z)" - using s(1)[THEN simple_functionD(1)] + using s[THEN simple_functionD(1)] unfolding simple_bochner_integral_def by (intro sum.mono_neutral_cong_left) (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) } note eq = this show ?thesis proof (rule borel_measurable_LIMSEQ_metric) fix i show "f' i \ borel_measurable N" unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong) next fix x assume x: "x \ space N" { assume int_f: "integrable M (f x)" have int_2f: "integrable M (\y. 2 * norm (f x y))" by (intro integrable_norm integrable_mult_right int_f) have "(\i. integral\<^sup>L M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" proof (rule integral_dominated_convergence) from int_f show "f x \ borel_measurable M" by auto show "\i. (\y. s i (x, y)) \ borel_measurable M" using x by simp show "AE xa in M. (\i. s i (x, xa)) \ f x xa" - using x s(2) by auto + using x * by auto show "\i. AE xa in M. norm (s i (x, xa)) \ 2 * norm (f x xa)" - using x s(3) by auto + using x * by auto qed fact moreover { fix i have "simple_bochner_integrable M (\y. s i (x, y))" proof (rule simple_bochner_integrableI_bounded) have "(\y. s i (x, y)) ` space M \ s i ` (space N \ space M)" using x by auto then show "simple_function M (\y. s i (x, y))" using simple_functionD(1)[OF s(1), of i] x by (intro simple_function_borel_measurable) (auto simp: space_pair_measure dest: finite_subset) have "(\\<^sup>+ y. ennreal (norm (s i (x, y))) \M) \ (\\<^sup>+ y. 2 * norm (f x y) \M)" - using x s by (intro nn_integral_mono) auto + using x * by (intro nn_integral_mono) auto also have "(\\<^sup>+ y. 2 * norm (f x y) \M) < \" using int_2f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ xa. ennreal (norm (s i (x, xa))) \M) < \" . qed then have "integral\<^sup>L M (\y. s i (x, y)) = simple_bochner_integral M (\y. s i (x, y))" by (rule simple_bochner_integrable_eq_integral[symmetric]) } ultimately have "(\i. simple_bochner_integral M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" by simp } then show "(\i. f' i x) \ integral\<^sup>L M (f x)" unfolding f'_def by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq) qed qed lemma (in pair_sigma_finite) integrable_product_swap: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "integrable (M1 \\<^sub>M M2) f" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x))" proof - interpret Q: pair_sigma_finite M2 M1 .. have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (rule integrable_distr[OF measurable_pair_swap']) (simp add: distr_pair_swap[symmetric] assms) qed lemma (in pair_sigma_finite) integrable_product_swap_iff: fixes f :: "_ \ _::{banach, second_countable_topology}" shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x)) \ integrable (M1 \\<^sub>M M2) f" proof - interpret Q: pair_sigma_finite M2 M1 .. from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] show ?thesis by auto qed lemma (in pair_sigma_finite) integral_product_swap: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\(x,y). f (y,x) \(M2 \\<^sub>M M1)) = integral\<^sup>L (M1 \\<^sub>M M2) f" proof - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) show ?thesis unfolding * by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) qed theorem (in pair_sigma_finite) Fubini_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" and integ1: "integrable M1 (\x. \ y. norm (f (x, y)) \M2)" and integ2: "AE x in M1. integrable M2 (\y. f (x, y))" shows "integrable (M1 \\<^sub>M M2) f" proof (rule integrableI_bounded) have "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M2)) = (\\<^sup>+ x. (\\<^sup>+ y. norm (f (x, y)) \M2) \M1)" by (simp add: M2.nn_integral_fst [symmetric]) also have "\ = (\\<^sup>+ x. \\y. norm (f (x, y)) \M2\ \M1)" apply (intro nn_integral_cong_AE) using integ2 proof eventually_elim fix x assume "integrable M2 (\y. f (x, y))" then have f: "integrable M2 (\y. norm (f (x, y)))" by simp then have "(\\<^sup>+y. ennreal (norm (f (x, y))) \M2) = ennreal (LINT y|M2. norm (f (x, y)))" by (rule nn_integral_eq_integral) simp also have "\ = ennreal \LINT y|M2. norm (f (x, y))\" using f by simp finally show "(\\<^sup>+y. ennreal (norm (f (x, y))) \M2) = ennreal \LINT y|M2. norm (f (x, y))\" . qed also have "\ < \" using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE) finally show "(\\<^sup>+ p. norm (f p) \(M1 \\<^sub>M M2)) < \" . qed fact lemma (in pair_sigma_finite) emeasure_pair_measure_finite: assumes A: "A \ sets (M1 \\<^sub>M M2)" and finite: "emeasure (M1 \\<^sub>M M2) A < \" shows "AE x in M1. emeasure M2 {y\space M2. (x, y) \ A} < \" proof - from M2.emeasure_pair_measure_alt[OF A] finite have "(\\<^sup>+ x. emeasure M2 (Pair x -` A) \M1) \ \" by simp then have "AE x in M1. emeasure M2 (Pair x -` A) \ \" by (rule nn_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A) moreover have "\x. x \ space M1 \ Pair x -` A = {y\space M2. (x, y) \ A}" using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) ultimately show ?thesis by (auto simp: less_top) qed lemma (in pair_sigma_finite) AE_integrable_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" shows "AE x in M1. integrable M2 (\y. f (x, y))" proof - have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" by (rule M2.nn_integral_fst) simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) \ \" using f unfolding integrable_iff_bounded by simp finally have "AE x in M1. (\\<^sup>+y. norm (f (x, y)) \M2) \ \" by (intro nn_integral_PInf_AE M2.borel_measurable_nn_integral ) (auto simp: measurable_split_conv) with AE_space show ?thesis by eventually_elim (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]] less_top) qed lemma (in pair_sigma_finite) integrable_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" shows "integrable M1 (\x. \y. f (x, y) \M2)" unfolding integrable_iff_bounded proof show "(\x. \ y. f (x, y) \M2) \ borel_measurable M1" by (rule M2.borel_measurable_lebesgue_integral) simp have "(\\<^sup>+ x. ennreal (norm (\ y. f (x, y) \M2)) \M1) \ (\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1)" using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal) also have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" by (rule M2.nn_integral_fst) simp also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) < \" using f unfolding integrable_iff_bounded by simp finally show "(\\<^sup>+ x. ennreal (norm (\ y. f (x, y) \M2)) \M1) < \" . qed proposition (in pair_sigma_finite) integral_fst': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) f" shows "(\x. (\y. f (x, y) \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) f" using f proof induct case (base A c) have A[measurable]: "A \ sets (M1 \\<^sub>M M2)" by fact have eq: "\x y. x \ space M1 \ indicator A (x, y) = indicator {y\space M2. (x, y) \ A} y" using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure) have int_A: "integrable (M1 \\<^sub>M M2) (indicator A :: _ \ real)" using base by (rule integrable_real_indicator) have "(\ x. \ y. indicator A (x, y) *\<^sub>R c \M2 \M1) = (\x. measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c \M1)" proof (intro integral_cong_AE, simp, simp) from AE_integrable_fst'[OF int_A] AE_space show "AE x in M1. (\y. indicator A (x, y) *\<^sub>R c \M2) = measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c" by eventually_elim (simp add: eq integrable_indicator_iff) qed also have "\ = measure (M1 \\<^sub>M M2) A *\<^sub>R c" proof (subst integral_scaleR_left) have "(\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = (\\<^sup>+x. emeasure M2 {y \ space M2. (x, y) \ A} \M1)" using emeasure_pair_measure_finite[OF base] by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure) also have "\ = emeasure (M1 \\<^sub>M M2) A" using sets.sets_into_space[OF A] by (subst M2.emeasure_pair_measure_alt) (auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure) finally have *: "(\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = emeasure (M1 \\<^sub>M M2) A" . from base * show "integrable M1 (\x. measure M2 {y \ space M2. (x, y) \ A})" by (simp add: integrable_iff_bounded) then have "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) = (\\<^sup>+x. ennreal (measure M2 {y \ space M2. (x, y) \ A}) \M1)" by (rule nn_integral_eq_integral[symmetric]) simp also note * finally show "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) *\<^sub>R c = measure (M1 \\<^sub>M M2) A *\<^sub>R c" using base by (simp add: emeasure_eq_ennreal_measure) qed also have "\ = (\ a. indicator A a *\<^sub>R c \(M1 \\<^sub>M M2))" using base by simp finally show ?case . next case (add f g) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "g \ borel_measurable (M1 \\<^sub>M M2)" by auto have "(\ x. \ y. f (x, y) + g (x, y) \M2 \M1) = (\ x. (\ y. f (x, y) \M2) + (\ y. g (x, y) \M2) \M1)" apply (rule integral_cong_AE) apply simp_all using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)] apply eventually_elim apply simp done also have "\ = (\ x. f x \(M1 \\<^sub>M M2)) + (\ x. g x \(M1 \\<^sub>M M2))" using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp finally show ?case using add by simp next case (lim f s) then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "\i. s i \ borel_measurable (M1 \\<^sub>M M2)" by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ integral\<^sup>L (M1 \\<^sub>M M2) f" proof (rule integral_dominated_convergence) show "integrable (M1 \\<^sub>M M2) (\x. 2 * norm (f x))" using lim(5) by auto qed (insert lim, auto) have "(\i. \ x. \ y. s i (x, y) \M2 \M1) \ \ x. \ y. f (x, y) \M2 \M1" proof (rule integral_dominated_convergence) have "AE x in M1. \i. integrable M2 (\y. s i (x, y))" unfolding AE_all_countable using AE_integrable_fst'[OF lim(1)] .. with AE_space AE_integrable_fst'[OF lim(5)] show "AE x in M1. (\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof eventually_elim fix x assume x: "x \ space M1" and s: "\i. integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" show "(\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof (rule integral_dominated_convergence) show "integrable M2 (\y. 2 * norm (f (x, y)))" using f by auto show "AE xa in M2. (\i. s i (x, xa)) \ f (x, xa)" using x lim(3) by (auto simp: space_pair_measure) show "\i. AE xa in M2. norm (s i (x, xa)) \ 2 * norm (f (x, xa))" using x lim(4) by (auto simp: space_pair_measure) qed (insert x, measurable) qed show "integrable M1 (\x. (\ y. 2 * norm (f (x, y)) \M2))" by (intro integrable_mult_right integrable_norm integrable_fst' lim) fix i show "AE x in M1. norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)] proof eventually_elim fix x assume x: "x \ space M1" and s: "integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" from s have "norm (\ y. s i (x, y) \M2) \ (\\<^sup>+y. norm (s i (x, y)) \M2)" by (rule integral_norm_bound_ennreal) also have "\ \ (\\<^sup>+y. 2 * norm (f (x, y)) \M2)" using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure) also have "\ = (\y. 2 * norm (f (x, y)) \M2)" using f by (intro nn_integral_eq_integral) auto finally show "norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" by simp qed qed simp_all then show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ \ x. \ y. f (x, y) \M2 \M1" using lim by simp qed qed lemma (in pair_sigma_finite) fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows AE_integrable_fst: "AE x in M1. integrable M2 (\y. f x y)" (is "?AE") and integrable_fst: "integrable M1 (\x. \y. f x y \M2)" (is "?INT") and integral_fst: "(\x. (\y. f x y \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) (\(x, y). f x y)" (is "?EQ") using AE_integrable_fst'[OF f] integrable_fst'[OF f] integral_fst'[OF f] by auto lemma (in pair_sigma_finite) fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" assumes f[measurable]: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows AE_integrable_snd: "AE y in M2. integrable M1 (\x. f x y)" (is "?AE") and integrable_snd: "integrable M2 (\y. \x. f x y \M1)" (is "?INT") and integral_snd: "(\y. (\x. f x y \M1) \M2) = integral\<^sup>L (M1 \\<^sub>M M2) (case_prod f)" (is "?EQ") proof - interpret Q: pair_sigma_finite M2 M1 .. have Q_int: "integrable (M2 \\<^sub>M M1) (\(x, y). f y x)" using f unfolding integrable_product_swap_iff[symmetric] by simp show ?AE using Q.AE_integrable_fst'[OF Q_int] by simp show ?INT using Q.integrable_fst'[OF Q_int] by simp show ?EQ using Q.integral_fst'[OF Q_int] using integral_product_swap[of "case_prod f"] by simp qed proposition (in pair_sigma_finite) Fubini_integral: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes f: "integrable (M1 \\<^sub>M M2) (case_prod f)" shows "(\y. (\x. f x y \M1) \M2) = (\x. (\y. f x y \M2) \M1)" unfolding integral_snd[OF assms] integral_fst[OF assms] .. lemma (in product_sigma_finite) product_integral_singleton: fixes f :: "_ \ _::{banach, second_countable_topology}" shows "f \ borel_measurable (M i) \ (\x. f (x i) \Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" apply (subst distr_singleton[symmetric]) apply (subst integral_distr) apply simp_all done lemma (in product_sigma_finite) product_integral_fold: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" and f: "integrable (Pi\<^sub>M (I \ J) M) f" shows "integral\<^sup>L (Pi\<^sub>M (I \ J) M) f = (\x. (\y. f (merge I J (x, y)) \Pi\<^sub>M J M) \Pi\<^sub>M I M)" proof - interpret I: finite_product_sigma_finite M I by standard fact interpret J: finite_product_sigma_finite M J by standard fact have "finite (I \ J)" using fin by auto interpret IJ: finite_product_sigma_finite M "I \ J" by standard fact interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" .. let ?M = "merge I J" let ?f = "\x. f (?M x)" from f have f_borel: "f \ borel_measurable (Pi\<^sub>M (I \ J) M)" by auto have P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)" using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) have f_int: "integrable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) ?f" by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) show ?thesis apply (subst distr_merge[symmetric, OF IJ fin]) apply (subst integral_distr[OF measurable_merge f_borel]) apply (subst P.integral_fst'[symmetric, OF f_int]) apply simp done qed lemma (in product_sigma_finite) product_integral_insert: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes I: "finite I" "i \ I" and f: "integrable (Pi\<^sub>M (insert i I) M) f" shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^sub>M I M)" proof - have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \ {i}) M) f" by simp also have "\ = (\x. (\y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) \Pi\<^sub>M I M)" using f I by (intro product_integral_fold) auto also have "\ = (\x. (\y. f (x(i := y)) \M i) \Pi\<^sub>M I M)" proof (rule integral_cong[OF refl], subst product_integral_singleton[symmetric]) fix x assume x: "x \ space (Pi\<^sub>M I M)" have f_borel: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" using f by auto show "(\y. f (x(i := y))) \ borel_measurable (M i)" using measurable_comp[OF measurable_component_update f_borel, OF x \i \ I\] unfolding comp_def . from x I show "(\ y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) = (\ xa. f (x(i := xa i)) \Pi\<^sub>M {i} M)" by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) qed finally show ?thesis . qed lemma (in product_sigma_finite) product_integrable_prod: fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" shows "integrable (Pi\<^sub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") proof (unfold integrable_iff_bounded, intro conjI) interpret finite_product_sigma_finite M I by standard fact show "?f \ borel_measurable (Pi\<^sub>M I M)" using assms by simp have "(\\<^sup>+ x. ennreal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) = (\\<^sup>+ x. (\i\I. ennreal (norm (f i (x i)))) \Pi\<^sub>M I M)" by (simp add: prod_norm prod_ennreal) also have "\ = (\i\I. \\<^sup>+ x. ennreal (norm (f i x)) \M i)" using assms by (intro product_nn_integral_prod) auto also have "\ < \" using integrable by (simp add: less_top[symmetric] ennreal_prod_eq_top integrable_iff_bounded) finally show "(\\<^sup>+ x. ennreal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) < \" . qed lemma (in product_sigma_finite) product_integral_prod: fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" assumes "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" shows "(\x. (\i\I. f i (x i)) \Pi\<^sub>M I M) = (\i\I. integral\<^sup>L (M i) (f i))" using assms proof induct case empty interpret finite_measure "Pi\<^sub>M {} M" by rule (simp add: space_PiM) show ?case by (simp add: space_PiM measure_def) next case (insert i I) then have iI: "finite (insert i I)" by auto then have prod: "\J. J \ insert i I \ integrable (Pi\<^sub>M J M) (\x. (\i\J. f i (x i)))" by (intro product_integrable_prod insert(4)) (auto intro: finite_subset) interpret I: finite_product_sigma_finite M I by standard fact have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using \i \ I\ by (auto intro!: prod.cong) show ?case unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] by (simp add: * insert prod subset_insertI) qed lemma integrable_subalgebra: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes borel: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integrable N f \ integrable M f" (is ?P) proof - have "f \ borel_measurable M" using assms by (auto simp: measurable_def) with assms show ?thesis using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra) qed lemma integral_subalgebra: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes borel: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integral\<^sup>L N f = integral\<^sup>L M f" proof cases assume "integrable N f" then show ?thesis proof induct case base with assms show ?case by (auto simp: subset_eq measure_def) next case (add f g) then have "(\ a. f a + g a \N) = integral\<^sup>L M f + integral\<^sup>L M g" by simp also have "\ = (\ a. f a + g a \M)" using add integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of g] by simp finally show ?case . next case (lim f s) then have M: "\i. integrable M (s i)" "integrable M f" using integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of "s i" for i] by simp_all show ?case proof (intro LIMSEQ_unique) show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L N f" apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim apply auto done show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L M f" unfolding lim apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim M N(2) apply auto done qed qed qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms]) hide_const (open) simple_bochner_integral hide_const (open) simple_bochner_integrable end diff --git a/src/HOL/Analysis/Borel_Space.thy b/src/HOL/Analysis/Borel_Space.thy --- a/src/HOL/Analysis/Borel_Space.thy +++ b/src/HOL/Analysis/Borel_Space.thy @@ -1,2086 +1,2086 @@ (* Title: HOL/Analysis/Borel_Space.thy Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Borel Space\ theory Borel_Space imports Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits begin lemma is_interval_real_ereal_oo: "is_interval (real_of_ereal ` {N<..i. {x\space M. P x i} \ sets M) \ {x\space M. eventually (P x) sequentially} \ sets M" unfolding eventually_sequentially by simp lemma topological_basis_trivial: "topological_basis {A. open A}" by (auto simp: topological_basis_def) proposition open_prod_generated: "open = generate_topology {A \ B | A B. open A \ open B}" proof - have "{A \ B :: ('a \ 'b) set | A B. open A \ open B} = ((\(a, b). a \ b) ` ({A. open A} \ {A. open A}))" by auto then show ?thesis by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis) qed proposition mono_on_imp_deriv_nonneg: assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)" assumes "x \ interior A" shows "D \ 0" proof (rule tendsto_lowerbound) let ?A' = "(\y. y - x) ` interior A" from deriv show "((\h. (f (x + h) - f x) / h) \ D) (at 0)" by (simp add: field_has_derivative_at has_field_derivative_def) from mono have mono': "mono_on f (interior A)" by (rule mono_on_subset) (rule interior_subset) show "eventually (\h. (f (x + h) - f x) / h \ 0) (at 0)" proof (subst eventually_at_topological, intro exI conjI ballI impI) have "open (interior A)" by simp hence "open ((+) (-x) ` interior A)" by (rule open_translation) also have "((+) (-x) ` interior A) = ?A'" by auto finally show "open ?A'" . next from \x \ interior A\ show "0 \ ?A'" by auto next fix h assume "h \ ?A'" hence "x + h \ interior A" by auto with mono' and \x \ interior A\ show "(f (x + h) - f x) / h \ 0" by (cases h rule: linorder_cases[of _ 0]) (simp_all add: divide_nonpos_neg divide_nonneg_pos mono_onD field_simps) qed qed simp proposition mono_on_ctble_discont: fixes f :: "real \ real" fixes A :: "real set" assumes "mono_on f A" shows "countable {a\A. \ continuous (at a within A) f}" proof - have mono: "\x y. x \ A \ y \ A \ x \ y \ f x \ f y" using \mono_on f A\ by (simp add: mono_on_def) have "\a \ {a\A. \ continuous (at a within A) f}. \q :: nat \ rat. (fst q = 0 \ of_rat (snd q) < f a \ (\x \ A. x < a \ f x < of_rat (snd q))) \ (fst q = 1 \ of_rat (snd q) > f a \ (\x \ A. x > a \ f x > of_rat (snd q)))" proof (clarsimp simp del: One_nat_def) fix a assume "a \ A" assume "\ continuous (at a within A) f" thus "\q1 q2. q1 = 0 \ real_of_rat q2 < f a \ (\x\A. x < a \ f x < real_of_rat q2) \ q1 = 1 \ f a < real_of_rat q2 \ (\x\A. a < x \ real_of_rat q2 < f x)" proof (auto simp add: continuous_within order_tendsto_iff eventually_at) fix l assume "l < f a" then obtain q2 where q2: "l < of_rat q2" "of_rat q2 < f a" using of_rat_dense by blast assume * [rule_format]: "\d>0. \x\A. x \ a \ dist x a < d \ \ l < f x" from q2 have "real_of_rat q2 < f a \ (\x\A. x < a \ f x < real_of_rat q2)" proof auto fix x assume "x \ A" "x < a" with q2 *[of "a - x"] show "f x < real_of_rat q2" apply (auto simp add: dist_real_def not_less) apply (subgoal_tac "f x \ f xa") by (auto intro: mono) qed thus ?thesis by auto next fix u assume "u > f a" then obtain q2 where q2: "f a < of_rat q2" "of_rat q2 < u" using of_rat_dense by blast assume *[rule_format]: "\d>0. \x\A. x \ a \ dist x a < d \ \ u > f x" from q2 have "real_of_rat q2 > f a \ (\x\A. x > a \ f x > real_of_rat q2)" proof auto fix x assume "x \ A" "x > a" with q2 *[of "x - a"] show "f x > real_of_rat q2" apply (auto simp add: dist_real_def) apply (subgoal_tac "f x \ f xa") by (auto intro: mono) qed thus ?thesis by auto qed qed - hence "\g :: real \ nat \ rat . \a \ {a\A. \ continuous (at a within A) f}. + then obtain g :: "real \ nat \ rat" where "\a \ {a\A. \ continuous (at a within A) f}. (fst (g a) = 0 \ of_rat (snd (g a)) < f a \ (\x \ A. x < a \ f x < of_rat (snd (g a)))) | (fst (g a) = 1 \ of_rat (snd (g a)) > f a \ (\x \ A. x > a \ f x > of_rat (snd (g a))))" - by (rule bchoice) - then guess g .. + by (rule bchoice [THEN exE]) blast hence g: "\a x. a \ A \ \ continuous (at a within A) f \ x \ A \ (fst (g a) = 0 \ of_rat (snd (g a)) < f a \ (x < a \ f x < of_rat (snd (g a)))) | (fst (g a) = 1 \ of_rat (snd (g a)) > f a \ (x > a \ f x > of_rat (snd (g a))))" by auto have "inj_on g {a\A. \ continuous (at a within A) f}" proof (auto simp add: inj_on_def) fix w z assume 1: "w \ A" and 2: "\ continuous (at w within A) f" and 3: "z \ A" and 4: "\ continuous (at z within A) f" and 5: "g w = g z" from g [OF 1 2 3] g [OF 3 4 1] 5 show "w = z" by auto qed thus ?thesis by (rule countableI') qed lemma mono_on_ctble_discont_open: fixes f :: "real \ real" fixes A :: "real set" assumes "open A" "mono_on f A" shows "countable {a\A. \isCont f a}" proof - have "{a\A. \isCont f a} = {a\A. \(continuous (at a within A) f)}" by (auto simp add: continuous_within_open [OF _ \open A\]) thus ?thesis apply (elim ssubst) by (rule mono_on_ctble_discont, rule assms) qed lemma mono_ctble_discont: fixes f :: "real \ real" assumes "mono f" shows "countable {a. \ isCont f a}" using assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto lemma has_real_derivative_imp_continuous_on: assumes "\x. x \ A \ (f has_real_derivative f' x) (at x)" shows "continuous_on A f" apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def) using assms differentiable_at_withinI real_differentiable_def by blast lemma continuous_interval_vimage_Int: assumes "continuous_on {a::real..b} g" and mono: "\x y. a \ x \ x \ y \ y \ b \ g x \ g y" assumes "a \ b" "(c::real) \ d" "{c..d} \ {g a..g b}" obtains c' d' where "{a..b} \ g -` {c..d} = {c'..d'}" "c' \ d'" "g c' = c" "g d' = d" proof- let ?A = "{a..b} \ g -` {c..d}" from IVT'[of g a c b, OF _ _ \a \ b\ assms(1)] assms(4,5) obtain c'' where c'': "c'' \ ?A" "g c'' = c" by auto from IVT'[of g a d b, OF _ _ \a \ b\ assms(1)] assms(4,5) obtain d'' where d'': "d'' \ ?A" "g d'' = d" by auto hence [simp]: "?A \ {}" by blast define c' where "c' = Inf ?A" define d' where "d' = Sup ?A" have "?A \ {c'..d'}" unfolding c'_def d'_def by (intro subsetI) (auto intro: cInf_lower cSup_upper) moreover from assms have "closed ?A" using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp hence c'd'_in_set: "c' \ ?A" "d' \ ?A" unfolding c'_def d'_def by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+ hence "{c'..d'} \ ?A" using assms by (intro subsetI) (auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x] intro!: mono) moreover have "c' \ d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto moreover have "g c' \ c" "g d' \ d" apply (insert c'' d'' c'd'_in_set) apply (subst c''(2)[symmetric]) apply (auto simp: c'_def intro!: mono cInf_lower c'') [] apply (subst d''(2)[symmetric]) apply (auto simp: d'_def intro!: mono cSup_upper d'') [] done with c'd'_in_set have "g c' = c" "g d' = d" by auto ultimately show ?thesis using that by blast qed subsection \Generic Borel spaces\ definition\<^marker>\tag important\ (in topological_space) borel :: "'a measure" where "borel = sigma UNIV {S. open S}" abbreviation "borel_measurable M \ measurable M borel" lemma in_borel_measurable: "f \ borel_measurable M \ (\S \ sigma_sets UNIV {S. open S}. f -` S \ space M \ sets M)" by (auto simp add: measurable_def borel_def) lemma in_borel_measurable_borel: "f \ borel_measurable M \ (\S \ sets borel. f -` S \ space M \ sets M)" by (auto simp add: measurable_def borel_def) lemma space_borel[simp]: "space borel = UNIV" unfolding borel_def by auto lemma space_in_borel[measurable]: "UNIV \ sets borel" unfolding borel_def by auto lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}" unfolding borel_def by (rule sets_measure_of) simp lemma measurable_sets_borel: "\f \ measurable borel M; A \ sets M\ \ f -` A \ sets borel" by (drule (1) measurable_sets) simp lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \ {x. P x} \ sets borel" unfolding borel_def pred_def by auto lemma borel_open[measurable (raw generic)]: assumes "open A" shows "A \ sets borel" proof - have "A \ {S. open S}" unfolding mem_Collect_eq using assms . thus ?thesis unfolding borel_def by auto qed lemma borel_closed[measurable (raw generic)]: assumes "closed A" shows "A \ sets borel" proof - have "space borel - (- A) \ sets borel" using assms unfolding closed_def by (blast intro: borel_open) thus ?thesis by simp qed lemma borel_singleton[measurable]: "A \ sets borel \ insert x A \ sets (borel :: 'a::t1_space measure)" unfolding insert_def by (rule sets.Un) auto lemma sets_borel_eq_count_space: "sets (borel :: 'a::{countable, t2_space} measure) = count_space UNIV" proof - have "(\a\A. {a}) \ sets borel" for A :: "'a set" by (intro sets.countable_UN') auto then show ?thesis by auto qed lemma borel_comp[measurable]: "A \ sets borel \ - A \ sets borel" unfolding Compl_eq_Diff_UNIV by simp lemma borel_measurable_vimage: fixes f :: "'a \ 'x::t2_space" assumes borel[measurable]: "f \ borel_measurable M" shows "f -` {x} \ space M \ sets M" by simp lemma borel_measurableI: fixes f :: "'a \ 'x::topological_space" assumes "\S. open S \ f -` S \ space M \ sets M" shows "f \ borel_measurable M" unfolding borel_def proof (rule measurable_measure_of, simp_all) fix S :: "'x set" assume "open S" thus "f -` S \ space M \ sets M" using assms[of S] by simp qed lemma borel_measurable_const: "(\x. c) \ borel_measurable M" by auto lemma borel_measurable_indicator: assumes A: "A \ sets M" shows "indicator A \ borel_measurable M" unfolding indicator_def [abs_def] using A by (auto intro!: measurable_If_set) lemma borel_measurable_count_space[measurable (raw)]: "f \ borel_measurable (count_space S)" unfolding measurable_def by auto lemma borel_measurable_indicator'[measurable (raw)]: assumes [measurable]: "{x\space M. f x \ A x} \ sets M" shows "(\x. indicator (A x) (f x)) \ borel_measurable M" unfolding indicator_def[abs_def] by (auto intro!: measurable_If) lemma borel_measurable_indicator_iff: "(indicator A :: 'a \ 'x::{t1_space, zero_neq_one}) \ borel_measurable M \ A \ space M \ sets M" (is "?I \ borel_measurable M \ _") proof assume "?I \ borel_measurable M" then have "?I -` {1} \ space M \ sets M" unfolding measurable_def by auto also have "?I -` {1} \ space M = A \ space M" unfolding indicator_def [abs_def] by auto finally show "A \ space M \ sets M" . next assume "A \ space M \ sets M" moreover have "?I \ borel_measurable M \ (indicator (A \ space M) :: 'a \ 'x) \ borel_measurable M" by (intro measurable_cong) (auto simp: indicator_def) ultimately show "?I \ borel_measurable M" by auto qed lemma borel_measurable_subalgebra: assumes "sets N \ sets M" "space N = space M" "f \ borel_measurable N" shows "f \ borel_measurable M" using assms unfolding measurable_def by auto lemma borel_measurable_restrict_space_iff_ereal: fixes f :: "'a \ ereal" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. f x * indicator \ x) \ borel_measurable M" by (subst measurable_restrict_space_iff) (auto simp: indicator_def of_bool_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) lemma borel_measurable_restrict_space_iff_ennreal: fixes f :: "'a \ ennreal" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. f x * indicator \ x) \ borel_measurable M" by (subst measurable_restrict_space_iff) (auto simp: indicator_def of_bool_def if_distrib[where f="\x. a * x" for a] cong del: if_weak_cong) lemma borel_measurable_restrict_space_iff: fixes f :: "'a \ 'b::real_normed_vector" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ (\x. indicator \ x *\<^sub>R f x) \ borel_measurable M" by (subst measurable_restrict_space_iff) (auto simp: indicator_def of_bool_def if_distrib[where f="\x. x *\<^sub>R a" for a] ac_simps cong del: if_weak_cong) lemma cbox_borel[measurable]: "cbox a b \ sets borel" by (auto intro: borel_closed) lemma box_borel[measurable]: "box a b \ sets borel" by (auto intro: borel_open) lemma borel_compact: "compact (A::'a::t2_space set) \ A \ sets borel" by (auto intro: borel_closed dest!: compact_imp_closed) lemma borel_sigma_sets_subset: "A \ sets borel \ sigma_sets UNIV A \ sets borel" using sets.sigma_sets_subset[of A borel] by simp lemma borel_eq_sigmaI1: fixes F :: "'i \ 'a::topological_space set" and X :: "'a::topological_space set set" assumes borel_eq: "borel = sigma UNIV X" assumes X: "\x. x \ X \ x \ sets (sigma UNIV (F ` A))" assumes F: "\i. i \ A \ F i \ sets borel" shows "borel = sigma UNIV (F ` A)" unfolding borel_def proof (intro sigma_eqI antisym) have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel" unfolding borel_def by simp also have "\ = sigma_sets UNIV X" unfolding borel_eq by simp also have "\ \ sigma_sets UNIV (F`A)" using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto finally show "sigma_sets UNIV {S. open S} \ sigma_sets UNIV (F`A)" . show "sigma_sets UNIV (F`A) \ sigma_sets UNIV {S. open S}" unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto qed auto lemma borel_eq_sigmaI2: fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'k \ 'a::topological_space set" assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`B)" assumes X: "\i j. (i, j) \ B \ G i j \ sets (sigma UNIV ((\(i, j). F i j) ` A))" assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" using assms by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` B" and F="(\(i, j). F i j)"]) auto lemma borel_eq_sigmaI3: fixes F :: "'i \ 'j \ 'a::topological_space set" and X :: "'a::topological_space set set" assumes borel_eq: "borel = sigma UNIV X" assumes X: "\x. x \ X \ x \ sets (sigma UNIV ((\(i, j). F i j) ` A))" assumes F: "\i j. (i, j) \ A \ F i j \ sets borel" shows "borel = sigma UNIV ((\(i, j). F i j) ` A)" using assms by (intro borel_eq_sigmaI1[where X=X and F="(\(i, j). F i j)"]) auto lemma borel_eq_sigmaI4: fixes F :: "'i \ 'a::topological_space set" and G :: "'l \ 'k \ 'a::topological_space set" assumes borel_eq: "borel = sigma UNIV ((\(i, j). G i j)`A)" assumes X: "\i j. (i, j) \ A \ G i j \ sets (sigma UNIV (range F))" assumes F: "\i. F i \ sets borel" shows "borel = sigma UNIV (range F)" using assms by (intro borel_eq_sigmaI1[where X="(\(i, j). G i j) ` A" and F=F]) auto lemma borel_eq_sigmaI5: fixes F :: "'i \ 'j \ 'a::topological_space set" and G :: "'l \ 'a::topological_space set" assumes borel_eq: "borel = sigma UNIV (range G)" assumes X: "\i. G i \ sets (sigma UNIV (range (\(i, j). F i j)))" assumes F: "\i j. F i j \ sets borel" shows "borel = sigma UNIV (range (\(i, j). F i j))" using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\(i, j). F i j)"]) auto theorem second_countable_borel_measurable: fixes X :: "'a::second_countable_topology set set" assumes eq: "open = generate_topology X" shows "borel = sigma UNIV X" unfolding borel_def proof (intro sigma_eqI sigma_sets_eqI) interpret X: sigma_algebra UNIV "sigma_sets UNIV X" by (rule sigma_algebra_sigma_sets) simp fix S :: "'a set" assume "S \ Collect open" then have "generate_topology X S" by (auto simp: eq) then show "S \ sigma_sets UNIV X" proof induction case (UN K) then have K: "\k. k \ K \ open k" unfolding eq by auto from ex_countable_basis obtain B :: "'a set set" where B: "\b. b \ B \ open b" "\X. open X \ \b\B. (\b) = X" and "countable B" by (auto simp: topological_basis_def) from B(2)[OF K] obtain m where m: "\k. k \ K \ m k \ B" "\k. k \ K \ \(m k) = k" by metis define U where "U = (\k\K. m k)" with m have "countable U" by (intro countable_subset[OF _ \countable B\]) auto have "\U = (\A\U. A)" by simp also have "\ = \K" unfolding U_def UN_simps by (simp add: m) finally have "\U = \K" . have "\b\U. \k\K. b \ k" using m by (auto simp: U_def) then obtain u where u: "\b. b \ U \ u b \ K" and "\b. b \ U \ b \ u b" by metis then have "(\b\U. u b) \ \K" "\U \ (\b\U. u b)" by auto then have "\K = (\b\U. u b)" unfolding \\U = \K\ by auto also have "\ \ sigma_sets UNIV X" using u UN by (intro X.countable_UN' \countable U\) auto finally show "\K \ sigma_sets UNIV X" . qed auto qed (auto simp: eq intro: generate_topology.Basis) lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)" unfolding borel_def proof (intro sigma_eqI sigma_sets_eqI, safe) fix x :: "'a set" assume "open x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect closed)" by (force intro: sigma_sets.Compl simp: \open x\) finally show "x \ sigma_sets UNIV (Collect closed)" by simp next fix x :: "'a set" assume "closed x" hence "x = UNIV - (UNIV - x)" by auto also have "\ \ sigma_sets UNIV (Collect open)" by (force intro: sigma_sets.Compl simp: \closed x\) finally show "x \ sigma_sets UNIV (Collect open)" by simp qed simp_all proposition borel_eq_countable_basis: fixes B::"'a::topological_space set set" assumes "countable B" assumes "topological_basis B" shows "borel = sigma UNIV B" unfolding borel_def proof (intro sigma_eqI sigma_sets_eqI, safe) interpret countable_basis "open" B using assms by (rule countable_basis_openI) fix X::"'a set" assume "open X" from open_countable_basisE[OF this] obtain B' where B': "B' \ B" "X = \ B'" . then show "X \ sigma_sets UNIV B" by (blast intro: sigma_sets_UNION \countable B\ countable_subset) next fix b assume "b \ B" hence "open b" by (rule topological_basis_open[OF assms(2)]) thus "b \ sigma_sets UNIV (Collect open)" by auto qed simp_all lemma borel_measurable_continuous_on_restrict: fixes f :: "'a::topological_space \ 'b::topological_space" assumes f: "continuous_on A f" shows "f \ borel_measurable (restrict_space borel A)" proof (rule borel_measurableI) fix S :: "'b set" assume "open S" with f obtain T where "f -` S \ A = T \ A" "open T" by (metis continuous_on_open_invariant) then show "f -` S \ space (restrict_space borel A) \ sets (restrict_space borel A)" by (force simp add: sets_restrict_space space_restrict_space) qed lemma borel_measurable_continuous_onI: "continuous_on UNIV f \ f \ borel_measurable borel" by (drule borel_measurable_continuous_on_restrict) simp lemma borel_measurable_continuous_on_if: "A \ sets borel \ continuous_on A f \ continuous_on (- A) g \ (\x. if x \ A then f x else g x) \ borel_measurable borel" by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq intro!: borel_measurable_continuous_on_restrict) lemma borel_measurable_continuous_countable_exceptions: fixes f :: "'a::t1_space \ 'b::topological_space" assumes X: "countable X" assumes "continuous_on (- X) f" shows "f \ borel_measurable borel" proof (rule measurable_discrete_difference[OF _ X]) have "X \ sets borel" by (rule sets.countable[OF _ X]) auto then show "(\x. if x \ X then undefined else f x) \ borel_measurable borel" by (intro borel_measurable_continuous_on_if assms continuous_intros) qed auto lemma borel_measurable_continuous_on: assumes f: "continuous_on UNIV f" and g: "g \ borel_measurable M" shows "(\x. f (g x)) \ borel_measurable M" using measurable_comp[OF g borel_measurable_continuous_onI[OF f]] by (simp add: comp_def) lemma borel_measurable_continuous_on_indicator: fixes f g :: "'a::topological_space \ 'b::real_normed_vector" shows "A \ sets borel \ continuous_on A f \ (\x. indicator A x *\<^sub>R f x) \ borel_measurable borel" by (subst borel_measurable_restrict_space_iff[symmetric]) (auto intro: borel_measurable_continuous_on_restrict) lemma borel_measurable_Pair[measurable (raw)]: fixes f :: "'a \ 'b::second_countable_topology" and g :: "'a \ 'c::second_countable_topology" assumes f[measurable]: "f \ borel_measurable M" assumes g[measurable]: "g \ borel_measurable M" shows "(\x. (f x, g x)) \ borel_measurable M" proof (subst borel_eq_countable_basis) let ?B = "SOME B::'b set set. countable B \ topological_basis B" let ?C = "SOME B::'c set set. countable B \ topological_basis B" let ?P = "(\(b, c). b \ c) ` (?B \ ?C)" show "countable ?P" "topological_basis ?P" by (auto intro!: countable_basis topological_basis_prod is_basis) show "(\x. (f x, g x)) \ measurable M (sigma UNIV ?P)" proof (rule measurable_measure_of) fix S assume "S \ ?P" then obtain b c where "b \ ?B" "c \ ?C" and S: "S = b \ c" by auto then have borel: "open b" "open c" by (auto intro: is_basis topological_basis_open) have "(\x. (f x, g x)) -` S \ space M = (f -` b \ space M) \ (g -` c \ space M)" unfolding S by auto also have "\ \ sets M" using borel by simp finally show "(\x. (f x, g x)) -` S \ space M \ sets M" . qed auto qed lemma borel_measurable_continuous_Pair: fixes f :: "'a \ 'b::second_countable_topology" and g :: "'a \ 'c::second_countable_topology" assumes [measurable]: "f \ borel_measurable M" assumes [measurable]: "g \ borel_measurable M" assumes H: "continuous_on UNIV (\x. H (fst x) (snd x))" shows "(\x. H (f x) (g x)) \ borel_measurable M" proof - have eq: "(\x. H (f x) (g x)) = (\x. (\x. H (fst x) (snd x)) (f x, g x))" by auto show ?thesis unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto qed subsection \Borel spaces on order topologies\ lemma [measurable]: fixes a b :: "'a::linorder_topology" shows lessThan_borel: "{..< a} \ sets borel" and greaterThan_borel: "{a <..} \ sets borel" and greaterThanLessThan_borel: "{a<.. sets borel" and atMost_borel: "{..a} \ sets borel" and atLeast_borel: "{a..} \ sets borel" and atLeastAtMost_borel: "{a..b} \ sets borel" and greaterThanAtMost_borel: "{a<..b} \ sets borel" and atLeastLessThan_borel: "{a.. sets borel" unfolding greaterThanAtMost_def atLeastLessThan_def by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan closed_atMost closed_atLeast closed_atLeastAtMost)+ lemma borel_Iio: "borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)" unfolding second_countable_borel_measurable[OF open_generated_order] proof (intro sigma_eqI sigma_sets_eqI) - from countable_dense_setE guess D :: "'a set" . note D = this + obtain D :: "'a set" where D: "countable D" "\X. open X \ X \ {} \ \d\D. d \ X" + by (rule countable_dense_setE) blast interpret L: sigma_algebra UNIV "sigma_sets UNIV (range lessThan)" by (rule sigma_algebra_sigma_sets) simp fix A :: "'a set" assume "A \ range lessThan \ range greaterThan" then obtain y where "A = {y <..} \ A = {..< y}" by blast then show "A \ sigma_sets UNIV (range lessThan)" proof assume A: "A = {y <..}" show ?thesis proof cases assume "\x>y. \d. y < d \ d < x" with D(2)[of "{y <..< x}" for x] have "\x>y. \d\D. y < d \ d < x" by (auto simp: set_eq_iff) then have "A = UNIV - (\d\{d\D. y < d}. {..< d})" by (auto simp: A) (metis less_asym) also have "\ \ sigma_sets UNIV (range lessThan)" using D(1) by (intro L.Diff L.top L.countable_INT'') auto finally show ?thesis . next assume "\ (\x>y. \d. y < d \ d < x)" then obtain x where "y < x" "\d. y < d \ \ d < x" by auto then have "A = UNIV - {..< x}" unfolding A by (auto simp: not_less[symmetric]) also have "\ \ sigma_sets UNIV (range lessThan)" by auto finally show ?thesis . qed qed auto qed auto lemma borel_Ioi: "borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)" unfolding second_countable_borel_measurable[OF open_generated_order] proof (intro sigma_eqI sigma_sets_eqI) - from countable_dense_setE guess D :: "'a set" . note D = this + obtain D :: "'a set" where D: "countable D" "\X. open X \ X \ {} \ \d\D. d \ X" + by (rule countable_dense_setE) blast interpret L: sigma_algebra UNIV "sigma_sets UNIV (range greaterThan)" by (rule sigma_algebra_sigma_sets) simp fix A :: "'a set" assume "A \ range lessThan \ range greaterThan" then obtain y where "A = {y <..} \ A = {..< y}" by blast then show "A \ sigma_sets UNIV (range greaterThan)" proof assume A: "A = {..< y}" show ?thesis proof cases assume "\xd. x < d \ d < y" with D(2)[of "{x <..< y}" for x] have "\xd\D. x < d \ d < y" by (auto simp: set_eq_iff) then have "A = UNIV - (\d\{d\D. d < y}. {d <..})" by (auto simp: A) (metis less_asym) also have "\ \ sigma_sets UNIV (range greaterThan)" using D(1) by (intro L.Diff L.top L.countable_INT'') auto finally show ?thesis . next assume "\ (\xd. x < d \ d < y)" then obtain x where "x < y" "\d. y > d \ x \ d" by (auto simp: not_less[symmetric]) then have "A = UNIV - {x <..}" unfolding A Compl_eq_Diff_UNIV[symmetric] by auto also have "\ \ sigma_sets UNIV (range greaterThan)" by auto finally show ?thesis . qed qed auto qed auto lemma borel_measurableI_less: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. f x < y} \ sets M) \ f \ borel_measurable M" unfolding borel_Iio by (rule measurable_measure_of) (auto simp: Int_def conj_commute) lemma borel_measurableI_greater: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. y < f x} \ sets M) \ f \ borel_measurable M" unfolding borel_Ioi by (rule measurable_measure_of) (auto simp: Int_def conj_commute) lemma borel_measurableI_le: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. f x \ y} \ sets M) \ f \ borel_measurable M" by (rule borel_measurableI_greater) (auto simp: not_le[symmetric]) lemma borel_measurableI_ge: fixes f :: "'a \ 'b::{linorder_topology, second_countable_topology}" shows "(\y. {x\space M. y \ f x} \ sets M) \ f \ borel_measurable M" by (rule borel_measurableI_less) (auto simp: not_le[symmetric]) lemma borel_measurable_less[measurable]: fixes f :: "'a \ 'b::{second_countable_topology, linorder_topology}" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "{w \ space M. f w < g w} \ sets M" proof - have "{w \ space M. f w < g w} = (\x. (f x, g x)) -` {x. fst x < snd x} \ space M" by auto also have "\ \ sets M" by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less] continuous_intros) finally show ?thesis . qed lemma fixes f :: "'a \ 'b::{second_countable_topology, linorder_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes g[measurable]: "g \ borel_measurable M" shows borel_measurable_le[measurable]: "{w \ space M. f w \ g w} \ sets M" and borel_measurable_eq[measurable]: "{w \ space M. f w = g w} \ sets M" and borel_measurable_neq: "{w \ space M. f w \ g w} \ sets M" unfolding eq_iff not_less[symmetric] by measurable lemma borel_measurable_SUP[measurable (raw)]: fixes F :: "_ \ _ \ _::{complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" shows "(\x. SUP i\I. F i x) \ borel_measurable M" by (rule borel_measurableI_greater) (simp add: less_SUP_iff) lemma borel_measurable_INF[measurable (raw)]: fixes F :: "_ \ _ \ _::{complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" shows "(\x. INF i\I. F i x) \ borel_measurable M" by (rule borel_measurableI_less) (simp add: INF_less_iff) lemma borel_measurable_cSUP[measurable (raw)]: fixes F :: "_ \ _ \ 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" assumes bdd: "\x. x \ space M \ bdd_above ((\i. F i x) ` I)" shows "(\x. SUP i\I. F i x) \ borel_measurable M" proof cases assume "I = {}" then show ?thesis unfolding \I = {}\ image_empty by simp next assume "I \ {}" show ?thesis proof (rule borel_measurableI_le) fix y have "{x \ space M. \i\I. F i x \ y} \ sets M" by measurable also have "{x \ space M. \i\I. F i x \ y} = {x \ space M. (SUP i\I. F i x) \ y}" by (simp add: cSUP_le_iff \I \ {}\ bdd cong: conj_cong) finally show "{x \ space M. (SUP i\I. F i x) \ y} \ sets M" . qed qed lemma borel_measurable_cINF[measurable (raw)]: fixes F :: "_ \ _ \ 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}" assumes [simp]: "countable I" assumes [measurable]: "\i. i \ I \ F i \ borel_measurable M" assumes bdd: "\x. x \ space M \ bdd_below ((\i. F i x) ` I)" shows "(\x. INF i\I. F i x) \ borel_measurable M" proof cases assume "I = {}" then show ?thesis unfolding \I = {}\ image_empty by simp next assume "I \ {}" show ?thesis proof (rule borel_measurableI_ge) fix y have "{x \ space M. \i\I. y \ F i x} \ sets M" by measurable also have "{x \ space M. \i\I. y \ F i x} = {x \ space M. y \ (INF i\I. F i x)}" by (simp add: le_cINF_iff \I \ {}\ bdd cong: conj_cong) finally show "{x \ space M. y \ (INF i\I. F i x)} \ sets M" . qed qed lemma borel_measurable_lfp[consumes 1, case_names continuity step]: fixes F :: "('a \ 'b) \ ('a \ 'b::{complete_linorder, linorder_topology, second_countable_topology})" assumes "sup_continuous F" assumes *: "\f. f \ borel_measurable M \ F f \ borel_measurable M" shows "lfp F \ borel_measurable M" proof - { fix i have "((F ^^ i) bot) \ borel_measurable M" by (induct i) (auto intro!: *) } then have "(\x. SUP i. (F ^^ i) bot x) \ borel_measurable M" by measurable also have "(\x. SUP i. (F ^^ i) bot x) = (SUP i. (F ^^ i) bot)" by (auto simp add: image_comp) also have "(SUP i. (F ^^ i) bot) = lfp F" by (rule sup_continuous_lfp[symmetric]) fact finally show ?thesis . qed lemma borel_measurable_gfp[consumes 1, case_names continuity step]: fixes F :: "('a \ 'b) \ ('a \ 'b::{complete_linorder, linorder_topology, second_countable_topology})" assumes "inf_continuous F" assumes *: "\f. f \ borel_measurable M \ F f \ borel_measurable M" shows "gfp F \ borel_measurable M" proof - { fix i have "((F ^^ i) top) \ borel_measurable M" by (induct i) (auto intro!: * simp: bot_fun_def) } then have "(\x. INF i. (F ^^ i) top x) \ borel_measurable M" by measurable also have "(\x. INF i. (F ^^ i) top x) = (INF i. (F ^^ i) top)" by (auto simp add: image_comp) also have "\ = gfp F" by (rule inf_continuous_gfp[symmetric]) fact finally show ?thesis . qed lemma borel_measurable_max[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" by (rule borel_measurableI_less) simp lemma borel_measurable_min[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" by (rule borel_measurableI_greater) simp lemma borel_measurable_Min[measurable (raw)]: "finite I \ (\i. i \ I \ f i \ borel_measurable M) \ (\x. Min ((\i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" proof (induct I rule: finite_induct) case (insert i I) then show ?case by (cases "I = {}") auto qed auto lemma borel_measurable_Max[measurable (raw)]: "finite I \ (\i. i \ I \ f i \ borel_measurable M) \ (\x. Max ((\i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \ borel_measurable M" proof (induct I rule: finite_induct) case (insert i I) then show ?case by (cases "I = {}") auto qed auto lemma borel_measurable_sup[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \ borel_measurable M" unfolding sup_max by measurable lemma borel_measurable_inf[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \ borel_measurable M" unfolding inf_min by measurable lemma [measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes "\i. f i \ borel_measurable M" shows borel_measurable_liminf: "(\x. liminf (\i. f i x)) \ borel_measurable M" and borel_measurable_limsup: "(\x. limsup (\i. f i x)) \ borel_measurable M" unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto lemma measurable_convergent[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes [measurable]: "\i. f i \ borel_measurable M" shows "Measurable.pred M (\x. convergent (\i. f i x))" unfolding convergent_ereal by measurable lemma sets_Collect_convergent[measurable]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "{x\space M. convergent (\i. f i x)} \ sets M" by measurable lemma borel_measurable_lim[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes [measurable]: "\i. f i \ borel_measurable M" shows "(\x. lim (\i. f i x)) \ borel_measurable M" proof - have "\x. lim (\i. f i x) = (if convergent (\i. f i x) then limsup (\i. f i x) else (THE i. False))" by (simp add: lim_def convergent_def convergent_limsup_cl) then show ?thesis by simp qed lemma borel_measurable_LIMSEQ_order: fixes u :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology}" assumes u': "\x. x \ space M \ (\i. u i x) \ u' x" and u: "\i. u i \ borel_measurable M" shows "u' \ borel_measurable M" proof - have "\x. x \ space M \ u' x = liminf (\n. u n x)" using u' by (simp add: lim_imp_Liminf[symmetric]) with u show ?thesis by (simp cong: measurable_cong) qed subsection \Borel spaces on topological monoids\ lemma borel_measurable_add[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, topological_monoid_add}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x + g x) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_sum[measurable (raw)]: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, topological_comm_monoid_add}" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" proof cases assume "finite S" thus ?thesis using assms by induct auto qed simp lemma borel_measurable_suminf_order[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "(\x. suminf (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp subsection \Borel spaces on Euclidean spaces\ lemma borel_measurable_inner[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, real_inner}" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. f x \ g x) \ borel_measurable M" using assms by (rule borel_measurable_continuous_Pair) (intro continuous_intros) notation eucl_less (infix " x \ b} = {x. a {..b}" and box_co: "{x. a \ x \ x {x. x sets borel" and "{x. a sets borel" and "{..a} \ sets borel" and "{a..} \ sets borel" and "{a..b} \ sets borel" and "{x. a x \ b} \ sets borel" and "{x. a \ x \ x sets borel" unfolding box_oc box_co by (auto intro: borel_open borel_closed) lemma fixes i :: "'a::{second_countable_topology, real_inner}" shows hafspace_less_borel: "{x. a < x \ i} \ sets borel" and hafspace_greater_borel: "{x. x \ i < a} \ sets borel" and hafspace_less_eq_borel: "{x. a \ x \ i} \ sets borel" and hafspace_greater_eq_borel: "{x. x \ i \ a} \ sets borel" by simp_all lemma borel_eq_box: "borel = sigma UNIV (range (\ (a, b). box a b :: 'a :: euclidean_space set))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI1[OF borel_def]) fix M :: "'a set" assume "M \ {S. open S}" then have "open M" by simp show "M \ ?SIGMA" apply (subst open_UNION_box[OF \open M\]) apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect) apply (auto intro: countable_rat) done qed (auto simp: box_def) lemma halfspace_gt_in_halfspace: assumes i: "i \ A" shows "{x::'a. a < x \ i} \ sigma_sets UNIV ((\ (a, i). {x::'a::euclidean_space. x \ i < a}) ` (UNIV \ A))" (is "?set \ ?SIGMA") proof - interpret sigma_algebra UNIV ?SIGMA by (intro sigma_algebra_sigma_sets) simp_all have *: "?set = (\n. UNIV - {x::'a. x \ i < a + 1 / real (Suc n)})" proof (safe, simp_all add: not_less del: of_nat_Suc) fix x :: 'a assume "a < x \ i" with reals_Archimedean[of "x \ i - a"] obtain n where "a + 1 / real (Suc n) < x \ i" by (auto simp: field_simps) then show "\n. a + 1 / real (Suc n) \ x \ i" by (blast intro: less_imp_le) next fix x n have "a < a + 1 / real (Suc n)" by auto also assume "\ \ x" finally show "a < x" . qed show "?set \ ?SIGMA" unfolding * by (auto intro!: Diff sigma_sets_Inter i) qed lemma borel_eq_halfspace_less: "borel = sigma UNIV ((\(a, i). {x::'a::euclidean_space. x \ i < a}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_box]) fix a b :: 'a have "box a b = {x\space ?SIGMA. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" by (auto simp: box_def) also have "\ \ sets ?SIGMA" by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const) (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat) finally show "box a b \ sets ?SIGMA" . qed auto lemma borel_eq_halfspace_le: "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. x \ i \ a}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" then have i: "i \ Basis" by auto have *: "{x::'a. x\i < a} = (\n. {x. x\i \ a - 1/real (Suc n)})" proof (safe, simp_all del: of_nat_Suc) fix x::'a assume *: "x\i < a" with reals_Archimedean[of "a - x\i"] obtain n where "x \ i < a - 1 / (real (Suc n))" by (auto simp: field_simps) then show "\n. x \ i \ a - 1 / (real (Suc n))" by (blast intro: less_imp_le) next fix x::'a and n assume "x\i \ a - 1 / real (Suc n)" also have "\ < a" by auto finally show "x\i < a" . qed show "{x. x\i < a} \ ?SIGMA" unfolding * by (intro sets.countable_UN) (auto intro: i) qed auto lemma borel_eq_halfspace_ge: "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. a \ x \ i}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less]) fix a :: real and i :: 'a assume i: "(a, i) \ UNIV \ Basis" have *: "{x::'a. x\i < a} = space ?SIGMA - {x::'a. a \ x\i}" by auto show "{x. x\i < a} \ ?SIGMA" unfolding * using i by (intro sets.compl_sets) auto qed auto lemma borel_eq_halfspace_greater: "borel = sigma UNIV ((\ (a, i). {x::'a::euclidean_space. a < x \ i}) ` (UNIV \ Basis))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le]) fix a :: real and i :: 'a assume "(a, i) \ (UNIV \ Basis)" then have i: "i \ Basis" by auto have *: "{x::'a. x\i \ a} = space ?SIGMA - {x::'a. a < x\i}" by auto show "{x. x\i \ a} \ ?SIGMA" unfolding * by (intro sets.compl_sets) (auto intro: i) qed auto lemma borel_eq_atMost: "borel = sigma UNIV (range (\a. {..a::'a::ordered_euclidean_space}))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le]) fix a :: real and i :: 'a assume "(a, i) \ UNIV \ Basis" then have "i \ Basis" by auto then have *: "{x::'a. x\i \ a} = (\k::nat. {.. (\n\Basis. (if n = i then a else real k)*\<^sub>R n)})" proof (safe, simp_all add: eucl_le[where 'a='a] split: if_split_asm) fix x :: 'a - from real_arch_simple[of "Max ((\i. x\i)`Basis)"] guess k::nat .. + obtain k where "Max ((\) x ` Basis) \ real k" + using real_arch_simple by blast then have "\i. i \ Basis \ x\i \ real k" by (subst (asm) Max_le_iff) auto then show "\k::nat. \ia\Basis. ia \ i \ x \ ia \ real k" by (auto intro!: exI[of _ k]) qed show "{x. x\i \ a} \ ?SIGMA" unfolding * by (intro sets.countable_UN) auto qed auto lemma borel_eq_greaterThan: "borel = sigma UNIV (range (\a::'a::ordered_euclidean_space. {x. a UNIV \ Basis" then have i: "i \ Basis" by auto have "{x::'a. x\i \ a} = UNIV - {x::'a. a < x\i}" by auto also have *: "{x::'a. a < x\i} = (\k::nat. {x. (\n\Basis. (if n = i then a else -real k) *\<^sub>R n) i. -x\i)`Basis)"] - guess k::nat .. note k = this + obtain k where k: "Max ((\) (- x) ` Basis) < real k" + using reals_Archimedean2 by blast { fix i :: 'a assume "i \ Basis" then have "-x\i < real k" using k by (subst (asm) Max_less_iff) auto then have "- real k < x\i" by simp } then show "\k::nat. \ia\Basis. ia \ i \ -real k < x \ ia" by (auto intro!: exI[of _ k]) qed finally show "{x. x\i \ a} \ ?SIGMA" apply (simp only:) apply (intro sets.countable_UN sets.Diff) apply (auto intro: sigma_sets_top) done qed auto lemma borel_eq_lessThan: "borel = sigma UNIV (range (\a::'a::ordered_euclidean_space. {x. x UNIV \ Basis" then have i: "i \ Basis" by auto have "{x::'a. a \ x\i} = UNIV - {x::'a. x\i < a}" by auto also have *: "{x::'a. x\i < a} = (\k::nat. {x. x n\Basis. (if n = i then a else real k) *\<^sub>R n)})" using \i\ Basis\ proof (safe, simp_all add: eucl_less_def split: if_split_asm) fix x :: 'a - from reals_Archimedean2[of "Max ((\i. x\i)`Basis)"] - guess k::nat .. note k = this + obtain k where k: "Max ((\) x ` Basis) < real k" + using reals_Archimedean2 by blast { fix i :: 'a assume "i \ Basis" then have "x\i < real k" using k by (subst (asm) Max_less_iff) auto then have "x\i < real k" by simp } then show "\k::nat. \ia\Basis. ia \ i \ x \ ia < real k" by (auto intro!: exI[of _ k]) qed finally show "{x. a \ x\i} \ ?SIGMA" apply (simp only:) apply (intro sets.countable_UN sets.Diff) apply (auto intro: sigma_sets_top ) done qed auto lemma borel_eq_atLeastAtMost: "borel = sigma UNIV (range (\(a,b). {a..b} ::'a::ordered_euclidean_space set))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) fix a::'a have *: "{..a} = (\n::nat. {- real n *\<^sub>R One .. a})" proof (safe, simp_all add: eucl_le[where 'a='a]) fix x :: 'a - from real_arch_simple[of "Max ((\i. - x\i)`Basis)"] - guess k::nat .. note k = this + obtain k where k: "Max ((\) (- x) ` Basis) \ real k" + using real_arch_simple by blast { fix i :: 'a assume "i \ Basis" with k have "- x\i \ real k" by (subst (asm) Max_le_iff) (auto simp: field_simps) then have "- real k \ x\i" by simp } then show "\n::nat. \i\Basis. - real n \ x \ i" by (auto intro!: exI[of _ k]) qed show "{..a} \ ?SIGMA" unfolding * by (intro sets.countable_UN) (auto intro!: sigma_sets_top) qed auto lemma borel_set_induct[consumes 1, case_names empty interval compl union]: assumes "A \ sets borel" assumes empty: "P {}" and int: "\a b. a \ b \ P {a..b}" and compl: "\A. A \ sets borel \ P A \ P (-A)" and un: "\f. disjoint_family f \ (\i. f i \ sets borel) \ (\i. P (f i)) \ P (\i::nat. f i)" shows "P (A::real set)" proof - let ?G = "range (\(a,b). {a..b::real})" have "Int_stable ?G" "?G \ Pow UNIV" "A \ sigma_sets UNIV ?G" using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def) thus ?thesis proof (induction rule: sigma_sets_induct_disjoint) case (union f) from union.hyps(2) have "\i. f i \ sets borel" by (auto simp: borel_eq_atLeastAtMost) with union show ?case by (auto intro: un) next case (basic A) then obtain a b where "A = {a .. b}" by auto then show ?case by (cases "a \ b") (auto intro: int empty) qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost) qed lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\(a, b). {a <.. b::real}))" proof (rule borel_eq_sigmaI5[OF borel_eq_atMost]) fix i :: real have "{..i} = (\j::nat. {-j <.. i})" by (auto simp: minus_less_iff reals_Archimedean2) also have "\ \ sets (sigma UNIV (range (\(i, j). {i<..j})))" by (intro sets.countable_nat_UN) auto finally show "{..i} \ sets (sigma UNIV (range (\(i, j). {i<..j})))" . qed simp lemma eucl_lessThan: "{x::real. x (a, b). {a ..< b :: real}))" (is "_ = ?SIGMA") proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan]) have move_uminus: "\x y::real. -x \ y \ -y \ x" by auto fix x :: real have "{..i::nat. {-real i ..< x})" by (auto simp: move_uminus real_arch_simple) then show "{y. y ?SIGMA" by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan) qed auto lemma borel_measurable_halfspacesI: fixes f :: "'a \ 'c::euclidean_space" assumes F: "borel = sigma UNIV (F ` (UNIV \ Basis))" and S_eq: "\a i. S a i = f -` F (a,i) \ space M" shows "f \ borel_measurable M = (\i\Basis. \a::real. S a i \ sets M)" proof safe fix a :: real and i :: 'b assume i: "i \ Basis" and f: "f \ borel_measurable M" then show "S a i \ sets M" unfolding assms by (auto intro!: measurable_sets simp: assms(1)) next assume a: "\i\Basis. \a. S a i \ sets M" then show "f \ borel_measurable M" by (auto intro!: measurable_measure_of simp: S_eq F) qed lemma borel_measurable_iff_halfspace_le: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. f w \ i \ a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto lemma borel_measurable_iff_halfspace_less: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. f w \ i < a} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto lemma borel_measurable_iff_halfspace_ge: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M = (\i\Basis. \a. {w \ space M. a \ f w \ i} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto lemma borel_measurable_iff_halfspace_greater: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. \a. {w \ space M. a < f w \ i} \ sets M)" by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto lemma borel_measurable_iff_le: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w \ a} \ sets M)" using borel_measurable_iff_halfspace_le[where 'c=real] by simp lemma borel_measurable_iff_less: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. f w < a} \ sets M)" using borel_measurable_iff_halfspace_less[where 'c=real] by simp lemma borel_measurable_iff_ge: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a \ f w} \ sets M)" using borel_measurable_iff_halfspace_ge[where 'c=real] by simp lemma borel_measurable_iff_greater: "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a < f w} \ sets M)" using borel_measurable_iff_halfspace_greater[where 'c=real] by simp lemma borel_measurable_euclidean_space: fixes f :: "'a \ 'c::euclidean_space" shows "f \ borel_measurable M \ (\i\Basis. (\x. f x \ i) \ borel_measurable M)" proof safe assume f: "\i\Basis. (\x. f x \ i) \ borel_measurable M" then show "f \ borel_measurable M" by (subst borel_measurable_iff_halfspace_le) auto qed auto subsection "Borel measurable operators" lemma borel_measurable_norm[measurable]: "norm \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \ 'a) \ borel_measurable borel" by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"]) (auto intro!: continuous_on_sgn continuous_on_id) lemma borel_measurable_uminus[measurable (raw)]: fixes g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes g: "g \ borel_measurable M" shows "(\x. - g x) \ borel_measurable M" by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros) lemma borel_measurable_diff[measurable (raw)]: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x - g x) \ borel_measurable M" using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def) lemma borel_measurable_times[measurable (raw)]: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_algebra}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x * g x) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_prod[measurable (raw)]: fixes f :: "'c \ 'a \ 'b::{second_countable_topology, real_normed_field}" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" proof cases assume "finite S" thus ?thesis using assms by induct auto qed simp lemma borel_measurable_dist[measurable (raw)]: fixes g f :: "'a \ 'b::{second_countable_topology, metric_space}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. dist (f x) (g x)) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_scaleR[measurable (raw)]: fixes g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "(\x. f x *\<^sub>R g x) \ borel_measurable M" using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros) lemma borel_measurable_uminus_eq [simp]: fixes f :: "'a \ 'b::{second_countable_topology, real_normed_vector}" shows "(\x. - f x) \ borel_measurable M \ f \ borel_measurable M" (is "?l = ?r") proof assume ?l from borel_measurable_uminus[OF this] show ?r by simp qed auto lemma affine_borel_measurable_vector: fixes f :: "'a \ 'x::real_normed_vector" assumes "f \ borel_measurable M" shows "(\x. a + b *\<^sub>R f x) \ borel_measurable M" proof (rule borel_measurableI) fix S :: "'x set" assume "open S" show "(\x. a + b *\<^sub>R f x) -` S \ space M \ sets M" proof cases assume "b \ 0" with \open S\ have "open ((\x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S") using open_affinity [of S "inverse b" "- a /\<^sub>R b"] by (auto simp: algebra_simps) hence "?S \ sets borel" by auto moreover from \b \ 0\ have "(\x. a + b *\<^sub>R f x) -` S = f -` ?S" apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all) ultimately show ?thesis using assms unfolding in_borel_measurable_borel by auto qed simp qed lemma borel_measurable_const_scaleR[measurable (raw)]: "f \ borel_measurable M \ (\x. b *\<^sub>R f x ::'a::real_normed_vector) \ borel_measurable M" using affine_borel_measurable_vector[of f M 0 b] by simp lemma borel_measurable_const_add[measurable (raw)]: "f \ borel_measurable M \ (\x. a + f x ::'a::real_normed_vector) \ borel_measurable M" using affine_borel_measurable_vector[of f M a 1] by simp lemma borel_measurable_inverse[measurable (raw)]: fixes f :: "'a \ 'b::real_normed_div_algebra" assumes f: "f \ borel_measurable M" shows "(\x. inverse (f x)) \ borel_measurable M" apply (rule measurable_compose[OF f]) apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) apply (auto intro!: continuous_on_inverse continuous_on_id) done lemma borel_measurable_divide[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \ borel_measurable M" by (simp add: divide_inverse) lemma borel_measurable_abs[measurable (raw)]: "f \ borel_measurable M \ (\x. \f x :: real\) \ borel_measurable M" unfolding abs_real_def by simp lemma borel_measurable_nth[measurable (raw)]: "(\x::real^'n. x $ i) \ borel_measurable borel" by (simp add: cart_eq_inner_axis) lemma convex_measurable: fixes A :: "'a :: euclidean_space set" shows "X \ borel_measurable M \ X ` space M \ A \ open A \ convex_on A q \ (\x. q (X x)) \ borel_measurable M" by (rule measurable_compose[where f=X and N="restrict_space borel A"]) (auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2) lemma borel_measurable_ln[measurable (raw)]: assumes f: "f \ borel_measurable M" shows "(\x. ln (f x :: real)) \ borel_measurable M" apply (rule measurable_compose[OF f]) apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"]) apply (auto intro!: continuous_on_ln continuous_on_id) done lemma borel_measurable_log[measurable (raw)]: "f \ borel_measurable M \ g \ borel_measurable M \ (\x. log (g x) (f x)) \ borel_measurable M" unfolding log_def by auto lemma borel_measurable_exp[measurable]: "(exp::'a::{real_normed_field,banach}\'a) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_at_imp_continuous_on ballI isCont_exp) lemma measurable_real_floor[measurable]: "(floor :: real \ int) \ measurable borel (count_space UNIV)" proof - have "\a x. \x\ = a \ (real_of_int a \ x \ x < real_of_int (a + 1))" by (auto intro: floor_eq2) then show ?thesis by (auto simp: vimage_def measurable_count_space_eq2_countable) qed lemma measurable_real_ceiling[measurable]: "(ceiling :: real \ int) \ measurable borel (count_space UNIV)" unfolding ceiling_def[abs_def] by simp lemma borel_measurable_real_floor: "(\x::real. real_of_int \x\) \ borel_measurable borel" by simp lemma borel_measurable_root [measurable]: "root n \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_sqrt [measurable]: "sqrt \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_power [measurable (raw)]: fixes f :: "_ \ 'b::{power,real_normed_algebra}" assumes f: "f \ borel_measurable M" shows "(\x. (f x) ^ n) \ borel_measurable M" by (intro borel_measurable_continuous_on [OF _ f] continuous_intros) lemma borel_measurable_Re [measurable]: "Re \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_Im [measurable]: "Im \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_of_real [measurable]: "(of_real :: _ \ (_::real_normed_algebra)) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_sin [measurable]: "(sin :: _ \ (_::{real_normed_field,banach})) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_cos [measurable]: "(cos :: _ \ (_::{real_normed_field,banach})) \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma borel_measurable_arctan [measurable]: "arctan \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_intros) lemma\<^marker>\tag important\ borel_measurable_complex_iff: "f \ borel_measurable M \ (\x. Re (f x)) \ borel_measurable M \ (\x. Im (f x)) \ borel_measurable M" apply auto apply (subst fun_complex_eq) apply (intro borel_measurable_add) apply auto done lemma powr_real_measurable [measurable]: assumes "f \ measurable M borel" "g \ measurable M borel" shows "(\x. f x powr g x :: real) \ measurable M borel" using assms by (simp_all add: powr_def) lemma measurable_of_bool[measurable]: "of_bool \ count_space UNIV \\<^sub>M borel" by simp subsection "Borel space on the extended reals" lemma borel_measurable_ereal[measurable (raw)]: assumes f: "f \ borel_measurable M" shows "(\x. ereal (f x)) \ borel_measurable M" using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id) lemma borel_measurable_real_of_ereal[measurable (raw)]: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" shows "(\x. real_of_ereal (f x)) \ borel_measurable M" apply (rule measurable_compose[OF f]) apply (rule borel_measurable_continuous_countable_exceptions[of "{\, -\ }"]) apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV) done lemma borel_measurable_ereal_cases: fixes f :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes H: "(\x. H (ereal (real_of_ereal (f x)))) \ borel_measurable M" shows "(\x. H (f x)) \ borel_measurable M" proof - let ?F = "\x. if f x = \ then H \ else if f x = - \ then H (-\) else H (ereal (real_of_ereal (f x)))" { fix x have "H (f x) = ?F x" by (cases "f x") auto } with f H show ?thesis by simp qed lemma fixes f :: "'a \ ereal" assumes f[measurable]: "f \ borel_measurable M" shows borel_measurable_ereal_abs[measurable(raw)]: "(\x. \f x\) \ borel_measurable M" and borel_measurable_ereal_inverse[measurable(raw)]: "(\x. inverse (f x) :: ereal) \ borel_measurable M" and borel_measurable_uminus_ereal[measurable(raw)]: "(\x. - f x :: ereal) \ borel_measurable M" by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If) lemma borel_measurable_uminus_eq_ereal[simp]: "(\x. - f x :: ereal) \ borel_measurable M \ f \ borel_measurable M" (is "?l = ?r") proof assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp qed auto lemma set_Collect_ereal2: fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" assumes H: "{x \ space M. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))} \ sets M" "{x \ space borel. H (-\) (ereal x)} \ sets borel" "{x \ space borel. H (\) (ereal x)} \ sets borel" "{x \ space borel. H (ereal x) (-\)} \ sets borel" "{x \ space borel. H (ereal x) (\)} \ sets borel" shows "{x \ space M. H (f x) (g x)} \ sets M" proof - let ?G = "\y x. if g x = \ then H y \ else if g x = -\ then H y (-\) else H y (ereal (real_of_ereal (g x)))" let ?F = "\x. if f x = \ then ?G \ x else if f x = -\ then ?G (-\) x else ?G (ereal (real_of_ereal (f x))) x" { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } note * = this from assms show ?thesis by (subst *) (simp del: space_borel split del: if_split) qed lemma borel_measurable_ereal_iff: shows "(\x. ereal (f x)) \ borel_measurable M \ f \ borel_measurable M" proof assume "(\x. ereal (f x)) \ borel_measurable M" from borel_measurable_real_of_ereal[OF this] show "f \ borel_measurable M" by auto qed auto lemma borel_measurable_erealD[measurable_dest]: "(\x. ereal (f x)) \ borel_measurable M \ g \ measurable N M \ (\x. f (g x)) \ borel_measurable N" unfolding borel_measurable_ereal_iff by simp theorem borel_measurable_ereal_iff_real: fixes f :: "'a \ ereal" shows "f \ borel_measurable M \ ((\x. real_of_ereal (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M \ f -` {-\} \ space M \ sets M)" proof safe assume *: "(\x. real_of_ereal (f x)) \ borel_measurable M" "f -` {\} \ space M \ sets M" "f -` {-\} \ space M \ sets M" have "f -` {\} \ space M = {x\space M. f x = \}" "f -` {-\} \ space M = {x\space M. f x = -\}" by auto with * have **: "{x\space M. f x = \} \ sets M" "{x\space M. f x = -\} \ sets M" by simp_all let ?f = "\x. if f x = \ then \ else if f x = -\ then -\ else ereal (real_of_ereal (f x))" have "?f \ borel_measurable M" using * ** by (intro measurable_If) auto also have "?f = f" by (auto simp: fun_eq_iff ereal_real) finally show "f \ borel_measurable M" . qed simp_all lemma borel_measurable_ereal_iff_Iio: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {..< a} \ space M \ sets M)" by (auto simp: borel_Iio measurable_iff_measure_of) lemma borel_measurable_ereal_iff_Ioi: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a <..} \ space M \ sets M)" by (auto simp: borel_Ioi measurable_iff_measure_of) lemma vimage_sets_compl_iff: "f -` A \ space M \ sets M \ f -` (- A) \ space M \ sets M" proof - { fix A assume "f -` A \ space M \ sets M" moreover have "f -` (- A) \ space M = space M - f -` A \ space M" by auto ultimately have "f -` (- A) \ space M \ sets M" by auto } from this[of A] this[of "-A"] show ?thesis by (metis double_complement) qed lemma borel_measurable_iff_Iic_ereal: "(f::'a\ereal) \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp lemma borel_measurable_iff_Ici_ereal: "(f::'a \ ereal) \ borel_measurable M \ (\a. f -` {a..} \ space M \ sets M)" unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp lemma borel_measurable_ereal2: fixes f g :: "'a \ ereal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" assumes H: "(\x. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))) \ borel_measurable M" "(\x. H (-\) (ereal (real_of_ereal (g x)))) \ borel_measurable M" "(\x. H (\) (ereal (real_of_ereal (g x)))) \ borel_measurable M" "(\x. H (ereal (real_of_ereal (f x))) (-\)) \ borel_measurable M" "(\x. H (ereal (real_of_ereal (f x))) (\)) \ borel_measurable M" shows "(\x. H (f x) (g x)) \ borel_measurable M" proof - let ?G = "\y x. if g x = \ then H y \ else if g x = - \ then H y (-\) else H y (ereal (real_of_ereal (g x)))" let ?F = "\x. if f x = \ then ?G \ x else if f x = - \ then ?G (-\) x else ?G (ereal (real_of_ereal (f x))) x" { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto } note * = this from assms show ?thesis unfolding * by simp qed lemma [measurable(raw)]: fixes f :: "'a \ ereal" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows borel_measurable_ereal_add: "(\x. f x + g x) \ borel_measurable M" and borel_measurable_ereal_times: "(\x. f x * g x) \ borel_measurable M" by (simp_all add: borel_measurable_ereal2) lemma [measurable(raw)]: fixes f g :: "'a \ ereal" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows borel_measurable_ereal_diff: "(\x. f x - g x) \ borel_measurable M" and borel_measurable_ereal_divide: "(\x. f x / g x) \ borel_measurable M" using assms by (simp_all add: minus_ereal_def divide_ereal_def) lemma borel_measurable_ereal_sum[measurable (raw)]: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" using assms by (induction S rule: infinite_finite_induct) auto lemma borel_measurable_ereal_prod[measurable (raw)]: fixes f :: "'c \ 'a \ ereal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" using assms by (induction S rule: infinite_finite_induct) auto lemma borel_measurable_extreal_suminf[measurable (raw)]: fixes f :: "nat \ 'a \ ereal" assumes [measurable]: "\i. f i \ borel_measurable M" shows "(\x. (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp subsection "Borel space on the extended non-negative reals" text \ \<^type>\ennreal\ is a topological monoid, so no rules for plus are required, also all order statements are usually done on type classes. \ lemma measurable_enn2ereal[measurable]: "enn2ereal \ borel \\<^sub>M borel" by (intro borel_measurable_continuous_onI continuous_on_enn2ereal) lemma measurable_e2ennreal[measurable]: "e2ennreal \ borel \\<^sub>M borel" by (intro borel_measurable_continuous_onI continuous_on_e2ennreal) lemma borel_measurable_enn2real[measurable (raw)]: "f \ M \\<^sub>M borel \ (\x. enn2real (f x)) \ M \\<^sub>M borel" unfolding enn2real_def[abs_def] by measurable definition\<^marker>\tag important\ [simp]: "is_borel f M \ f \ borel_measurable M" lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (=) is_borel is_borel" unfolding is_borel_def[abs_def] proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1]) fix f and M :: "'a measure" show "f \ borel_measurable M" if f: "enn2ereal \ f \ borel_measurable M" using measurable_compose[OF f measurable_e2ennreal] by simp qed simp context includes ennreal.lifting begin lemma measurable_ennreal[measurable]: "ennreal \ borel \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_ennreal_iff[simp]: assumes [simp]: "\x. x \ space M \ 0 \ f x" shows "(\x. ennreal (f x)) \ M \\<^sub>M borel \ f \ M \\<^sub>M borel" proof safe assume "(\x. ennreal (f x)) \ M \\<^sub>M borel" then have "(\x. enn2real (ennreal (f x))) \ M \\<^sub>M borel" by measurable then show "f \ M \\<^sub>M borel" by (rule measurable_cong[THEN iffD1, rotated]) auto qed measurable lemma borel_measurable_times_ennreal[measurable (raw)]: fixes f g :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ g \ M \\<^sub>M borel \ (\x. f x * g x) \ M \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_inverse_ennreal[measurable (raw)]: fixes f :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ (\x. inverse (f x)) \ M \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_divide_ennreal[measurable (raw)]: fixes f :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ g \ M \\<^sub>M borel \ (\x. f x / g x) \ M \\<^sub>M borel" unfolding divide_ennreal_def by simp lemma borel_measurable_minus_ennreal[measurable (raw)]: fixes f :: "'a \ ennreal" shows "f \ M \\<^sub>M borel \ g \ M \\<^sub>M borel \ (\x. f x - g x) \ M \\<^sub>M borel" unfolding is_borel_def[symmetric] by transfer simp lemma borel_measurable_power_ennreal [measurable (raw)]: fixes f :: "_ \ ennreal" assumes f: "f \ borel_measurable M" shows "(\x. (f x) ^ n) \ borel_measurable M" by (induction n) (use f in auto) lemma borel_measurable_prod_ennreal[measurable (raw)]: fixes f :: "'c \ 'a \ ennreal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" using assms by (induction S rule: infinite_finite_induct) auto end hide_const (open) is_borel subsection \LIMSEQ is borel measurable\ lemma borel_measurable_LIMSEQ_real: fixes u :: "nat \ 'a \ real" assumes u': "\x. x \ space M \ (\i. u i x) \ u' x" and u: "\i. u i \ borel_measurable M" shows "u' \ borel_measurable M" proof - have "\x. x \ space M \ liminf (\n. ereal (u n x)) = ereal (u' x)" using u' by (simp add: lim_imp_Liminf) moreover from u have "(\x. liminf (\n. ereal (u n x))) \ borel_measurable M" by auto ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) qed lemma borel_measurable_LIMSEQ_metric: fixes f :: "nat \ 'a \ 'b :: metric_space" assumes [measurable]: "\i. f i \ borel_measurable M" assumes lim: "\x. x \ space M \ (\i. f i x) \ g x" shows "g \ borel_measurable M" unfolding borel_eq_closed proof (safe intro!: measurable_measure_of) fix A :: "'b set" assume "closed A" have [measurable]: "(\x. infdist (g x) A) \ borel_measurable M" proof (rule borel_measurable_LIMSEQ_real) show "\x. x \ space M \ (\i. infdist (f i x) A) \ infdist (g x) A" by (intro tendsto_infdist lim) show "\i. (\x. infdist (f i x) A) \ borel_measurable M" by (intro borel_measurable_continuous_on[where f="\x. infdist x A"] continuous_at_imp_continuous_on ballI continuous_infdist continuous_ident) auto qed show "g -` A \ space M \ sets M" proof cases assume "A \ {}" then have "\x. infdist x A = 0 \ x \ A" using \closed A\ by (simp add: in_closed_iff_infdist_zero) then have "g -` A \ space M = {x\space M. infdist (g x) A = 0}" by auto also have "\ \ sets M" by measurable finally show ?thesis . qed simp qed auto lemma sets_Collect_Cauchy[measurable]: fixes f :: "nat \ 'a => 'b::{metric_space, second_countable_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "{x\space M. Cauchy (\i. f i x)} \ sets M" unfolding metric_Cauchy_iff2 using f by auto lemma borel_measurable_lim_metric[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "(\x. lim (\i. f i x)) \ borel_measurable M" proof - define u' where "u' x = lim (\i. if Cauchy (\i. f i x) then f i x else 0)" for x then have *: "\x. lim (\i. f i x) = (if Cauchy (\i. f i x) then u' x else (THE x. False))" by (auto simp: lim_def convergent_eq_Cauchy[symmetric]) have "u' \ borel_measurable M" proof (rule borel_measurable_LIMSEQ_metric) fix x have "convergent (\i. if Cauchy (\i. f i x) then f i x else 0)" by (cases "Cauchy (\i. f i x)") (auto simp add: convergent_eq_Cauchy[symmetric] convergent_def) then show "(\i. if Cauchy (\i. f i x) then f i x else 0) \ u' x" unfolding u'_def by (rule convergent_LIMSEQ_iff[THEN iffD1]) qed measurable then show ?thesis unfolding * by measurable qed lemma borel_measurable_suminf[measurable (raw)]: fixes f :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "\i. f i \ borel_measurable M" shows "(\x. suminf (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp lemma Collect_closed_imp_pred_borel: "closed {x. P x} \ Measurable.pred borel P" by (simp add: pred_def) (* Proof by Jeremy Avigad and Luke Serafin *) lemma isCont_borel_pred[measurable]: fixes f :: "'b::metric_space \ 'a::metric_space" shows "Measurable.pred borel (isCont f)" proof (subst measurable_cong) let ?I = "\j. inverse(real (Suc j))" show "isCont f x = (\i. \j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i)" for x unfolding continuous_at_eps_delta proof safe fix i assume "\e>0. \d>0. \y. dist y x < d \ dist (f y) (f x) < e" moreover have "0 < ?I i / 2" by simp ultimately obtain d where d: "0 < d" "\y. dist x y < d \ dist (f y) (f x) < ?I i / 2" by (metis dist_commute) then obtain j where j: "?I j < d" by (metis reals_Archimedean) show "\j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i" proof (safe intro!: exI[where x=j]) fix y z assume *: "dist x y < ?I j" "dist x z < ?I j" have "dist (f y) (f z) \ dist (f y) (f x) + dist (f z) (f x)" by (rule dist_triangle2) also have "\ < ?I i / 2 + ?I i / 2" by (intro add_strict_mono d less_trans[OF _ j] *) also have "\ \ ?I i" by (simp add: field_simps) finally show "dist (f y) (f z) \ ?I i" by simp qed next fix e::real assume "0 < e" then obtain n where n: "?I n < e" by (metis reals_Archimedean) assume "\i. \j. \y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I i" from this[THEN spec, of "Suc n"] obtain j where j: "\y z. dist x y < ?I j \ dist x z < ?I j \ dist (f y) (f z) \ ?I (Suc n)" by auto show "\d>0. \y. dist y x < d \ dist (f y) (f x) < e" proof (safe intro!: exI[of _ "?I j"]) fix y assume "dist y x < ?I j" then have "dist (f y) (f x) \ ?I (Suc n)" by (intro j) (auto simp: dist_commute) also have "?I (Suc n) < ?I n" by simp also note n finally show "dist (f y) (f x) < e" . qed simp qed qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros) lemma isCont_borel: fixes f :: "'b::metric_space \ 'a::metric_space" shows "{x. isCont f x} \ sets borel" by simp lemma is_real_interval: assumes S: "is_interval S" shows "\a b::real. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" using S unfolding is_interval_1 by (blast intro: interval_cases) lemma real_interval_borel_measurable: assumes "is_interval (S::real set)" shows "S \ sets borel" proof - from assms is_real_interval have "\a b::real. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" by auto - then guess a .. - then guess b .. - thus ?thesis + then show ?thesis by auto qed text \The next lemmas hold in any second countable linorder (including ennreal or ereal for instance), but in the current state they are restricted to reals.\ lemma borel_measurable_mono_on_fnc: fixes f :: "real \ real" and A :: "real set" assumes "mono_on f A" shows "f \ borel_measurable (restrict_space borel A)" apply (rule measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]]) apply (auto intro!: image_eqI[where x="{x}" for x] simp: sets_restrict_space) apply (auto simp add: sets_restrict_restrict_space continuous_on_eq_continuous_within cong: measurable_cong_sets intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset) done lemma borel_measurable_piecewise_mono: fixes f::"real \ real" and C::"real set set" assumes "countable C" "\c. c \ C \ c \ sets borel" "\c. c \ C \ mono_on f c" "(\C) = UNIV" shows "f \ borel_measurable borel" by (rule measurable_piecewise_restrict[of C], auto intro: borel_measurable_mono_on_fnc simp: assms) lemma borel_measurable_mono: fixes f :: "real \ real" shows "mono f \ f \ borel_measurable borel" using borel_measurable_mono_on_fnc[of f UNIV] by (simp add: mono_def mono_on_def) lemma measurable_bdd_below_real[measurable (raw)]: fixes F :: "'a \ 'i \ real" assumes [simp]: "countable I" and [measurable]: "\i. i \ I \ F i \ M \\<^sub>M borel" shows "Measurable.pred M (\x. bdd_below ((\i. F i x)`I))" proof (subst measurable_cong) show "bdd_below ((\i. F i x)`I) \ (\q\\. \i\I. q \ F i x)" for x by (auto simp: bdd_below_def intro!: bexI[of _ "of_int (floor _)"] intro: order_trans of_int_floor_le) show "Measurable.pred M (\w. \q\\. \i\I. q \ F i w)" using countable_int by measurable qed lemma borel_measurable_cINF_real[measurable (raw)]: fixes F :: "_ \ _ \ real" assumes [simp]: "countable I" assumes F[measurable]: "\i. i \ I \ F i \ borel_measurable M" shows "(\x. INF i\I. F i x) \ borel_measurable M" proof (rule measurable_piecewise_restrict) let ?\ = "{x\space M. bdd_below ((\i. F i x)`I)}" show "countable {?\, - ?\}" "space M \ \{?\, - ?\}" "\X. X \ {?\, - ?\} \ X \ space M \ sets M" by auto fix X assume "X \ {?\, - ?\}" then show "(\x. INF i\I. F i x) \ borel_measurable (restrict_space M X)" proof safe show "(\x. INF i\I. F i x) \ borel_measurable (restrict_space M ?\)" by (intro borel_measurable_cINF measurable_restrict_space1 F) (auto simp: space_restrict_space) show "(\x. INF i\I. F i x) \ borel_measurable (restrict_space M (-?\))" proof (subst measurable_cong) fix x assume "x \ space (restrict_space M (-?\))" then have "\ (\i\I. - F i x \ y)" for y by (auto simp: space_restrict_space bdd_above_def bdd_above_uminus[symmetric]) then show "(INF i\I. F i x) = - (THE x. False)" by (auto simp: space_restrict_space Inf_real_def Sup_real_def Least_def simp del: Set.ball_simps(10)) qed simp qed qed lemma borel_Ici: "borel = sigma UNIV (range (\x::real. {x ..}))" proof (safe intro!: borel_eq_sigmaI1[OF borel_Iio]) fix x :: real have eq: "{.. sets (sigma UNIV (range atLeast))" unfolding eq by (intro sets.compl_sets) auto qed auto lemma borel_measurable_pred_less[measurable (raw)]: fixes f :: "'a \ 'b::{second_countable_topology, linorder_topology}" shows "f \ borel_measurable M \ g \ borel_measurable M \ Measurable.pred M (\w. f w < g w)" unfolding Measurable.pred_def by (rule borel_measurable_less) no_notation eucl_less (infix " _ \ 'a::{second_countable_topology, dense_linorder, linorder_topology}" assumes "finite I" and [measurable]: "\i. f i \ borel_measurable M" shows "(\x. Max{f i x |i. i \ I}) \ borel_measurable M" by (simp add: borel_measurable_Max[OF assms(1), where ?f=f and ?M=M] Setcompr_eq_image) lemma measurable_compose_n [measurable (raw)]: assumes "T \ measurable M M" shows "(T^^n) \ measurable M M" by (induction n, auto simp add: measurable_compose[OF _ assms]) lemma measurable_real_imp_nat: fixes f::"'a \ nat" assumes [measurable]: "(\x. real(f x)) \ borel_measurable M" shows "f \ measurable M (count_space UNIV)" proof - let ?g = "(\x. real(f x))" have "\(n::nat). ?g-`({real n}) \ space M = f-`{n} \ space M" by auto moreover have "\(n::nat). ?g-`({real n}) \ space M \ sets M" using assms by measurable ultimately have "\(n::nat). f-`{n} \ space M \ sets M" by simp then show ?thesis using measurable_count_space_eq2_countable by blast qed lemma measurable_equality_set [measurable]: fixes f g::"_\ 'a::{second_countable_topology, t2_space}" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows "{x \ space M. f x = g x} \ sets M" proof - define A where "A = {x \ space M. f x = g x}" define B where "B = {y. \x::'a. y = (x,x)}" have "A = (\x. (f x, g x))-`B \ space M" unfolding A_def B_def by auto moreover have "(\x. (f x, g x)) \ borel_measurable M" by simp moreover have "B \ sets borel" unfolding B_def by (simp add: closed_diagonal) ultimately have "A \ sets M" by simp then show ?thesis unfolding A_def by simp qed lemma measurable_inequality_set [measurable]: fixes f g::"_ \ 'a::{second_countable_topology, linorder_topology}" assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows "{x \ space M. f x \ g x} \ sets M" "{x \ space M. f x < g x} \ sets M" "{x \ space M. f x \ g x} \ sets M" "{x \ space M. f x > g x} \ sets M" proof - define F where "F = (\x. (f x, g x))" have * [measurable]: "F \ borel_measurable M" unfolding F_def by simp have "{x \ space M. f x \ g x} = F-`{(x, y) | x y. x \ y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x \ (y::'a)} \ sets borel" using closed_subdiagonal borel_closed by blast ultimately show "{x \ space M. f x \ g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) have "{x \ space M. f x < g x} = F-`{(x, y) | x y. x < y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x < (y::'a)} \ sets borel" using open_subdiagonal borel_open by blast ultimately show "{x \ space M. f x < g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) have "{x \ space M. f x \ g x} = F-`{(x, y) | x y. x \ y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x \ (y::'a)} \ sets borel" using closed_superdiagonal borel_closed by blast ultimately show "{x \ space M. f x \ g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) have "{x \ space M. f x > g x} = F-`{(x, y) | x y. x > y} \ space M" unfolding F_def by auto moreover have "{(x, y) | x y. x > (y::'a)} \ sets borel" using open_superdiagonal borel_open by blast ultimately show "{x \ space M. f x > g x} \ sets M" using * by (metis (mono_tags, lifting) measurable_sets) qed proposition measurable_limit [measurable]: fixes f::"nat \ 'a \ 'b::first_countable_topology" assumes [measurable]: "\n::nat. f n \ borel_measurable M" shows "Measurable.pred M (\x. (\n. f n x) \ c)" proof - obtain A :: "nat \ 'b set" where A: "\i. open (A i)" "\i. c \ A i" "\S. open S \ c \ S \ eventually (\i. A i \ S) sequentially" by (rule countable_basis_at_decseq) blast have [measurable]: "\N i. (f N)-`(A i) \ space M \ sets M" using A(1) by auto then have mes: "(\i. \n. \N\{n..}. (f N)-`(A i) \ space M) \ sets M" by blast have "(u \ c) \ (\i. eventually (\n. u n \ A i) sequentially)" for u::"nat \ 'b" proof assume "u \ c" then have "eventually (\n. u n \ A i) sequentially" for i using A(1)[of i] A(2)[of i] by (simp add: topological_tendstoD) then show "(\i. eventually (\n. u n \ A i) sequentially)" by auto next assume H: "(\i. eventually (\n. u n \ A i) sequentially)" show "(u \ c)" proof (rule topological_tendstoI) fix S assume "open S" "c \ S" with A(3)[OF this] obtain i where "A i \ S" using eventually_False_sequentially eventually_mono by blast moreover have "eventually (\n. u n \ A i) sequentially" using H by simp ultimately show "\\<^sub>F n in sequentially. u n \ S" by (simp add: eventually_mono subset_eq) qed qed then have "{x. (\n. f n x) \ c} = (\i. \n. \N\{n..}. (f N)-`(A i))" by (auto simp add: atLeast_def eventually_at_top_linorder) then have "{x \ space M. (\n. f n x) \ c} = (\i. \n. \N\{n..}. (f N)-`(A i) \ space M)" by auto then have "{x \ space M. (\n. f n x) \ c} \ sets M" using mes by simp then show ?thesis by auto qed lemma measurable_limit2 [measurable]: fixes u::"nat \ 'a \ real" assumes [measurable]: "\n. u n \ borel_measurable M" "v \ borel_measurable M" shows "Measurable.pred M (\x. (\n. u n x) \ v x)" proof - define w where "w = (\n x. u n x - v x)" have [measurable]: "w n \ borel_measurable M" for n unfolding w_def by auto have "((\n. u n x) \ v x) \ ((\n. w n x) \ 0)" for x unfolding w_def using Lim_null by auto then show ?thesis using measurable_limit by auto qed lemma measurable_P_restriction [measurable (raw)]: assumes [measurable]: "Measurable.pred M P" "A \ sets M" shows "{x \ A. P x} \ sets M" proof - have "A \ space M" using sets.sets_into_space[OF assms(2)]. then have "{x \ A. P x} = A \ {x \ space M. P x}" by blast then show ?thesis by auto qed lemma measurable_sum_nat [measurable (raw)]: fixes f :: "'c \ 'a \ nat" assumes "\i. i \ S \ f i \ measurable M (count_space UNIV)" shows "(\x. \i\S. f i x) \ measurable M (count_space UNIV)" proof cases assume "finite S" then show ?thesis using assms by induct auto qed simp lemma measurable_abs_powr [measurable]: fixes p::real assumes [measurable]: "f \ borel_measurable M" shows "(\x. \f x\ powr p) \ borel_measurable M" by simp text \The next one is a variation around \measurable_restrict_space\.\ lemma measurable_restrict_space3: assumes "f \ measurable M N" and "f \ A \ B" shows "f \ measurable (restrict_space M A) (restrict_space N B)" proof - have "f \ measurable (restrict_space M A) N" using assms(1) measurable_restrict_space1 by auto then show ?thesis by (metis Int_iff funcsetI funcset_mem measurable_restrict_space2[of f, of "restrict_space M A", of B, of N] assms(2) space_restrict_space) qed lemma measurable_restrict_mono: assumes f: "f \ restrict_space M A \\<^sub>M N" and "B \ A" shows "f \ restrict_space M B \\<^sub>M N" by (rule measurable_compose[OF measurable_restrict_space3 f]) (insert \B \ A\, auto) text \The next one is a variation around \measurable_piecewise_restrict\.\ lemma measurable_piecewise_restrict2: assumes [measurable]: "\n. A n \ sets M" and "space M = (\(n::nat). A n)" "\n. \h \ measurable M N. (\x \ A n. f x = h x)" shows "f \ measurable M N" proof (rule measurableI) fix B assume [measurable]: "B \ sets N" { fix n::nat obtain h where [measurable]: "h \ measurable M N" and "\x \ A n. f x = h x" using assms(3) by blast then have *: "f-`B \ A n = h-`B \ A n" by auto have "h-`B \ A n = h-`B \ space M \ A n" using assms(2) sets.sets_into_space by auto then have "h-`B \ A n \ sets M" by simp then have "f-`B \ A n \ sets M" using * by simp } then have "(\n. f-`B \ A n) \ sets M" by measurable moreover have "f-`B \ space M = (\n. f-`B \ A n)" using assms(2) by blast ultimately show "f-`B \ space M \ sets M" by simp next fix x assume "x \ space M" then obtain n where "x \ A n" using assms(2) by blast obtain h where [measurable]: "h \ measurable M N" and "\x \ A n. f x = h x" using assms(3) by blast then have "f x = h x" using \x \ A n\ by blast moreover have "h x \ space N" by (metis measurable_space \x \ space M\ \h \ measurable M N\) ultimately show "f x \ space N" by simp qed end diff --git a/src/HOL/Analysis/Caratheodory.thy b/src/HOL/Analysis/Caratheodory.thy --- a/src/HOL/Analysis/Caratheodory.thy +++ b/src/HOL/Analysis/Caratheodory.thy @@ -1,891 +1,900 @@ (* Title: HOL/Analysis/Caratheodory.thy Author: Lawrence C Paulson Author: Johannes Hölzl, TU München *) section \Caratheodory Extension Theorem\ theory Caratheodory imports Measure_Space begin text \ Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson. \ lemma suminf_ennreal_2dimen: fixes f:: "nat \ nat \ ennreal" assumes "\m. g m = (\n. f (m,n))" shows "(\i. f (prod_decode i)) = suminf g" proof - have g_def: "g = (\m. (\n. f (m,n)))" using assms by (simp add: fun_eq_iff) have reindex: "\B. (\x\B. f (prod_decode x)) = sum f (prod_decode ` B)" by (simp add: sum.reindex[OF inj_prod_decode] comp_def) have "(SUP n. \i UNIV \ UNIV. \in sum f ({.. {..a b. sum f (prod_decode ` {.. sum f ({.. {.. ?M" by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) } then have "sum f ({.. {.. sum f ?M" by (auto intro!: sum_mono2) then show "\n. sum f ({.. {.. sum f (prod_decode ` {.. = (SUP p. \in. f (i, n))" unfolding suminf_sum[OF summableI, symmetric] by (simp add: suminf_eq_SUP SUP_pair sum.swap[of _ "{..< fst _}"]) finally show ?thesis unfolding g_def by (simp add: suminf_eq_SUP) qed subsection \Characterizations of Measures\ definition\<^marker>\tag important\ outer_measure_space where "outer_measure_space M f \ positive M f \ increasing M f \ countably_subadditive M f" subsubsection\<^marker>\tag important\ \Lambda Systems\ definition\<^marker>\tag important\ lambda_system :: "'a set \ 'a set set \ ('a set \ ennreal) \ 'a set set" where "lambda_system \ M f = {l \ M. \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" lemma (in algebra) lambda_system_eq: "lambda_system \ M f = {l \ M. \x \ M. f (x \ l) + f (x - l) = f x}" proof - have [simp]: "\l x. l \ M \ x \ M \ (\ - l) \ x = x - l" by (metis Int_Diff Int_absorb1 Int_commute sets_into_space) show ?thesis by (auto simp add: lambda_system_def) (metis Int_commute)+ qed lemma (in algebra) lambda_system_empty: "positive M f \ {} \ lambda_system \ M f" by (auto simp add: positive_def lambda_system_eq) lemma lambda_system_sets: "x \ lambda_system \ M f \ x \ M" by (simp add: lambda_system_def) lemma (in algebra) lambda_system_Compl: fixes f:: "'a set \ ennreal" assumes x: "x \ lambda_system \ M f" shows "\ - x \ lambda_system \ M f" proof - have "x \ \" by (metis sets_into_space lambda_system_sets x) hence "\ - (\ - x) = x" by (metis double_diff equalityE) with x show ?thesis by (force simp add: lambda_system_def ac_simps) qed lemma (in algebra) lambda_system_Int: fixes f:: "'a set \ ennreal" assumes xl: "x \ lambda_system \ M f" and yl: "y \ lambda_system \ M f" shows "x \ y \ lambda_system \ M f" proof - from xl yl show ?thesis proof (auto simp add: positive_def lambda_system_eq Int) fix u assume x: "x \ M" and y: "y \ M" and u: "u \ M" and fx: "\z\M. f (z \ x) + f (z - x) = f z" and fy: "\z\M. f (z \ y) + f (z - y) = f z" have "u - x \ y \ M" by (metis Diff Diff_Int Un u x y) moreover have "(u - (x \ y)) \ y = u \ y - x" by blast moreover have "u - x \ y - y = u - y" by blast ultimately have ey: "f (u - x \ y) = f (u \ y - x) + f (u - y)" using fy by force have "f (u \ (x \ y)) + f (u - x \ y) = (f (u \ (x \ y)) + f (u \ y - x)) + f (u - y)" by (simp add: ey ac_simps) also have "... = (f ((u \ y) \ x) + f (u \ y - x)) + f (u - y)" by (simp add: Int_ac) also have "... = f (u \ y) + f (u - y)" using fx [THEN bspec, of "u \ y"] Int y u by force also have "... = f u" by (metis fy u) finally show "f (u \ (x \ y)) + f (u - x \ y) = f u" . qed qed lemma (in algebra) lambda_system_Un: fixes f:: "'a set \ ennreal" assumes xl: "x \ lambda_system \ M f" and yl: "y \ lambda_system \ M f" shows "x \ y \ lambda_system \ M f" proof - have "(\ - x) \ (\ - y) \ M" by (metis Diff_Un Un compl_sets lambda_system_sets xl yl) moreover have "x \ y = \ - ((\ - x) \ (\ - y))" by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+ ultimately show ?thesis by (metis lambda_system_Compl lambda_system_Int xl yl) qed lemma (in algebra) lambda_system_algebra: "positive M f \ algebra \ (lambda_system \ M f)" apply (auto simp add: algebra_iff_Un) apply (metis lambda_system_sets subsetD sets_into_space) apply (metis lambda_system_empty) apply (metis lambda_system_Compl) apply (metis lambda_system_Un) done lemma (in algebra) lambda_system_strong_additive: assumes z: "z \ M" and disj: "x \ y = {}" and xl: "x \ lambda_system \ M f" and yl: "y \ lambda_system \ M f" shows "f (z \ (x \ y)) = f (z \ x) + f (z \ y)" proof - have "z \ x = (z \ (x \ y)) \ x" using disj by blast moreover have "z \ y = (z \ (x \ y)) - x" using disj by blast moreover have "(z \ (x \ y)) \ M" by (metis Int Un lambda_system_sets xl yl z) ultimately show ?thesis using xl yl by (simp add: lambda_system_eq) qed lemma (in algebra) lambda_system_additive: "additive (lambda_system \ M f) f" proof (auto simp add: additive_def) fix x and y assume disj: "x \ y = {}" and xl: "x \ lambda_system \ M f" and yl: "y \ lambda_system \ M f" hence "x \ M" "y \ M" by (blast intro: lambda_system_sets)+ thus "f (x \ y) = f x + f y" using lambda_system_strong_additive [OF top disj xl yl] by (simp add: Un) qed lemma lambda_system_increasing: "increasing M f \ increasing (lambda_system \ M f) f" by (simp add: increasing_def lambda_system_def) lemma lambda_system_positive: "positive M f \ positive (lambda_system \ M f) f" by (simp add: positive_def lambda_system_def) lemma (in algebra) lambda_system_strong_sum: fixes A:: "nat \ 'a set" and f :: "'a set \ ennreal" assumes f: "positive M f" and a: "a \ M" and A: "range A \ lambda_system \ M f" and disj: "disjoint_family A" shows "(\i = 0..A i)) = f (a \ (\i\{0.. \ (A ` {0.. lambda_system \ M f" using A by blast interpret l: algebra \ "lambda_system \ M f" using f by (rule lambda_system_algebra) have 4: "\ (A ` {0.. lambda_system \ M f" using A l.UNION_in_sets by simp from Suc.hyps show ?case by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4]) qed proposition (in sigma_algebra) lambda_system_caratheodory: assumes oms: "outer_measure_space M f" and A: "range A \ lambda_system \ M f" and disj: "disjoint_family A" shows "(\i. A i) \ lambda_system \ M f \ (\i. f (A i)) = f (\i. A i)" proof - have pos: "positive M f" and inc: "increasing M f" and csa: "countably_subadditive M f" by (metis oms outer_measure_space_def)+ have sa: "subadditive M f" by (metis countably_subadditive_subadditive csa pos) have A': "\S. A`S \ (lambda_system \ M f)" using A by auto interpret ls: algebra \ "lambda_system \ M f" using pos by (rule lambda_system_algebra) have A'': "range A \ M" by (metis A image_subset_iff lambda_system_sets) have U_in: "(\i. A i) \ M" by (metis A'' countable_UN) have U_eq: "f (\i. A i) = (\i. f (A i))" proof (rule antisym) show "f (\i. A i) \ (\i. f (A i))" using csa[unfolded countably_subadditive_def] A'' disj U_in by auto have dis: "\N. disjoint_family_on A {..i. f (A i)) \ f (\i. A i)" using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] A'' by (intro suminf_le_const[OF summableI]) (auto intro!: increasingD[OF inc] countable_UN) qed have "f (a \ (\i. A i)) + f (a - (\i. A i)) = f a" if a [iff]: "a \ M" for a proof (rule antisym) have "range (\i. a \ A i) \ M" using A'' by blast moreover have "disjoint_family (\i. a \ A i)" using disj by (auto simp add: disjoint_family_on_def) moreover have "a \ (\i. A i) \ M" by (metis Int U_in a) ultimately have "f (a \ (\i. A i)) \ (\i. f (a \ A i))" using csa[unfolded countably_subadditive_def, rule_format, of "(\i. a \ A i)"] by (simp add: o_def) hence "f (a \ (\i. A i)) + f (a - (\i. A i)) \ (\i. f (a \ A i)) + f (a - (\i. A i))" by (rule add_right_mono) also have "\ \ f a" proof (intro ennreal_suminf_bound_add) fix n have UNION_in: "(\i\{0.. M" by (metis A'' UNION_in_sets) have le_fa: "f (\ (A ` {0.. a) \ f a" using A'' by (blast intro: increasingD [OF inc] A'' UNION_in_sets) have ls: "(\i\{0.. lambda_system \ M f" using ls.UNION_in_sets by (simp add: A) hence eq_fa: "f a = f (a \ (\i\{0..i\{0..i. A i)) \ f (a - (\i\{0..i A i)) + f (a - (\i. A i)) \ f a" by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric]) qed finally show "f (a \ (\i. A i)) + f (a - (\i. A i)) \ f a" by simp next have "f a \ f (a \ (\i. A i) \ (a - (\i. A i)))" by (blast intro: increasingD [OF inc] U_in) also have "... \ f (a \ (\i. A i)) + f (a - (\i. A i))" by (blast intro: subadditiveD [OF sa] U_in) finally show "f a \ f (a \ (\i. A i)) + f (a - (\i. A i))" . qed thus ?thesis by (simp add: lambda_system_eq sums_iff U_eq U_in) qed proposition (in sigma_algebra) caratheodory_lemma: assumes oms: "outer_measure_space M f" defines "L \ lambda_system \ M f" shows "measure_space \ L f" proof - have pos: "positive M f" by (metis oms outer_measure_space_def) have alg: "algebra \ L" using lambda_system_algebra [of f, OF pos] by (simp add: algebra_iff_Un L_def) then have "sigma_algebra \ L" using lambda_system_caratheodory [OF oms] by (simp add: sigma_algebra_disjoint_iff L_def) moreover have "countably_additive L f" "positive L f" using pos lambda_system_caratheodory [OF oms] by (auto simp add: lambda_system_sets L_def countably_additive_def positive_def) ultimately show ?thesis using pos by (simp add: measure_space_def) qed definition\<^marker>\tag important\ outer_measure :: "'a set set \ ('a set \ ennreal) \ 'a set \ ennreal" where "outer_measure M f X = (INF A\{A. range A \ M \ disjoint_family A \ X \ (\i. A i)}. \i. f (A i))" lemma (in ring_of_sets) outer_measure_agrees: assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \ M" shows "outer_measure M f s = f s" unfolding outer_measure_def proof (safe intro!: antisym INF_greatest) fix A :: "nat \ 'a set" assume A: "range A \ M" and dA: "disjoint_family A" and sA: "s \ (\x. A x)" have inc: "increasing M f" by (metis additive_increasing ca countably_additive_additive posf) have "f s = f (\i. A i \ s)" using sA by (auto simp: Int_absorb1) also have "\ = (\i. f (A i \ s))" using sA dA A s by (intro ca[unfolded countably_additive_def, rule_format, symmetric]) (auto simp: Int_absorb1 disjoint_family_on_def) also have "... \ (\i. f (A i))" using A s by (auto intro!: suminf_le increasingD[OF inc]) finally show "f s \ (\i. f (A i))" . next have "(\i. f (if i = 0 then s else {})) \ f s" using positiveD1[OF posf] by (subst suminf_finite[of "{0}"]) auto with s show "(INF A\{A. range A \ M \ disjoint_family A \ s \ \(A ` UNIV)}. \i. f (A i)) \ f s" by (intro INF_lower2[of "\i. if i = 0 then s else {}"]) (auto simp: disjoint_family_on_def) qed lemma outer_measure_empty: "positive M f \ {} \ M \ outer_measure M f {} = 0" unfolding outer_measure_def by (intro antisym INF_lower2[of "\_. {}"]) (auto simp: disjoint_family_on_def positive_def) lemma (in ring_of_sets) positive_outer_measure: assumes "positive M f" shows "positive (Pow \) (outer_measure M f)" unfolding positive_def by (auto simp: assms outer_measure_empty) lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \) (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: assumes pos: "positive M f" and inc: "increasing M f" and A: "range A \ M" and X: "X \ (\i. A i)" shows "outer_measure M f X \ (\i. f (A i))" unfolding outer_measure_def proof (safe intro!: INF_lower2[of "disjointed A"] del: subsetI) show dA: "range (disjointed A) \ M" by (auto intro!: A range_disjointed_sets) have "\n. f (disjointed A n) \ f (A n)" by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A) then show "(\i. f (disjointed A i)) \ (\i. f (A i))" by (blast intro!: suminf_le) qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed) lemma (in ring_of_sets) outer_measure_close: "outer_measure M f X < e \ \A. range A \ M \ disjoint_family A \ X \ (\i. A i) \ (\i. f (A i)) < e" unfolding outer_measure_def INF_less_iff by auto lemma (in ring_of_sets) countably_subadditive_outer_measure: assumes posf: "positive M f" and inc: "increasing M f" shows "countably_subadditive (Pow \) (outer_measure M f)" proof (simp add: countably_subadditive_def, safe) fix A :: "nat \ _" assume A: "range A \ Pow (\)" and sb: "(\i. A i) \ \" let ?O = "outer_measure M f" show "?O (\i. A i) \ (\n. ?O (A n))" proof (rule ennreal_le_epsilon) fix b and e :: real assume "0 < e" "(\n. outer_measure M f (A n)) < top" then have *: "\n. outer_measure M f (A n) < outer_measure M f (A n) + e * (1/2)^Suc n" by (auto simp add: less_top dest!: ennreal_suminf_lessD) obtain B where B: "\n. range (B n) \ M" and sbB: "\n. A n \ (\i. B n i)" and Ble: "\n. (\i. f (B n i)) \ ?O (A n) + e * (1/2)^(Suc n)" by (metis less_imp_le outer_measure_close[OF *]) define C where "C = case_prod B o prod_decode" from B have B_in_M: "\i j. B i j \ M" by (rule range_subsetD) then have C: "range C \ M" by (auto simp add: C_def split_def) have A_C: "(\i. A i) \ (\i. C i)" using sbB by (auto simp add: C_def subset_eq) (metis prod.case prod_encode_inverse) have "?O (\i. A i) \ ?O (\i. C i)" using A_C A C by (intro increasing_outer_measure[THEN increasingD]) (auto dest!: sets_into_space) also have "\ \ (\i. f (C i))" using C by (intro outer_measure_le[OF posf inc]) auto also have "\ = (\n. \i. f (B n i))" using B_in_M unfolding C_def comp_def by (intro suminf_ennreal_2dimen) auto also have "\ \ (\n. ?O (A n) + e * (1/2) ^ Suc n)" using B_in_M by (intro suminf_le suminf_nonneg allI Ble) auto also have "... = (\n. ?O (A n)) + (\n. ennreal e * ennreal ((1/2) ^ Suc n))" using \0 < e\ by (subst suminf_add[symmetric]) (auto simp del: ennreal_suminf_cmult simp add: ennreal_mult[symmetric]) also have "\ = (\n. ?O (A n)) + e" unfolding ennreal_suminf_cmult by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto finally show "?O (\i. A i) \ (\n. ?O (A n)) + e" . qed qed lemma (in ring_of_sets) outer_measure_space_outer_measure: "positive M f \ increasing M f \ outer_measure_space (Pow \) (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: assumes posf: "positive M f" and inc: "increasing M f" and add: "additive M f" shows "M \ lambda_system \ (Pow \) (outer_measure M f)" proof (auto dest: sets_into_space simp add: algebra.lambda_system_eq [OF algebra_Pow]) fix x s assume x: "x \ M" and s: "s \ \" have [simp]: "\x. x \ M \ s \ (\ - x) = s - x" using s by blast have "outer_measure M f (s \ x) + outer_measure M f (s - x) \ outer_measure M f s" unfolding outer_measure_def[of M f s] proof (safe intro!: INF_greatest) fix A :: "nat \ 'a set" assume A: "disjoint_family A" "range A \ M" "s \ (\i. A i)" have "outer_measure M f (s \ x) \ (\i. f (A i \ x))" unfolding outer_measure_def proof (safe intro!: INF_lower2[of "\i. A i \ x"]) from A(1) show "disjoint_family (\i. A i \ x)" by (rule disjoint_family_on_bisimulation) auto qed (insert x A, auto) moreover have "outer_measure M f (s - x) \ (\i. f (A i - x))" unfolding outer_measure_def proof (safe intro!: INF_lower2[of "\i. A i - x"]) from A(1) show "disjoint_family (\i. A i - x)" by (rule disjoint_family_on_bisimulation) auto qed (insert x A, auto) ultimately have "outer_measure M f (s \ x) + outer_measure M f (s - x) \ (\i. f (A i \ x)) + (\i. f (A i - x))" by (rule add_mono) also have "\ = (\i. f (A i \ x) + f (A i - x))" using A(2) x posf by (subst suminf_add) (auto simp: positive_def) also have "\ = (\i. f (A i))" using A x by (subst add[THEN additiveD, symmetric]) (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f]) finally show "outer_measure M f (s \ x) + outer_measure M f (s - x) \ (\i. f (A i))" . qed moreover have "outer_measure M f s \ outer_measure M f (s \ x) + outer_measure M f (s - x)" proof - have "outer_measure M f s = outer_measure M f ((s \ x) \ (s - x))" by (metis Un_Diff_Int Un_commute) also have "... \ outer_measure M f (s \ x) + outer_measure M f (s - x)" apply (rule subadditiveD) apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow]) apply (simp add: positive_def outer_measure_empty[OF posf]) apply (rule countably_subadditive_outer_measure) using s by (auto intro!: posf inc) finally show ?thesis . qed ultimately show "outer_measure M f (s \ x) + outer_measure M f (s - x) = outer_measure M f s" by (rule order_antisym) qed lemma measure_down: "measure_space \ N \ \ sigma_algebra \ M \ M \ N \ measure_space \ M \" by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq) subsection \Caratheodory's theorem\ theorem (in ring_of_sets) caratheodory': assumes posf: "positive M f" and ca: "countably_additive M f" shows "\\ :: 'a set \ ennreal. (\s \ M. \ s = f s) \ measure_space \ (sigma_sets \ M) \" proof - have inc: "increasing M f" by (metis additive_increasing ca countably_additive_additive posf) let ?O = "outer_measure M f" define ls where "ls = lambda_system \ (Pow \) ?O" have mls: "measure_space \ ls ?O" using sigma_algebra.caratheodory_lemma [OF sigma_algebra_Pow outer_measure_space_outer_measure [OF posf inc]] by (simp add: ls_def) hence sls: "sigma_algebra \ ls" by (simp add: measure_space_def) have "M \ ls" by (simp add: ls_def) (metis ca posf inc countably_additive_additive algebra_subset_lambda_system) hence sgs_sb: "sigma_sets (\) (M) \ ls" using sigma_algebra.sigma_sets_subset [OF sls, of "M"] by simp have "measure_space \ (sigma_sets \ M) ?O" by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets) (simp_all add: sgs_sb space_closed) thus ?thesis using outer_measure_agrees [OF posf ca] by (intro exI[of _ ?O]) auto qed lemma (in ring_of_sets) caratheodory_empty_continuous: assumes f: "positive M f" "additive M f" and fin: "\A. A \ M \ f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" shows "\\ :: 'a set \ ennreal. (\s \ M. \ s = f s) \ measure_space \ (sigma_sets \ M) \" proof (intro caratheodory' empty_continuous_imp_countably_additive f) show "\A\M. f A \ \" using fin by auto qed (rule cont) subsection \Volumes\ definition\<^marker>\tag important\ volume :: "'a set set \ ('a set \ ennreal) \ bool" where "volume M f \ (f {} = 0) \ (\a\M. 0 \ f a) \ (\C\M. disjoint C \ finite C \ \C \ M \ f (\C) = (\c\C. f c))" lemma volumeI: assumes "f {} = 0" assumes "\a. a \ M \ 0 \ f a" assumes "\C. C \ M \ disjoint C \ finite C \ \C \ M \ f (\C) = (\c\C. f c)" shows "volume M f" using assms by (auto simp: volume_def) lemma volume_positive: "volume M f \ a \ M \ 0 \ f a" by (auto simp: volume_def) lemma volume_empty: "volume M f \ f {} = 0" by (auto simp: volume_def) proposition volume_finite_additive: assumes "volume M f" assumes A: "\i. i \ I \ A i \ M" "disjoint_family_on A I" "finite I" "\(A ` I) \ M" shows "f (\(A ` I)) = (\i\I. f (A i))" proof - have "A`I \ M" "disjoint (A`I)" "finite (A`I)" "\(A`I) \ M" using A by (auto simp: disjoint_family_on_disjoint_image) with \volume M f\ have "f (\(A`I)) = (\a\A`I. f a)" unfolding volume_def by blast also have "\ = (\i\I. f (A i))" proof (subst sum.reindex_nontrivial) fix i j assume "i \ I" "j \ I" "i \ j" "A i = A j" with \disjoint_family_on A I\ have "A i = {}" by (auto simp: disjoint_family_on_def) then show "f (A i) = 0" using volume_empty[OF \volume M f\] by simp qed (auto intro: \finite I\) finally show "f (\(A ` I)) = (\i\I. f (A i))" by simp qed lemma (in ring_of_sets) volume_additiveI: assumes pos: "\a. a \ M \ 0 \ \ a" assumes [simp]: "\ {} = 0" assumes add: "\a b. a \ M \ b \ M \ a \ b = {} \ \ (a \ b) = \ a + \ b" shows "volume M \" proof (unfold volume_def, safe) fix C assume "finite C" "C \ M" "disjoint C" then show "\ (\C) = sum \ C" proof (induct C) case (insert c C) from insert(1,2,4,5) have "\ (\(insert c C)) = \ c + \ (\C)" by (auto intro!: add simp: disjoint_def) with insert show ?case by (simp add: disjoint_def) qed simp qed fact+ proposition (in semiring_of_sets) extend_volume: assumes "volume M \" shows "\\'. volume generated_ring \' \ (\a\M. \' a = \ a)" proof - let ?R = generated_ring have "\a\?R. \m. \C\M. a = \C \ finite C \ disjoint C \ m = (\c\C. \ c)" by (auto simp: generated_ring_def) - from bchoice[OF this] guess \' .. note \'_spec = this - + from bchoice[OF this] obtain \' + where \'_spec: "\x\generated_ring. \C\M. x = \ C \ finite C \ disjoint C \ \' x = sum \ C" + by blast { fix C assume C: "C \ M" "finite C" "disjoint C" fix D assume D: "D \ M" "finite D" "disjoint D" assume "\C = \D" have "(\d\D. \ d) = (\d\D. \c\C. \ (c \ d))" proof (intro sum.cong refl) fix d assume "d \ D" have Un_eq_d: "(\c\C. c \ d) = d" using \d \ D\ \\C = \D\ by auto moreover have "\ (\c\C. c \ d) = (\c\C. \ (c \ d))" proof (rule volume_finite_additive) { fix c assume "c \ C" then show "c \ d \ M" using C D \d \ D\ by auto } show "(\a\C. a \ d) \ M" unfolding Un_eq_d using \d \ D\ D by auto show "disjoint_family_on (\a. a \ d) C" using \disjoint C\ by (auto simp: disjoint_family_on_def disjoint_def) qed fact+ ultimately show "\ d = (\c\C. \ (c \ d))" by simp qed } note split_sum = this { fix C assume C: "C \ M" "finite C" "disjoint C" fix D assume D: "D \ M" "finite D" "disjoint D" assume "\C = \D" with split_sum[OF C D] split_sum[OF D C] have "(\d\D. \ d) = (\c\C. \ c)" by (simp, subst sum.swap, simp add: ac_simps) } note sum_eq = this { fix C assume C: "C \ M" "finite C" "disjoint C" then have "\C \ ?R" by (auto simp: generated_ring_def) with \'_spec[THEN bspec, of "\C"] obtain D where D: "D \ M" "finite D" "disjoint D" "\C = \D" and "\' (\C) = (\d\D. \ d)" by auto with sum_eq[OF C D] have "\' (\C) = (\c\C. \ c)" by simp } note \' = this show ?thesis proof (intro exI conjI ring_of_sets.volume_additiveI[OF generating_ring] ballI) fix a assume "a \ M" with \'[of "{a}"] show "\' a = \ a" by (simp add: disjoint_def) next - fix a assume "a \ ?R" then guess Ca .. note Ca = this + fix a assume "a \ ?R" + then obtain Ca where Ca: "finite Ca" "disjoint Ca" "Ca \ M" "a = \ Ca" .. with \'[of Ca] \volume M \\[THEN volume_positive] show "0 \ \' a" by (auto intro!: sum_nonneg) next show "\' {} = 0" using \'[of "{}"] by auto next - fix a assume "a \ ?R" then guess Ca .. note Ca = this - fix b assume "b \ ?R" then guess Cb .. note Cb = this + 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" + by (meson generated_ringE) assume "a \ b = {}" with Ca Cb have "Ca \ Cb \ {{}}" by auto then have C_Int_cases: "Ca \ Cb = {{}} \ Ca \ Cb = {}" by auto from \a \ b = {}\ have "\' (\(Ca \ Cb)) = (\c\Ca \ Cb. \ c)" using Ca Cb by (intro \') (auto intro!: disjoint_union) also have "\ = (\c\Ca \ Cb. \ c) + (\c\Ca \ Cb. \ c)" using C_Int_cases volume_empty[OF \volume M \\] by (elim disjE) simp_all also have "\ = (\c\Ca. \ c) + (\c\Cb. \ c)" using Ca Cb by (simp add: sum.union_inter) also have "\ = \' a + \' b" using Ca Cb by (simp add: \') finally show "\' (a \ b) = \' a + \' b" using Ca Cb by simp qed qed subsubsection\<^marker>\tag important\ \Caratheodory on semirings\ theorem (in semiring_of_sets) caratheodory: assumes pos: "positive M \" and ca: "countably_additive M \" shows "\\' :: 'a set \ ennreal. (\s \ M. \' s = \ s) \ measure_space \ (sigma_sets \ M) \'" proof - have "volume M \" proof (rule volumeI) { fix a assume "a \ M" then show "0 \ \ a" using pos unfolding positive_def by auto } note p = this fix C assume sets_C: "C \ M" "\C \ M" and "disjoint C" "finite C" have "\F'. bij_betw F' {..finite C\]) auto - then guess F' .. note F' = this + then obtain F' where "bij_betw F' {.. j" with F' have "F' i \ C" "F' j \ C" "F' i \ F' j" unfolding inj_on_def by auto with \disjoint C\[THEN disjointD] have "F' i \ F' j = {}" by auto } note F'_disj = this define F where "F i = (if i < card C then F' i else {})" for i then have "disjoint_family F" using F'_disj by (auto simp: disjoint_family_on_def) moreover from F' have "(\i. F i) = \C" by (auto simp add: F_def split: if_split_asm cong del: SUP_cong) moreover have sets_F: "\i. F i \ M" using F' sets_C by (auto simp: F_def) moreover note sets_C ultimately have "\ (\C) = (\i. \ (F i))" using ca[unfolded countably_additive_def, THEN spec, of F] by auto also have "\ = (\i (F' i))" proof - have "(\i. if i \ {..< card C} then \ (F' i) else 0) sums (\i (F' i))" by (rule sums_If_finite_set) auto also have "(\i. if i \ {..< card C} then \ (F' i) else 0) = (\i. \ (F i))" using pos by (auto simp: positive_def F_def) finally show "(\i. \ (F i)) = (\i (F' i))" by (simp add: sums_iff) qed also have "\ = (\c\C. \ c)" using F'(2) by (subst (2) F') (simp add: sum.reindex) finally show "\ (\C) = (\c\C. \ c)" . next show "\ {} = 0" using \positive M \\ by (rule positiveD1) qed from extend_volume[OF this] obtain \_r where V: "volume generated_ring \_r" "\a. a \ M \ \ a = \_r a" by auto interpret G: ring_of_sets \ generated_ring by (rule generating_ring) have pos: "positive generated_ring \_r" using V unfolding positive_def by (auto simp: positive_def intro!: volume_positive volume_empty) have "countably_additive generated_ring \_r" proof (rule countably_additiveI) - fix A' :: "nat \ 'a set" assume A': "range A' \ generated_ring" "disjoint_family A'" + fix A' :: "nat \ 'a set" + assume A': "range A' \ generated_ring" "disjoint_family A'" and Un_A: "(\i. A' i) \ generated_ring" - from generated_ringE[OF Un_A] guess C' . note C' = this + obtain C' where C': "finite C'" "disjoint C'" "C' \ M" "\ (range A') = \ C'" + using generated_ringE[OF Un_A] by auto { fix c assume "c \ C'" moreover define A where [abs_def]: "A i = A' i \ c" for i ultimately have A: "range A \ generated_ring" "disjoint_family A" and Un_A: "(\i. A i) \ generated_ring" using A' C' by (auto intro!: G.Int G.finite_Union intro: generated_ringI_Basic simp: disjoint_family_on_def) from A C' \c \ C'\ have UN_eq: "(\i. A i) = c" by (auto simp: A_def) have "\i::nat. \f::nat \ 'a set. \_r (A i) = (\j. \_r (f j)) \ disjoint_family f \ \(range f) = A i \ (\j. f j \ M)" (is "\i. ?P i") proof fix i from A have Ai: "A i \ generated_ring" by auto - from generated_ringE[OF this] guess C . note C = this - + from generated_ringE[OF this] obtain C + where C: "finite C" "disjoint C" "C \ M" "A i = \ C" by auto have "\F'. bij_betw F' {..finite C\]) auto - then guess F .. note F = this + then obtain F where F: "bij_betw F {..j. f j \ M" by (auto simp: Pi_iff f_def dest!: bij_betw_imp_funcset) moreover from f C have d_f: "disjoint_family_on f {..x(range f) = A i" using f by (auto simp add: f_def) moreover { have "(\j. \_r (f j)) = (\j. if j \ {..< card C} then \_r (f j) else 0)" using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def) also have "\ = (\j_r (f j))" by (rule sums_If_finite_set[THEN sums_unique, symmetric]) simp also have "\ = \_r (A i)" using C f[THEN bij_betw_imp_funcset] unfolding Ai_eq by (intro volume_finite_additive[OF V(1) _ d_f, symmetric]) (auto simp: Pi_iff Ai_eq intro: generated_ringI_Basic) finally have "\_r (A i) = (\j. \_r (f j))" .. } ultimately show "?P i" by blast qed - from choice[OF this] guess f .. note f = this + from choice[OF this] obtain f + where f: "\x. \_r (A x) = (\j. \_r (f x j)) \ disjoint_family (f x) \ \ (range (f x)) = A x \ (\j. f x j \ M)" + .. then have UN_f_eq: "(\i. case_prod f (prod_decode i)) = (\i. A i)" unfolding UN_extend_simps surj_prod_decode by (auto simp: set_eq_iff) have d: "disjoint_family (\i. case_prod f (prod_decode i))" unfolding disjoint_family_on_def proof (intro ballI impI) fix m n :: nat assume "m \ n" then have neq: "prod_decode m \ prod_decode n" using inj_prod_decode[of UNIV] by (auto simp: inj_on_def) show "case_prod f (prod_decode m) \ case_prod f (prod_decode n) = {}" proof cases assume "fst (prod_decode m) = fst (prod_decode n)" then show ?thesis using neq f by (fastforce simp: disjoint_family_on_def) next assume neq: "fst (prod_decode m) \ fst (prod_decode n)" have "case_prod f (prod_decode m) \ A (fst (prod_decode m))" "case_prod f (prod_decode n) \ A (fst (prod_decode n))" using f[THEN spec, of "fst (prod_decode m)"] using f[THEN spec, of "fst (prod_decode n)"] by (auto simp: set_eq_iff) with f A neq show ?thesis by (fastforce simp: disjoint_family_on_def subset_eq set_eq_iff) qed qed from f have "(\n. \_r (A n)) = (\n. \_r (case_prod f (prod_decode n)))" by (intro suminf_ennreal_2dimen[symmetric] generated_ringI_Basic) (auto split: prod.split) also have "\ = (\n. \ (case_prod f (prod_decode n)))" using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split) also have "\ = \ (\i. case_prod f (prod_decode i))" using f \c \ C'\ C' by (intro ca[unfolded countably_additive_def, rule_format]) (auto split: prod.split simp: UN_f_eq d UN_eq) finally have "(\n. \_r (A' n \ c)) = \ c" using UN_f_eq UN_eq by (simp add: A_def) } note eq = this have "(\n. \_r (A' n)) = (\n. \c\C'. \_r (A' n \ c))" using C' A' by (subst volume_finite_additive[symmetric, OF V(1)]) (auto simp: disjoint_def disjoint_family_on_def intro!: G.Int G.finite_Union arg_cong[where f="\X. suminf (\i. \_r (X i))"] ext intro: generated_ringI_Basic) also have "\ = (\c\C'. \n. \_r (A' n \ c))" using C' A' by (intro suminf_sum G.Int G.finite_Union) (auto intro: generated_ringI_Basic) also have "\ = (\c\C'. \_r c)" using eq V C' by (auto intro!: sum.cong) also have "\ = \_r (\C')" using C' Un_A by (subst volume_finite_additive[symmetric, OF V(1)]) (auto simp: disjoint_family_on_def disjoint_def intro: generated_ringI_Basic) finally show "(\n. \_r (A' n)) = \_r (\i. A' i)" using C' by simp qed - from G.caratheodory'[OF \positive generated_ring \_r\ \countably_additive generated_ring \_r\] - guess \' .. + obtain \' where "(\s\generated_ring. \' s = \_r s) \ measure_space \ (sigma_sets \ generated_ring) \'" + using G.caratheodory'[OF pos \countably_additive generated_ring \_r\] by auto with V show ?thesis unfolding sigma_sets_generated_ring_eq by (intro exI[of _ \']) (auto intro: generated_ringI_Basic) qed lemma extend_measure_caratheodory: fixes G :: "'i \ 'a set" assumes M: "M = extend_measure \ I G \" assumes "i \ I" assumes "semiring_of_sets \ (G ` I)" assumes empty: "\i. i \ I \ G i = {} \ \ i = 0" assumes inj: "\i j. i \ I \ j \ I \ G i = G j \ \ i = \ j" assumes nonneg: "\i. i \ I \ 0 \ \ i" assumes add: "\A::nat \ 'i. \j. A \ UNIV \ I \ j \ I \ disjoint_family (G \ A) \ (\i. G (A i)) = G j \ (\n. \ (A n)) = \ j" shows "emeasure M (G i) = \ i" proof - interpret semiring_of_sets \ "G ` I" by fact have "\g\G`I. \i\I. g = G i" by auto then obtain sel where sel: "\g. g \ G ` I \ sel g \ I" "\g. g \ G ` I \ G (sel g) = g" by metis have "\\'. (\s\G ` I. \' s = \ (sel s)) \ measure_space \ (sigma_sets \ (G ` I)) \'" proof (rule caratheodory) show "positive (G ` I) (\s. \ (sel s))" by (auto simp: positive_def intro!: empty sel nonneg) show "countably_additive (G ` I) (\s. \ (sel s))" proof (rule countably_additiveI) fix A :: "nat \ 'a set" assume "range A \ G ` I" "disjoint_family A" "(\i. A i) \ G ` I" then show "(\i. \ (sel (A i))) = \ (sel (\i. A i))" by (intro add) (auto simp: sel image_subset_iff_funcset comp_def Pi_iff intro!: sel) qed qed then obtain \' where \': "\s\G ` I. \' s = \ (sel s)" "measure_space \ (sigma_sets \ (G ` I)) \'" by metis show ?thesis proof (rule emeasure_extend_measure[OF M]) { fix i assume "i \ I" then show "\' (G i) = \ i" using \' by (auto intro!: inj sel) } show "G ` I \ Pow \" by (rule space_closed) then show "positive (sets M) \'" "countably_additive (sets M) \'" using \' by (simp_all add: M sets_extend_measure measure_space_def) qed fact qed proposition extend_measure_caratheodory_pair: fixes G :: "'i \ 'j \ 'a set" assumes M: "M = extend_measure \ {(a, b). P a b} (\(a, b). G a b) (\(a, b). \ a b)" assumes "P i j" assumes semiring: "semiring_of_sets \ {G a b | a b. P a b}" assumes empty: "\i j. P i j \ G i j = {} \ \ i j = 0" assumes inj: "\i j k l. P i j \ P k l \ G i j = G k l \ \ i j = \ k l" assumes nonneg: "\i j. P i j \ 0 \ \ i j" assumes add: "\A::nat \ 'i. \B::nat \ 'j. \j k. (\n. P (A n) (B n)) \ P j k \ disjoint_family (\n. G (A n) (B n)) \ (\i. G (A i) (B i)) = G j k \ (\n. \ (A n) (B n)) = \ j k" shows "emeasure M (G i j) = \ i j" proof - have "emeasure M ((\(a, b). G a b) (i, j)) = (\(a, b). \ a b) (i, j)" proof (rule extend_measure_caratheodory[OF M]) show "semiring_of_sets \ ((\(a, b). G a b) ` {(a, b). P a b})" using semiring by (simp add: image_def conj_commute) next fix A :: "nat \ ('i \ 'j)" and j assume "A \ UNIV \ {(a, b). P a b}" "j \ {(a, b). P a b}" "disjoint_family ((\(a, b). G a b) \ A)" "(\i. case A i of (a, b) \ G a b) = (case j of (a, b) \ G a b)" then show "(\n. case A n of (a, b) \ \ a b) = (case j of (a, b) \ \ a b)" using add[of "\i. fst (A i)" "\i. snd (A i)" "fst j" "snd j"] by (simp add: split_beta' comp_def Pi_iff) qed (auto split: prod.splits intro: assms) then show ?thesis by simp qed end diff --git a/src/HOL/Analysis/Complete_Measure.thy b/src/HOL/Analysis/Complete_Measure.thy --- a/src/HOL/Analysis/Complete_Measure.thy +++ b/src/HOL/Analysis/Complete_Measure.thy @@ -1,1159 +1,1164 @@ (* Title: HOL/Analysis/Complete_Measure.thy Author: Robert Himmelmann, Johannes Hoelzl, TU Muenchen *) section \Complete Measures\ theory Complete_Measure imports Bochner_Integration begin locale\<^marker>\tag important\ complete_measure = fixes M :: "'a measure" assumes complete: "\A B. B \ A \ A \ null_sets M \ B \ sets M" definition\<^marker>\tag important\ "split_completion M A p = (if A \ sets M then p = (A, {}) else \N'. A = fst p \ snd p \ fst p \ snd p = {} \ fst p \ sets M \ snd p \ N' \ N' \ null_sets M)" definition\<^marker>\tag important\ "main_part M A = fst (Eps (split_completion M A))" definition\<^marker>\tag important\ "null_part M A = snd (Eps (split_completion M A))" definition\<^marker>\tag important\ completion :: "'a measure \ 'a measure" where "completion M = measure_of (space M) { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' } (emeasure M \ main_part M)" lemma completion_into_space: "{ S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' } \ Pow (space M)" using sets.sets_into_space by auto lemma space_completion[simp]: "space (completion M) = space M" unfolding completion_def using space_measure_of[OF completion_into_space] by simp lemma completionI: assumes "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" shows "A \ { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" using assms by auto lemma completionE: assumes "A \ { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" obtains S N N' where "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" using assms by auto lemma sigma_algebra_completion: "sigma_algebra (space M) { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" (is "sigma_algebra _ ?A") unfolding sigma_algebra_iff2 proof (intro conjI ballI allI impI) show "?A \ Pow (space M)" using sets.sets_into_space by auto next show "{} \ ?A" by auto next let ?C = "space M" - fix A assume "A \ ?A" from completionE[OF this] guess S N N' . + fix A assume "A \ ?A" + then obtain S N N' + where "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" + by (rule completionE) then show "space M - A \ ?A" by (intro completionI[of _ "(?C - S) \ (?C - N')" "(?C - S) \ N' \ (?C - N)"]) auto next fix A :: "nat \ 'a set" assume A: "range A \ ?A" then have "\n. \S N N'. A n = S \ N \ S \ sets M \ N' \ null_sets M \ N \ N'" by (auto simp: image_subset_iff) - from choice[OF this] guess S .. - from choice[OF this] guess N .. - from choice[OF this] guess N' .. + then obtain S N N' where "\x. A x = S x \ N x \ S x \ sets M \ N' x \ null_sets M \ N x \ N' x" + by metis then show "\(A ` UNIV) \ ?A" using null_sets_UN[of N'] by (intro completionI[of _ "\(S ` UNIV)" "\(N ` UNIV)" "\(N' ` UNIV)"]) auto qed lemma\<^marker>\tag important\ sets_completion: "sets (completion M) = { S \ N |S N N'. S \ sets M \ N' \ null_sets M \ N \ N' }" using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion] by (simp add: completion_def) lemma sets_completionE: assumes "A \ sets (completion M)" obtains S N N' where "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" using assms unfolding sets_completion by auto lemma sets_completionI: assumes "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" shows "A \ sets (completion M)" using assms unfolding sets_completion by auto lemma sets_completionI_sets[intro, simp]: "A \ sets M \ A \ sets (completion M)" unfolding sets_completion by force lemma\<^marker>\tag important\ measurable_completion: "f \ M \\<^sub>M N \ f \ completion M \\<^sub>M N" by (auto simp: measurable_def) lemma null_sets_completion: assumes "N' \ null_sets M" "N \ N'" shows "N \ sets (completion M)" using assms by (intro sets_completionI[of N "{}" N N']) auto lemma\<^marker>\tag important\ split_completion: assumes "A \ sets (completion M)" shows "split_completion M A (main_part M A, null_part M A)" proof cases assume "A \ sets M" then show ?thesis by (simp add: split_completion_def[abs_def] main_part_def null_part_def) next assume nA: "A \ sets M" show ?thesis unfolding main_part_def null_part_def if_not_P[OF nA] proof (rule someI2_ex) - from assms[THEN sets_completionE] guess S N N' . note A = this + from assms obtain S N N' where A: "A = S \ N" "N \ N'" "N' \ null_sets M" "S \ sets M" + by (blast intro: sets_completionE) let ?P = "(S, N - S)" show "\p. split_completion M A p" unfolding split_completion_def if_not_P[OF nA] using A proof (intro exI conjI) show "A = fst ?P \ snd ?P" using A by auto show "snd ?P \ N'" using A by auto qed auto qed auto qed lemma sets_restrict_space_subset: assumes S: "S \ sets (completion M)" shows "sets (restrict_space (completion M) S) \ sets (completion M)" by (metis assms sets.Int_space_eq2 sets_restrict_space_iff subsetI) lemma assumes "S \ sets (completion M)" shows main_part_sets[intro, simp]: "main_part M S \ sets M" and main_part_null_part_Un[simp]: "main_part M S \ null_part M S = S" and main_part_null_part_Int[simp]: "main_part M S \ null_part M S = {}" using split_completion[OF assms] by (auto simp: split_completion_def split: if_split_asm) lemma main_part[simp]: "S \ sets M \ main_part M S = S" using split_completion[of S M] by (auto simp: split_completion_def split: if_split_asm) lemma null_part: assumes "S \ sets (completion M)" shows "\N. N\null_sets M \ null_part M S \ N" using split_completion[OF assms] by (auto simp: split_completion_def split: if_split_asm) lemma null_part_sets[intro, simp]: assumes "S \ sets M" shows "null_part M S \ sets M" "emeasure M (null_part M S) = 0" proof - have S: "S \ sets (completion M)" using assms by auto - have "S - main_part M S \ sets M" using assms by auto - moreover + have *: "S - main_part M S \ sets M" using assms by auto from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S] have "S - main_part M S = null_part M S" by auto - ultimately show sets: "null_part M S \ sets M" by auto - from null_part[OF S] guess N .. + with * show sets: "null_part M S \ sets M" by auto + from null_part[OF S] obtain N where "N \ null_sets M \ null_part M S \ N" .. with emeasure_eq_0[of N _ "null_part M S"] sets show "emeasure M (null_part M S) = 0" by auto qed lemma emeasure_main_part_UN: fixes S :: "nat \ 'a set" assumes "range S \ sets (completion M)" shows "emeasure M (main_part M (\i. (S i))) = emeasure M (\i. main_part M (S i))" proof - have S: "\i. S i \ sets (completion M)" using assms by auto then have UN: "(\i. S i) \ sets (completion M)" by auto have "\i. \N. N \ null_sets M \ null_part M (S i) \ N" using null_part[OF S] by auto - from choice[OF this] guess N .. note N = this + then obtain N where N: "\x. N x \ null_sets M \ null_part M (S x) \ N x" by metis then have UN_N: "(\i. N i) \ null_sets M" by (intro null_sets_UN) auto - have "(\i. S i) \ sets (completion M)" using S by auto - from null_part[OF this] guess N' .. note N' = this + from S have "(\i. S i) \ sets (completion M)" by auto + from null_part[OF this] obtain N' where N': "N' \ null_sets M \ null_part M (\ (range S)) \ N'" .. let ?N = "(\i. N i) \ N'" have null_set: "?N \ null_sets M" using N' UN_N by (intro null_sets.Un) auto have "main_part M (\i. S i) \ ?N = (main_part M (\i. S i) \ null_part M (\i. S i)) \ ?N" using N' by auto also have "\ = (\i. main_part M (S i) \ null_part M (S i)) \ ?N" unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto also have "\ = (\i. main_part M (S i)) \ ?N" using N by auto finally have *: "main_part M (\i. S i) \ ?N = (\i. main_part M (S i)) \ ?N" . have "emeasure M (main_part M (\i. S i)) = emeasure M (main_part M (\i. S i) \ ?N)" using null_set UN by (intro emeasure_Un_null_set[symmetric]) auto also have "\ = emeasure M ((\i. main_part M (S i)) \ ?N)" unfolding * .. also have "\ = emeasure M (\i. main_part M (S i))" using null_set S by (intro emeasure_Un_null_set) auto finally show ?thesis . qed lemma\<^marker>\tag important\ emeasure_completion[simp]: assumes S: "S \ sets (completion M)" shows "emeasure (completion M) S = emeasure M (main_part M S)" proof (subst emeasure_measure_of[OF completion_def completion_into_space]) let ?\ = "emeasure M \ main_part M" show "S \ sets (completion M)" "?\ S = emeasure M (main_part M S) " using S by simp_all show "positive (sets (completion M)) ?\" by (simp add: positive_def) show "countably_additive (sets (completion M)) ?\" proof (intro countably_additiveI) fix A :: "nat \ 'a set" assume A: "range A \ sets (completion M)" "disjoint_family A" have "disjoint_family (\i. main_part M (A i))" proof (intro disjoint_family_on_bisimulation[OF A(2)]) fix n m assume "A n \ A m = {}" then have "(main_part M (A n) \ null_part M (A n)) \ (main_part M (A m) \ null_part M (A m)) = {}" using A by (subst (1 2) main_part_null_part_Un) auto then show "main_part M (A n) \ main_part M (A m) = {}" by auto qed then have "(\n. emeasure M (main_part M (A n))) = emeasure M (\i. main_part M (A i))" using A by (auto intro!: suminf_emeasure) then show "(\n. ?\ (A n)) = ?\ (\(A ` UNIV))" by (simp add: completion_def emeasure_main_part_UN[OF A(1)]) qed qed lemma measure_completion[simp]: "S \ sets M \ measure (completion M) S = measure M S" unfolding measure_def by auto lemma emeasure_completion_UN: "range S \ sets (completion M) \ emeasure (completion M) (\i::nat. (S i)) = emeasure M (\i. main_part M (S i))" by (subst emeasure_completion) (auto simp add: emeasure_main_part_UN) lemma emeasure_completion_Un: assumes S: "S \ sets (completion M)" and T: "T \ sets (completion M)" shows "emeasure (completion M) (S \ T) = emeasure M (main_part M S \ main_part M T)" proof (subst emeasure_completion) have UN: "(\i. binary (main_part M S) (main_part M T) i) = (\i. main_part M (binary S T i))" unfolding binary_def by (auto split: if_split_asm) show "emeasure M (main_part M (S \ T)) = emeasure M (main_part M S \ main_part M T)" using emeasure_main_part_UN[of "binary S T" M] assms by (simp add: range_binary_eq, simp add: Un_range_binary UN) qed (auto intro: S T) lemma sets_completionI_sub: assumes N: "N' \ null_sets M" "N \ N'" shows "N \ sets (completion M)" using assms by (intro sets_completionI[of _ "{}" N N']) auto lemma completion_ex_simple_function: assumes f: "simple_function (completion M) f" shows "\f'. simple_function M f' \ (AE x in M. f x = f' x)" proof - let ?F = "\x. f -` {x} \ space M" have F: "\x. ?F x \ sets (completion M)" and fin: "finite (f`space M)" using simple_functionD[OF f] simple_functionD[OF f] by simp_all have "\x. \N. N \ null_sets M \ null_part M (?F x) \ N" using F null_part by auto from choice[OF this] obtain N where N: "\x. null_part M (?F x) \ N x" "\x. N x \ null_sets M" by auto let ?N = "\x\f`space M. N x" let ?f' = "\x. if x \ ?N then undefined else f x" have sets: "?N \ null_sets M" using N fin by (intro null_sets.finite_UN) auto show ?thesis unfolding simple_function_def proof (safe intro!: exI[of _ ?f']) have "?f' ` space M \ f`space M \ {undefined}" by auto from finite_subset[OF this] simple_functionD(1)[OF f] show "finite (?f' ` space M)" by auto next fix x assume "x \ space M" have "?f' -` {?f' x} \ space M = (if x \ ?N then ?F undefined \ ?N else if f x = undefined then ?F (f x) \ ?N else ?F (f x) - ?N)" using N(2) sets.sets_into_space by (auto split: if_split_asm simp: null_sets_def) moreover { fix y have "?F y \ ?N \ sets M" proof cases assume y: "y \ f`space M" have "?F y \ ?N = (main_part M (?F y) \ null_part M (?F y)) \ ?N" using main_part_null_part_Un[OF F] by auto also have "\ = main_part M (?F y) \ ?N" using y N by auto finally show ?thesis using F sets by auto next assume "y \ f`space M" then have "?F y = {}" by auto then show ?thesis using sets by auto qed } moreover { have "?F (f x) - ?N = main_part M (?F (f x)) \ null_part M (?F (f x)) - ?N" using main_part_null_part_Un[OF F] by auto also have "\ = main_part M (?F (f x)) - ?N" using N \x \ space M\ by auto finally have "?F (f x) - ?N \ sets M" using F sets by auto } ultimately show "?f' -` {?f' x} \ space M \ sets M" by auto next show "AE x in M. f x = ?f' x" by (rule AE_I', rule sets) auto qed qed lemma\<^marker>\tag important\ completion_ex_borel_measurable: fixes g :: "'a \ ennreal" assumes g: "g \ borel_measurable (completion M)" shows "\g'\borel_measurable M. (AE x in M. g x = g' x)" proof - - from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this - from this(1)[THEN completion_ex_simple_function] + obtain f :: "nat \ 'a \ ennreal" + where *: "\i. simple_function (completion M) (f i)" + and **: "\x. (SUP i. f i x) = g x" + using g[THEN borel_measurable_implies_simple_function_sequence'] by metis + from *[THEN completion_ex_simple_function] have "\i. \f'. simple_function M f' \ (AE x in M. f i x = f' x)" .. - from this[THEN choice] obtain f' where - sf: "\i. simple_function M (f' i)" and - AE: "\i. AE x in M. f i x = f' i x" by auto + then obtain f' + where sf: "\i. simple_function M (f' i)" + and AE: "\i. AE x in M. f i x = f' i x" + by metis show ?thesis proof (intro bexI) from AE[unfolded AE_all_countable[symmetric]] show "AE x in M. g x = (SUP i. f' i x)" (is "AE x in M. g x = ?f x") proof (elim AE_mp, safe intro!: AE_I2) fix x assume eq: "\i. f i x = f' i x" - moreover have "g x = (SUP i. f i x)" - unfolding f by (auto split: split_max) - ultimately show "g x = ?f x" by auto + have "g x = (SUP i. f i x)" by (auto simp: ** split: split_max) + with eq show "g x = ?f x" by auto qed show "?f \ borel_measurable M" using sf[THEN borel_measurable_simple_function] by auto qed qed lemma null_sets_completionI: "N \ null_sets M \ N \ null_sets (completion M)" by (auto simp: null_sets_def) lemma AE_completion: "(AE x in M. P x) \ (AE x in completion M. P x)" unfolding eventually_ae_filter by (auto intro: null_sets_completionI) lemma null_sets_completion_iff: "N \ sets M \ N \ null_sets (completion M) \ N \ null_sets M" by (auto simp: null_sets_def) lemma sets_completion_AE: "(AE x in M. \ P x) \ Measurable.pred (completion M) P" unfolding pred_def sets_completion eventually_ae_filter by auto lemma null_sets_completion_iff2: "A \ null_sets (completion M) \ (\N'\null_sets M. A \ N')" proof safe assume "A \ null_sets (completion M)" then have A: "A \ sets (completion M)" and "main_part M A \ null_sets M" by (auto simp: null_sets_def) moreover obtain N where "N \ null_sets M" "null_part M A \ N" using null_part[OF A] by auto ultimately show "\N'\null_sets M. A \ N'" proof (intro bexI) show "A \ N \ main_part M A" using \null_part M A \ N\ by (subst main_part_null_part_Un[OF A, symmetric]) auto qed auto next fix N assume "N \ null_sets M" "A \ N" then have "A \ sets (completion M)" and N: "N \ sets M" "A \ N" "emeasure M N = 0" by (auto intro: null_sets_completion) moreover have "emeasure (completion M) A = 0" using N by (intro emeasure_eq_0[of N _ A]) auto ultimately show "A \ null_sets (completion M)" by auto qed lemma null_sets_completion_subset: "B \ A \ A \ null_sets (completion M) \ B \ null_sets (completion M)" unfolding null_sets_completion_iff2 by auto interpretation completion: complete_measure "completion M" for M proof show "B \ A \ A \ null_sets (completion M) \ B \ sets (completion M)" for B A using null_sets_completion_subset[of B A M] by (simp add: null_sets_def) qed lemma null_sets_restrict_space: "\ \ sets M \ A \ null_sets (restrict_space M \) \ A \ \ \ A \ null_sets M" by (auto simp: null_sets_def emeasure_restrict_space sets_restrict_space) lemma completion_ex_borel_measurable_real: fixes g :: "'a \ real" assumes g: "g \ borel_measurable (completion M)" shows "\g'\borel_measurable M. (AE x in M. g x = g' x)" proof - have "(\x. ennreal (g x)) \ completion M \\<^sub>M borel" "(\x. ennreal (- g x)) \ completion M \\<^sub>M borel" using g by auto from this[THEN completion_ex_borel_measurable] obtain pf nf :: "'a \ ennreal" where [measurable]: "nf \ M \\<^sub>M borel" "pf \ M \\<^sub>M borel" and ae: "AE x in M. pf x = ennreal (g x)" "AE x in M. nf x = ennreal (- g x)" by (auto simp: eq_commute) then have "AE x in M. pf x = ennreal (g x) \ nf x = ennreal (- g x)" by auto then obtain N where "N \ null_sets M" "{x\space M. pf x \ ennreal (g x) \ nf x \ ennreal (- g x)} \ N" by (auto elim!: AE_E) show ?thesis proof let ?F = "\x. indicator (space M - N) x * (enn2real (pf x) - enn2real (nf x))" show "?F \ M \\<^sub>M borel" using \N \ null_sets M\ by auto show "AE x in M. g x = ?F x" using \N \ null_sets M\[THEN AE_not_in] ae AE_space apply eventually_elim subgoal for x by (cases "0::real" "g x" rule: linorder_le_cases) (auto simp: ennreal_neg) done qed qed lemma simple_function_completion: "simple_function M f \ simple_function (completion M) f" by (simp add: simple_function_def) lemma simple_integral_completion: "simple_function M f \ simple_integral (completion M) f = simple_integral M f" unfolding simple_integral_def by simp lemma nn_integral_completion: "nn_integral (completion M) f = nn_integral M f" unfolding nn_integral_def proof (safe intro!: SUP_eq) fix s assume s: "simple_function (completion M) s" and "s \ f" then obtain s' where s': "simple_function M s'" "AE x in M. s x = s' x" by (auto dest: completion_ex_simple_function) then obtain N where N: "N \ null_sets M" "{x\space M. s x \ s' x} \ N" by (auto elim!: AE_E) then have ae_N: "AE x in M. (s x \ s' x \ x \ N) \ x \ N" by (auto dest: AE_not_in) define s'' where "s'' x = (if x \ N then 0 else s x)" for x then have ae_s_eq_s'': "AE x in completion M. s x = s'' x" using s' ae_N by (intro AE_completion) auto have s'': "simple_function M s''" proof (subst simple_function_cong) show "t \ space M \ s'' t = (if t \ N then 0 else s' t)" for t using N by (auto simp: s''_def dest: sets.sets_into_space) show "simple_function M (\t. if t \ N then 0 else s' t)" unfolding s''_def[abs_def] using N by (auto intro!: simple_function_If s') qed show "\j\{g. simple_function M g \ g \ f}. integral\<^sup>S (completion M) s \ integral\<^sup>S M j" proof (safe intro!: bexI[of _ s'']) have "integral\<^sup>S (completion M) s = integral\<^sup>S (completion M) s''" by (intro simple_integral_cong_AE s simple_function_completion s'' ae_s_eq_s'') then show "integral\<^sup>S (completion M) s \ integral\<^sup>S M s''" using s'' by (simp add: simple_integral_completion) from \s \ f\ show "s'' \ f" unfolding s''_def le_fun_def by auto qed fact next fix s assume "simple_function M s" "s \ f" then show "\j\{g. simple_function (completion M) g \ g \ f}. integral\<^sup>S M s \ integral\<^sup>S (completion M) j" by (intro bexI[of _ s]) (auto simp: simple_integral_completion simple_function_completion) qed lemma integrable_completion: fixes f :: "'a \ 'b::{banach, second_countable_topology}" shows "f \ M \\<^sub>M borel \ integrable (completion M) f \ integrable M f" by (rule integrable_subalgebra[symmetric]) auto lemma integral_completion: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f: "f \ M \\<^sub>M borel" shows "integral\<^sup>L (completion M) f = integral\<^sup>L M f" by (rule integral_subalgebra[symmetric]) (auto intro: f) locale\<^marker>\tag important\ semifinite_measure = fixes M :: "'a measure" assumes semifinite: "\A. A \ sets M \ emeasure M A = \ \ \B\sets M. B \ A \ emeasure M B < \" locale\<^marker>\tag important\ locally_determined_measure = semifinite_measure + assumes locally_determined: "\A. A \ space M \ (\B. B \ sets M \ emeasure M B < \ \ A \ B \ sets M) \ A \ sets M" locale\<^marker>\tag important\ cld_measure = complete_measure M + locally_determined_measure M for M :: "'a measure" definition\<^marker>\tag important\ outer_measure_of :: "'a measure \ 'a set \ ennreal" where "outer_measure_of M A = (INF B \ {B\sets M. A \ B}. emeasure M B)" lemma outer_measure_of_eq[simp]: "A \ sets M \ outer_measure_of M A = emeasure M A" by (auto simp: outer_measure_of_def intro!: INF_eqI emeasure_mono) lemma outer_measure_of_mono: "A \ B \ outer_measure_of M A \ outer_measure_of M B" unfolding outer_measure_of_def by (intro INF_superset_mono) auto lemma outer_measure_of_attain: assumes "A \ space M" shows "\E\sets M. A \ E \ outer_measure_of M A = emeasure M E" proof - have "emeasure M ` {B \ sets M. A \ B} \ {}" using \A \ space M\ by auto from ennreal_Inf_countable_INF[OF this] obtain f where f: "range f \ emeasure M ` {B \ sets M. A \ B}" "decseq f" and "outer_measure_of M A = (INF i. f i)" unfolding outer_measure_of_def by auto have "\E. \n. (E n \ sets M \ A \ E n \ emeasure M (E n) \ f n) \ E (Suc n) \ E n" proof (rule dependent_nat_choice) show "\x. x \ sets M \ A \ x \ emeasure M x \ f 0" using f(1) by (fastforce simp: image_subset_iff image_iff intro: eq_refl[OF sym]) next fix E n assume "E \ sets M \ A \ E \ emeasure M E \ f n" moreover obtain F where "F \ sets M" "A \ F" "f (Suc n) = emeasure M F" using f(1) by (auto simp: image_subset_iff image_iff) ultimately show "\y. (y \ sets M \ A \ y \ emeasure M y \ f (Suc n)) \ y \ E" by (auto intro!: exI[of _ "F \ E"] emeasure_mono) qed then obtain E where [simp]: "\n. E n \ sets M" and "\n. A \ E n" and le_f: "\n. emeasure M (E n) \ f n" and "decseq E" by (auto simp: decseq_Suc_iff) show ?thesis proof cases assume fin: "\i. emeasure M (E i) < \" show ?thesis proof (intro bexI[of _ "\i. E i"] conjI) show "A \ (\i. E i)" "(\i. E i) \ sets M" using \\n. A \ E n\ by auto have " (INF i. emeasure M (E i)) \ outer_measure_of M A" unfolding \outer_measure_of M A = (INF n. f n)\ by (intro INF_superset_mono le_f) auto moreover have "outer_measure_of M A \ (INF i. outer_measure_of M (E i))" by (intro INF_greatest outer_measure_of_mono \\n. A \ E n\) ultimately have "outer_measure_of M A = (INF i. emeasure M (E i))" by auto also have "\ = emeasure M (\i. E i)" using fin by (intro INF_emeasure_decseq' \decseq E\) (auto simp: less_top) finally show "outer_measure_of M A = emeasure M (\i. E i)" . qed next assume "\i. emeasure M (E i) < \" then have "f n = \" for n using le_f by (auto simp: not_less top_unique) moreover have "\E\sets M. A \ E \ f 0 = emeasure M E" using f by auto ultimately show ?thesis unfolding \outer_measure_of M A = (INF n. f n)\ by simp qed qed lemma SUP_outer_measure_of_incseq: assumes A: "\n. A n \ space M" and "incseq A" shows "(SUP n. outer_measure_of M (A n)) = outer_measure_of M (\i. A i)" proof (rule antisym) obtain E where E: "\n. E n \ sets M" "\n. A n \ E n" "\n. outer_measure_of M (A n) = emeasure M (E n)" using outer_measure_of_attain[OF A] by metis define F where "F n = (\i\{n ..}. E i)" for n with E have F: "incseq F" "\n. F n \ sets M" by (auto simp: incseq_def) have "A n \ F n" for n using incseqD[OF \incseq A\, of n] \\n. A n \ E n\ by (auto simp: F_def) have eq: "outer_measure_of M (A n) = outer_measure_of M (F n)" for n proof (intro antisym) have "outer_measure_of M (F n) \ outer_measure_of M (E n)" by (intro outer_measure_of_mono) (auto simp add: F_def) with E show "outer_measure_of M (F n) \ outer_measure_of M (A n)" by auto show "outer_measure_of M (A n) \ outer_measure_of M (F n)" by (intro outer_measure_of_mono \A n \ F n\) qed have "outer_measure_of M (\n. A n) \ outer_measure_of M (\n. F n)" using \\n. A n \ F n\ by (intro outer_measure_of_mono) auto also have "\ = (SUP n. emeasure M (F n))" using F by (simp add: SUP_emeasure_incseq subset_eq) finally show "outer_measure_of M (\n. A n) \ (SUP n. outer_measure_of M (A n))" by (simp add: eq F) qed (auto intro: SUP_least outer_measure_of_mono) definition\<^marker>\tag important\ measurable_envelope :: "'a measure \ 'a set \ 'a set \ bool" where "measurable_envelope M A E \ (A \ E \ E \ sets M \ (\F\sets M. emeasure M (F \ E) = outer_measure_of M (F \ A)))" lemma measurable_envelopeD: assumes "measurable_envelope M A E" shows "A \ E" and "E \ sets M" and "\F. F \ sets M \ emeasure M (F \ E) = outer_measure_of M (F \ A)" and "A \ space M" using assms sets.sets_into_space[of E] by (auto simp: measurable_envelope_def) lemma measurable_envelopeD1: assumes E: "measurable_envelope M A E" and F: "F \ sets M" "F \ E - A" shows "emeasure M F = 0" proof - have "emeasure M F = emeasure M (F \ E)" using F by (intro arg_cong2[where f=emeasure]) auto also have "\ = outer_measure_of M (F \ A)" using measurable_envelopeD[OF E] \F \ sets M\ by (auto simp: measurable_envelope_def) also have "\ = outer_measure_of M {}" using \F \ E - A\ by (intro arg_cong2[where f=outer_measure_of]) auto finally show "emeasure M F = 0" by simp qed lemma measurable_envelope_eq1: assumes "A \ E" "E \ sets M" shows "measurable_envelope M A E \ (\F\sets M. F \ E - A \ emeasure M F = 0)" proof safe assume *: "\F\sets M. F \ E - A \ emeasure M F = 0" show "measurable_envelope M A E" unfolding measurable_envelope_def proof (rule ccontr, auto simp add: \E \ sets M\ \A \ E\) fix F assume "F \ sets M" "emeasure M (F \ E) \ outer_measure_of M (F \ A)" then have "outer_measure_of M (F \ A) < emeasure M (F \ E)" using outer_measure_of_mono[of "F \ A" "F \ E" M] \A \ E\ \E \ sets M\ by (auto simp: less_le) then obtain G where G: "G \ sets M" "F \ A \ G" and less: "emeasure M G < emeasure M (E \ F)" unfolding outer_measure_of_def INF_less_iff by (auto simp: ac_simps) have le: "emeasure M (G \ E \ F) \ emeasure M G" using \E \ sets M\ \G \ sets M\ \F \ sets M\ by (auto intro!: emeasure_mono) from G have "E \ F - G \ sets M" "E \ F - G \ E - A" using \F \ sets M\ \E \ sets M\ by auto with * have "0 = emeasure M (E \ F - G)" by auto also have "E \ F - G = E \ F - (G \ E \ F)" by auto also have "emeasure M (E \ F - (G \ E \ F)) = emeasure M (E \ F) - emeasure M (G \ E \ F)" using \E \ sets M\ \F \ sets M\ le less G by (intro emeasure_Diff) (auto simp: top_unique) also have "\ > 0" using le less by (intro diff_gr0_ennreal) auto finally show False by auto qed qed (rule measurable_envelopeD1) lemma measurable_envelopeD2: assumes E: "measurable_envelope M A E" shows "emeasure M E = outer_measure_of M A" proof - from \measurable_envelope M A E\ have "emeasure M (E \ E) = outer_measure_of M (E \ A)" by (auto simp: measurable_envelope_def) with measurable_envelopeD[OF E] show "emeasure M E = outer_measure_of M A" by (auto simp: Int_absorb1) qed lemma\<^marker>\tag important\ measurable_envelope_eq2: assumes "A \ E" "E \ sets M" "emeasure M E < \" shows "measurable_envelope M A E \ (emeasure M E = outer_measure_of M A)" proof safe assume *: "emeasure M E = outer_measure_of M A" show "measurable_envelope M A E" unfolding measurable_envelope_eq1[OF \A \ E\ \E \ sets M\] proof (intro conjI ballI impI assms) fix F assume F: "F \ sets M" "F \ E - A" with \E \ sets M\ have le: "emeasure M F \ emeasure M E" by (intro emeasure_mono) auto from F \A \ E\ have "outer_measure_of M A \ outer_measure_of M (E - F)" by (intro outer_measure_of_mono) auto then have "emeasure M E - 0 \ emeasure M (E - F)" using * \E \ sets M\ \F \ sets M\ by simp also have "\ = emeasure M E - emeasure M F" using \E \ sets M\ \emeasure M E < \\ F le by (intro emeasure_Diff) (auto simp: top_unique) finally show "emeasure M F = 0" using ennreal_mono_minus_cancel[of "emeasure M E" 0 "emeasure M F"] le assms by auto qed qed (auto intro: measurable_envelopeD2) lemma measurable_envelopeI_countable: fixes A :: "nat \ 'a set" assumes E: "\n. measurable_envelope M (A n) (E n)" shows "measurable_envelope M (\n. A n) (\n. E n)" proof (subst measurable_envelope_eq1) show "(\n. A n) \ (\n. E n)" "(\n. E n) \ sets M" using measurable_envelopeD(1,2)[OF E] by auto show "\F\sets M. F \ (\n. E n) - (\n. A n) \ emeasure M F = 0" proof safe fix F assume F: "F \ sets M" "F \ (\n. E n) - (\n. A n)" then have "F \ E n \ sets M" "F \ E n \ E n - A n" "F \ (\n. E n)" for n using measurable_envelopeD(1,2)[OF E] by auto then have "emeasure M (\n. F \ E n) = 0" by (intro emeasure_UN_eq_0 measurable_envelopeD1[OF E]) auto then show "emeasure M F = 0" using \F \ (\n. E n)\ by (auto simp: Int_absorb2) qed qed lemma measurable_envelopeI_countable_cover: fixes A and C :: "nat \ 'a set" assumes C: "A \ (\n. C n)" "\n. C n \ sets M" "\n. emeasure M (C n) < \" shows "\E\(\n. C n). measurable_envelope M A E" proof - have "A \ C n \ space M" for n using \C n \ sets M\ by (auto dest: sets.sets_into_space) then have "\n. \E\sets M. A \ C n \ E \ outer_measure_of M (A \ C n) = emeasure M E" using outer_measure_of_attain[of "A \ C n" M for n] by auto then obtain E where E: "\n. E n \ sets M" "\n. A \ C n \ E n" and eq: "\n. outer_measure_of M (A \ C n) = emeasure M (E n)" by metis have "outer_measure_of M (A \ C n) \ outer_measure_of M (E n \ C n)" for n using E by (intro outer_measure_of_mono) auto moreover have "outer_measure_of M (E n \ C n) \ outer_measure_of M (E n)" for n by (intro outer_measure_of_mono) auto ultimately have eq: "outer_measure_of M (A \ C n) = emeasure M (E n \ C n)" for n using E C by (intro antisym) (auto simp: eq) { fix n have "outer_measure_of M (A \ C n) \ outer_measure_of M (C n)" by (intro outer_measure_of_mono) simp also have "\ < \" using assms by auto finally have "emeasure M (E n \ C n) < \" using eq by simp } then have "measurable_envelope M (\n. A \ C n) (\n. E n \ C n)" using E C by (intro measurable_envelopeI_countable measurable_envelope_eq2[THEN iffD2]) (auto simp: eq) with \A \ (\n. C n)\ show ?thesis by (intro exI[of _ "(\n. E n \ C n)"]) (auto simp add: Int_absorb2) qed lemma (in complete_measure) complete_sets_sandwich: assumes [measurable]: "A \ sets M" "C \ sets M" and subset: "A \ B" "B \ C" and measure: "emeasure M A = emeasure M C" "emeasure M A < \" shows "B \ sets M" proof - have "B - A \ sets M" proof (rule complete) show "B - A \ C - A" using subset by auto show "C - A \ null_sets M" using measure subset by(simp add: emeasure_Diff null_setsI) qed then have "A \ (B - A) \ sets M" by measurable also have "A \ (B - A) = B" using \A \ B\ by auto finally show ?thesis . qed lemma (in complete_measure) complete_sets_sandwich_fmeasurable: assumes [measurable]: "A \ fmeasurable M" "C \ fmeasurable M" and subset: "A \ B" "B \ C" and measure: "measure M A = measure M C" shows "B \ fmeasurable M" proof (rule fmeasurableI2) show "B \ C" "C \ fmeasurable M" by fact+ show "B \ sets M" proof (rule complete_sets_sandwich) show "A \ sets M" "C \ sets M" "A \ B" "B \ C" using assms by auto show "emeasure M A < \" using \A \ fmeasurable M\ by (auto simp: fmeasurable_def) show "emeasure M A = emeasure M C" using assms by (simp add: emeasure_eq_measure2) qed qed lemma AE_completion_iff: "(AE x in completion M. P x) \ (AE x in M. P x)" proof assume "AE x in completion M. P x" then obtain N where "N \ null_sets (completion M)" and P: "{x\space M. \ P x} \ N" unfolding eventually_ae_filter by auto then obtain N' where N': "N' \ null_sets M" and "N \ N'" unfolding null_sets_completion_iff2 by auto from P \N \ N'\ have "{x\space M. \ P x} \ N'" by auto with N' show "AE x in M. P x" unfolding eventually_ae_filter by auto qed (rule AE_completion) lemma null_part_null_sets: "S \ completion M \ null_part M S \ null_sets (completion M)" by (auto dest!: null_part intro: null_sets_completionI null_sets_completion_subset) lemma AE_notin_null_part: "S \ completion M \ (AE x in M. x \ null_part M S)" by (auto dest!: null_part_null_sets AE_not_in simp: AE_completion_iff) lemma completion_upper: assumes A: "A \ sets (completion M)" shows "\A'\sets M. A \ A' \ emeasure (completion M) A = emeasure M A'" proof - from AE_notin_null_part[OF A] obtain N where N: "N \ null_sets M" "null_part M A \ N" unfolding eventually_ae_filter using null_part_null_sets[OF A, THEN null_setsD2, THEN sets.sets_into_space] by auto show ?thesis proof (intro bexI conjI) show "A \ main_part M A \ N" using \null_part M A \ N\ by (subst main_part_null_part_Un[symmetric, OF A]) auto show "emeasure (completion M) A = emeasure M (main_part M A \ N)" using A \N \ null_sets M\ by (simp add: emeasure_Un_null_set) qed (use A N in auto) qed lemma AE_in_main_part: assumes A: "A \ completion M" shows "AE x in M. x \ main_part M A \ x \ A" using AE_notin_null_part[OF A] by (subst (2) main_part_null_part_Un[symmetric, OF A]) auto lemma completion_density_eq: assumes ae: "AE x in M. 0 < f x" and [measurable]: "f \ M \\<^sub>M borel" shows "completion (density M f) = density (completion M) f" proof (intro measure_eqI) have "N' \ sets M \ (AE x\N' in M. f x = 0) \ N' \ null_sets M" for N' proof safe assume N': "N' \ sets M" and ae_N': "AE x\N' in M. f x = 0" from ae_N' ae have "AE x in M. x \ N'" by eventually_elim auto then show "N' \ null_sets M" using N' by (simp add: AE_iff_null_sets) next assume N': "N' \ null_sets M" then show "N' \ sets M" "AE x\N' in M. f x = 0" using ae AE_not_in[OF N'] by (auto simp: less_le) qed then show sets_eq: "sets (completion (density M f)) = sets (density (completion M) f)" by (simp add: sets_completion null_sets_density_iff) fix A assume A: \A \ completion (density M f)\ moreover have "A \ completion M" using A unfolding sets_eq by simp moreover have "main_part (density M f) A \ M" using A main_part_sets[of A "density M f"] unfolding sets_density sets_eq by simp moreover have "AE x in M. x \ main_part (density M f) A \ x \ A" using AE_in_main_part[OF \A \ completion (density M f)\] ae by (auto simp add: AE_density) ultimately show "emeasure (completion (density M f)) A = emeasure (density (completion M) f) A" by (auto simp add: emeasure_density measurable_completion nn_integral_completion intro!: nn_integral_cong_AE) qed lemma null_sets_subset: "B \ null_sets M \ A \ sets M \ A \ B \ A \ null_sets M" using emeasure_mono[of A B M] by (simp add: null_sets_def) lemma (in complete_measure) complete2: "A \ B \ B \ null_sets M \ A \ null_sets M" using complete[of A B] null_sets_subset[of B M A] by simp lemma (in complete_measure) AE_iff_null_sets: "(AE x in M. P x) \ {x\space M. \ P x} \ null_sets M" unfolding eventually_ae_filter by (auto intro: complete2) lemma (in complete_measure) null_sets_iff_AE: "A \ null_sets M \ ((AE x in M. x \ A) \ A \ space M)" unfolding AE_iff_null_sets by (auto cong: rev_conj_cong dest: sets.sets_into_space simp: subset_eq) lemma (in complete_measure) in_sets_AE: assumes ae: "AE x in M. x \ A \ x \ B" and A: "A \ sets M" and B: "B \ space M" shows "B \ sets M" proof - have "(AE x in M. x \ B - A \ x \ A - B)" using ae by eventually_elim auto then have "B - A \ null_sets M" "A - B \ null_sets M" using A B unfolding null_sets_iff_AE by (auto dest: sets.sets_into_space) then have "A \ (B - A) - (A - B) \ sets M" using A by blast also have "A \ (B - A) - (A - B) = B" by auto finally show "B \ sets M" . qed lemma (in complete_measure) vimage_null_part_null_sets: assumes f: "f \ M \\<^sub>M N" and eq: "null_sets N \ null_sets (distr M N f)" and A: "A \ completion N" shows "f -` null_part N A \ space M \ null_sets M" proof - obtain N' where "N' \ null_sets N" "null_part N A \ N'" using null_part[OF A] by auto then have N': "N' \ null_sets (distr M N f)" using eq by auto show ?thesis proof (rule complete2) show "f -` null_part N A \ space M \ f -` N' \ space M" using \null_part N A \ N'\ by auto show "f -` N' \ space M \ null_sets M" using f N' by (auto simp: null_sets_def emeasure_distr) qed qed lemma (in complete_measure) vimage_null_part_sets: "f \ M \\<^sub>M N \ null_sets N \ null_sets (distr M N f) \ A \ completion N \ f -` null_part N A \ space M \ sets M" using vimage_null_part_null_sets[of f N A] by auto lemma (in complete_measure) measurable_completion2: assumes f: "f \ M \\<^sub>M N" and null_sets: "null_sets N \ null_sets (distr M N f)" shows "f \ M \\<^sub>M completion N" proof (rule measurableI) show "x \ space M \ f x \ space (completion N)" for x using f[THEN measurable_space] by auto fix A assume A: "A \ sets (completion N)" have "f -` A \ space M = (f -` main_part N A \ space M) \ (f -` null_part N A \ space M)" using main_part_null_part_Un[OF A] by auto then show "f -` A \ space M \ sets M" using f A null_sets by (auto intro: vimage_null_part_sets measurable_sets) qed lemma (in complete_measure) completion_distr_eq: assumes X: "X \ M \\<^sub>M N" and null_sets: "null_sets (distr M N X) = null_sets N" shows "completion (distr M N X) = distr M (completion N) X" proof (rule measure_eqI) show eq: "sets (completion (distr M N X)) = sets (distr M (completion N) X)" by (simp add: sets_completion null_sets) fix A assume A: "A \ completion (distr M N X)" moreover have A': "A \ completion N" using A by (simp add: eq) moreover have "main_part (distr M N X) A \ sets N" using main_part_sets[OF A] by simp moreover have "X -` A \ space M = (X -` main_part (distr M N X) A \ space M) \ (X -` null_part (distr M N X) A \ space M)" using main_part_null_part_Un[OF A] by auto moreover have "X -` null_part (distr M N X) A \ space M \ null_sets M" using X A by (intro vimage_null_part_null_sets) (auto cong: distr_cong) ultimately show "emeasure (completion (distr M N X)) A = emeasure (distr M (completion N) X) A" using X by (auto simp: emeasure_distr measurable_completion null_sets measurable_completion2 intro!: emeasure_Un_null_set[symmetric]) qed lemma distr_completion: "X \ M \\<^sub>M N \ distr (completion M) N X = distr M N X" by (intro measure_eqI) (auto simp: emeasure_distr measurable_completion) proposition (in complete_measure) fmeasurable_inner_outer: "S \ fmeasurable M \ (\e>0. \T\fmeasurable M. \U\fmeasurable M. T \ S \ S \ U \ \measure M T - measure M U\ < e)" (is "_ \ ?approx") proof safe let ?\ = "measure M" let ?D = "\T U . \?\ T - ?\ U\" assume ?approx have "\A. \n. (fst (A n) \ fmeasurable M \ snd (A n) \ fmeasurable M \ fst (A n) \ S \ S \ snd (A n) \ ?D (fst (A n)) (snd (A n)) < 1/Suc n) \ (fst (A n) \ fst (A (Suc n)) \ snd (A (Suc n)) \ snd (A n))" (is "\A. \n. ?P n (A n) \ ?Q (A n) (A (Suc n))") proof (intro dependent_nat_choice) show "\A. ?P 0 A" using \?approx\[THEN spec, of 1] by auto next fix A n assume "?P n A" moreover from \?approx\[THEN spec, of "1/Suc (Suc n)"] obtain T U where *: "T \ fmeasurable M" "U \ fmeasurable M" "T \ S" "S \ U" "?D T U < 1 / Suc (Suc n)" by auto ultimately have "?\ T \ ?\ (T \ fst A)" "?\ (U \ snd A) \ ?\ U" "?\ T \ ?\ U" "?\ (T \ fst A) \ ?\ (U \ snd A)" by (auto intro!: measure_mono_fmeasurable) then have "?D (T \ fst A) (U \ snd A) \ ?D T U" by auto also have "?D T U < 1/Suc (Suc n)" by fact finally show "\B. ?P (Suc n) B \ ?Q A B" using \?P n A\ * by (intro exI[of _ "(T \ fst A, U \ snd A)"] conjI) auto qed then obtain A where lm: "\n. fst (A n) \ fmeasurable M" "\n. snd (A n) \ fmeasurable M" and set_bound: "\n. fst (A n) \ S" "\n. S \ snd (A n)" and mono: "\n. fst (A n) \ fst (A (Suc n))" "\n. snd (A (Suc n)) \ snd (A n)" and bound: "\n. ?D (fst (A n)) (snd (A n)) < 1/Suc n" by metis have INT_sA: "(\n. snd (A n)) \ fmeasurable M" using lm by (intro fmeasurable_INT[of _ 0]) auto have UN_fA: "(\n. fst (A n)) \ fmeasurable M" using lm order_trans[OF set_bound(1) set_bound(2)[of 0]] by (intro fmeasurable_UN[of _ _ "snd (A 0)"]) auto have "(\n. ?\ (fst (A n)) - ?\ (snd (A n))) \ 0" using bound by (subst tendsto_rabs_zero_iff[symmetric]) (intro tendsto_sandwich[OF _ _ tendsto_const LIMSEQ_inverse_real_of_nat]; auto intro!: always_eventually less_imp_le simp: divide_inverse) moreover have "(\n. ?\ (fst (A n)) - ?\ (snd (A n))) \ ?\ (\n. fst (A n)) - ?\ (\n. snd (A n))" proof (intro tendsto_diff Lim_measure_incseq Lim_measure_decseq) show "range (\i. fst (A i)) \ sets M" "range (\i. snd (A i)) \ sets M" "incseq (\i. fst (A i))" "decseq (\n. snd (A n))" using mono lm by (auto simp: incseq_Suc_iff decseq_Suc_iff intro!: measure_mono_fmeasurable) show "emeasure M (\x. fst (A x)) \ \" "emeasure M (snd (A n)) \ \" for n using lm(2)[of n] UN_fA by (auto simp: fmeasurable_def) qed ultimately have eq: "0 = ?\ (\n. fst (A n)) - ?\ (\n. snd (A n))" by (rule LIMSEQ_unique) show "S \ fmeasurable M" using UN_fA INT_sA proof (rule complete_sets_sandwich_fmeasurable) show "(\n. fst (A n)) \ S" "S \ (\n. snd (A n))" using set_bound by auto show "?\ (\n. fst (A n)) = ?\ (\n. snd (A n))" using eq by auto qed qed (auto intro!: bexI[of _ S]) lemma (in complete_measure) fmeasurable_measure_inner_outer: "(S \ fmeasurable M \ m = measure M S) \ (\e>0. \T\fmeasurable M. T \ S \ m - e < measure M T) \ (\e>0. \U\fmeasurable M. S \ U \ measure M U < m + e)" (is "?lhs = ?rhs") proof assume RHS: ?rhs then have T: "\e. 0 < e \ (\T\fmeasurable M. T \ S \ m - e < measure M T)" and U: "\e. 0 < e \ (\U\fmeasurable M. S \ U \ measure M U < m + e)" by auto have "S \ fmeasurable M" proof (subst fmeasurable_inner_outer, safe) fix e::real assume "0 < e" with RHS obtain T U where T: "T \ fmeasurable M" "T \ S" "m - e/2 < measure M T" and U: "U \ fmeasurable M" "S \ U" "measure M U < m + e/2" by (meson half_gt_zero)+ moreover have "measure M U - measure M T < (m + e/2) - (m - e/2)" by (intro diff_strict_mono) fact+ moreover have "measure M T \ measure M U" using T U by (intro measure_mono_fmeasurable) auto ultimately show "\T\fmeasurable M. \U\fmeasurable M. T \ S \ S \ U \ \measure M T - measure M U\ < e" apply (rule_tac bexI[OF _ \T \ fmeasurable M\]) apply (rule_tac bexI[OF _ \U \ fmeasurable M\]) by auto qed moreover have "m = measure M S" using \S \ fmeasurable M\ U[of "measure M S - m"] T[of "m - measure M S"] by (cases m "measure M S" rule: linorder_cases) (auto simp: not_le[symmetric] measure_mono_fmeasurable) ultimately show ?lhs by simp qed (auto intro!: bexI[of _ S]) lemma (in complete_measure) null_sets_outer: "S \ null_sets M \ (\e>0. \T\fmeasurable M. S \ T \ measure M T < e)" proof - have "S \ null_sets M \ (S \ fmeasurable M \ 0 = measure M S)" by (auto simp: null_sets_def emeasure_eq_measure2 intro: fmeasurableI) (simp add: measure_def) also have "\ = (\e>0. \T\fmeasurable M. S \ T \ measure M T < e)" unfolding fmeasurable_measure_inner_outer by auto finally show ?thesis . qed lemma (in complete_measure) fmeasurable_measure_inner_outer_le: "(S \ fmeasurable M \ m = measure M S) \ (\e>0. \T\fmeasurable M. T \ S \ m - e \ measure M T) \ (\e>0. \U\fmeasurable M. S \ U \ measure M U \ m + e)" (is ?T1) and null_sets_outer_le: "S \ null_sets M \ (\e>0. \T\fmeasurable M. S \ T \ measure M T \ e)" (is ?T2) proof - have "e > 0 \ m - e/2 \ t \ m - e < t" "e > 0 \ t \ m + e/2 \ t < m + e" "e > 0 \ e/2 > 0" for e t by auto then show ?T1 ?T2 unfolding fmeasurable_measure_inner_outer null_sets_outer by (meson dense le_less_trans less_imp_le)+ qed lemma (in cld_measure) notin_sets_outer_measure_of_cover: assumes E: "E \ space M" "E \ sets M" shows "\B\sets M. 0 < emeasure M B \ emeasure M B < \ \ outer_measure_of M (B \ E) = emeasure M B \ outer_measure_of M (B - E) = emeasure M B" proof - from locally_determined[OF \E \ space M\] \E \ sets M\ obtain F where [measurable]: "F \ sets M" and "emeasure M F < \" "E \ F \ sets M" by blast then obtain H H' where H: "measurable_envelope M (F \ E) H" and H': "measurable_envelope M (F - E) H'" using measurable_envelopeI_countable_cover[of "F \ E" "\_. F" M] measurable_envelopeI_countable_cover[of "F - E" "\_. F" M] by auto note measurable_envelopeD(2)[OF H', measurable] measurable_envelopeD(2)[OF H, measurable] from measurable_envelopeD(1)[OF H'] measurable_envelopeD(1)[OF H] have subset: "F - H' \ F \ E" "F \ E \ F \ H" by auto moreover define G where "G = (F \ H) - (F - H')" ultimately have G: "G = F \ H \ H'" by auto have "emeasure M (F \ H) \ 0" proof assume "emeasure M (F \ H) = 0" then have "F \ H \ null_sets M" by auto with \E \ F \ sets M\ show False using complete[OF \F \ E \ F \ H\] by (auto simp: Int_commute) qed moreover have "emeasure M (F - H') \ emeasure M (F \ H)" proof assume "emeasure M (F - H') = emeasure M (F \ H)" with \E \ F \ sets M\ emeasure_mono[of "F \ H" F M] \emeasure M F < \\ have "F \ E \ sets M" by (intro complete_sets_sandwich[OF _ _ subset]) auto with \E \ F \ sets M\ show False by (simp add: Int_commute) qed moreover have "emeasure M (F - H') \ emeasure M (F \ H)" using subset by (intro emeasure_mono) auto ultimately have "emeasure M G \ 0" unfolding G_def using subset by (subst emeasure_Diff) (auto simp: top_unique diff_eq_0_iff_ennreal) show ?thesis proof (intro bexI conjI) have "emeasure M G \ emeasure M F" unfolding G by (auto intro!: emeasure_mono) with \emeasure M F < \\ show "0 < emeasure M G" "emeasure M G < \" using \emeasure M G \ 0\ by (auto simp: zero_less_iff_neq_zero) show [measurable]: "G \ sets M" unfolding G by auto have "emeasure M G = outer_measure_of M (F \ H' \ (F \ E))" using measurable_envelopeD(3)[OF H, of "F \ H'"] unfolding G by (simp add: ac_simps) also have "\ \ outer_measure_of M (G \ E)" using measurable_envelopeD(1)[OF H] by (intro outer_measure_of_mono) (auto simp: G) finally show "outer_measure_of M (G \ E) = emeasure M G" using outer_measure_of_mono[of "G \ E" G M] by auto have "emeasure M G = outer_measure_of M (F \ H \ (F - E))" using measurable_envelopeD(3)[OF H', of "F \ H"] unfolding G by (simp add: ac_simps) also have "\ \ outer_measure_of M (G - E)" using measurable_envelopeD(1)[OF H'] by (intro outer_measure_of_mono) (auto simp: G) finally show "outer_measure_of M (G - E) = emeasure M G" using outer_measure_of_mono[of "G - E" G M] by auto qed qed text \The following theorem is a specialization of D.H. Fremlin, Measure Theory vol 4I (413G). We only show one direction and do not use a inner regular family \K\.\ lemma (in cld_measure) borel_measurable_cld: fixes f :: "'a \ real" assumes "\A a b. A \ sets M \ 0 < emeasure M A \ emeasure M A < \ \ a < b \ min (outer_measure_of M {x\A. f x \ a}) (outer_measure_of M {x\A. b \ f x}) < emeasure M A" shows "f \ M \\<^sub>M borel" proof (rule ccontr) let ?E = "\a. {x\space M. f x \ a}" and ?F = "\a. {x\space M. a \ f x}" assume "f \ M \\<^sub>M borel" then obtain a where "?E a \ sets M" unfolding borel_measurable_iff_le by blast from notin_sets_outer_measure_of_cover[OF _ this] obtain K where K: "K \ sets M" "0 < emeasure M K" "emeasure M K < \" and eq1: "outer_measure_of M (K \ ?E a) = emeasure M K" and eq2: "outer_measure_of M (K - ?E a) = emeasure M K" by auto then have me_K: "measurable_envelope M (K \ ?E a) K" by (subst measurable_envelope_eq2) auto define b where "b n = a + inverse (real (Suc n))" for n have "(SUP n. outer_measure_of M (K \ ?F (b n))) = outer_measure_of M (\n. K \ ?F (b n))" proof (intro SUP_outer_measure_of_incseq) have "x \ y \ b y \ b x" for x y by (auto simp: b_def field_simps) then show "incseq (\n. K \ {x \ space M. b n \ f x})" by (auto simp: incseq_def intro: order_trans) qed auto also have "(\n. K \ ?F (b n)) = K - ?E a" proof - have "b \ a" unfolding b_def by (rule LIMSEQ_inverse_real_of_nat_add) then have "\n. \ b n \ f x \ f x \ a" for x by (rule LIMSEQ_le_const) (auto intro: less_imp_le simp: not_le) moreover have "\ b n \ a" for n by (auto simp: b_def) ultimately show ?thesis using \K \ sets M\[THEN sets.sets_into_space] by (auto simp: subset_eq intro: order_trans) qed finally have "0 < (SUP n. outer_measure_of M (K \ ?F (b n)))" using K by (simp add: eq2) then obtain n where pos_b: "0 < outer_measure_of M (K \ ?F (b n))" and "a < b n" unfolding less_SUP_iff by (auto simp: b_def) from measurable_envelopeI_countable_cover[of "K \ ?F (b n)" "\_. K" M] K obtain K' where "K' \ K" and me_K': "measurable_envelope M (K \ ?F (b n)) K'" by auto then have K'_le_K: "emeasure M K' \ emeasure M K" by (intro emeasure_mono K) have "K' \ sets M" using me_K' by (rule measurable_envelopeD) have "min (outer_measure_of M {x\K'. f x \ a}) (outer_measure_of M {x\K'. b n \ f x}) < emeasure M K'" proof (rule assms) show "0 < emeasure M K'" "emeasure M K' < \" using measurable_envelopeD2[OF me_K'] pos_b K K'_le_K by auto qed fact+ also have "{x\K'. f x \ a} = K' \ (K \ ?E a)" using \K' \ sets M\[THEN sets.sets_into_space] \K' \ K\ by auto also have "{x\K'. b n \ f x} = K' \ (K \ ?F (b n))" using \K' \ sets M\[THEN sets.sets_into_space] \K' \ K\ by auto finally have "min (emeasure M K) (emeasure M K') < emeasure M K'" unfolding measurable_envelopeD(3)[OF me_K \K' \ sets M\, symmetric] measurable_envelopeD(3)[OF me_K' \K' \ sets M\, symmetric] using \K' \ K\ by (simp add: Int_absorb1 Int_absorb2) with K'_le_K show False by (auto simp: min_def split: if_split_asm) qed end diff --git a/src/HOL/Analysis/Finite_Product_Measure.thy b/src/HOL/Analysis/Finite_Product_Measure.thy --- a/src/HOL/Analysis/Finite_Product_Measure.thy +++ b/src/HOL/Analysis/Finite_Product_Measure.thy @@ -1,1343 +1,1380 @@ (* Title: HOL/Analysis/Finite_Product_Measure.thy Author: Johannes Hölzl, TU München *) section \Finite Product Measure\ theory Finite_Product_Measure imports Binary_Product_Measure Function_Topology begin lemma PiE_choice: "(\f\Pi\<^sub>E I F. \i\I. P i (f i)) \ (\i\I. \x\F i. P i x)" by (auto simp: Bex_def PiE_iff Ball_def dest!: choice_iff'[THEN iffD1]) (force intro: exI[of _ "restrict f I" for f]) lemma case_prod_const: "(\(i, j). c) = (\_. c)" by auto subsection\<^marker>\tag unimportant\ \More about Function restricted by \<^const>\extensional\\ definition "merge I J = (\(x, y) i. if i \ I then x i else if i \ J then y i else undefined)" lemma merge_apply[simp]: "I \ J = {} \ i \ I \ merge I J (x, y) i = x i" "I \ J = {} \ i \ J \ merge I J (x, y) i = y i" "J \ I = {} \ i \ I \ merge I J (x, y) i = x i" "J \ I = {} \ i \ J \ merge I J (x, y) i = y i" "i \ I \ i \ J \ merge I J (x, y) i = undefined" unfolding merge_def by auto lemma merge_commute: "I \ J = {} \ merge I J (x, y) = merge J I (y, x)" by (force simp: merge_def) lemma Pi_cancel_merge_range[simp]: "I \ J = {} \ x \ Pi I (merge I J (A, B)) \ x \ Pi I A" "I \ J = {} \ x \ Pi I (merge J I (B, A)) \ x \ Pi I A" "J \ I = {} \ x \ Pi I (merge I J (A, B)) \ x \ Pi I A" "J \ I = {} \ x \ Pi I (merge J I (B, A)) \ x \ Pi I A" by (auto simp: Pi_def) lemma Pi_cancel_merge[simp]: "I \ J = {} \ merge I J (x, y) \ Pi I B \ x \ Pi I B" "J \ I = {} \ merge I J (x, y) \ Pi I B \ x \ Pi I B" "I \ J = {} \ merge I J (x, y) \ Pi J B \ y \ Pi J B" "J \ I = {} \ merge I J (x, y) \ Pi J B \ y \ Pi J B" by (auto simp: Pi_def) lemma extensional_merge[simp]: "merge I J (x, y) \ extensional (I \ J)" by (auto simp: extensional_def) lemma restrict_merge[simp]: "I \ J = {} \ restrict (merge I J (x, y)) I = restrict x I" "I \ J = {} \ restrict (merge I J (x, y)) J = restrict y J" "J \ I = {} \ restrict (merge I J (x, y)) I = restrict x I" "J \ I = {} \ restrict (merge I J (x, y)) J = restrict y J" by (auto simp: restrict_def) lemma split_merge: "P (merge I J (x,y) i) \ (i \ I \ P (x i)) \ (i \ J - I \ P (y i)) \ (i \ I \ J \ P undefined)" unfolding merge_def by auto lemma PiE_cancel_merge[simp]: "I \ J = {} \ merge I J (x, y) \ Pi\<^sub>E (I \ J) B \ x \ Pi I B \ y \ Pi J B" by (auto simp: PiE_def restrict_Pi_cancel) lemma merge_singleton[simp]: "i \ I \ merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)" unfolding merge_def by (auto simp: fun_eq_iff) lemma extensional_merge_sub: "I \ J \ K \ merge I J (x, y) \ extensional K" unfolding merge_def extensional_def by auto lemma merge_restrict[simp]: "merge I J (restrict x I, y) = merge I J (x, y)" "merge I J (x, restrict y J) = merge I J (x, y)" unfolding merge_def by auto lemma merge_x_x_eq_restrict[simp]: "merge I J (x, x) = restrict x (I \ J)" unfolding merge_def by auto lemma injective_vimage_restrict: assumes J: "J \ I" and sets: "A \ (\\<^sub>E i\J. S i)" "B \ (\\<^sub>E i\J. S i)" and ne: "(\\<^sub>E i\I. S i) \ {}" and eq: "(\x. restrict x J) -` A \ (\\<^sub>E i\I. S i) = (\x. restrict x J) -` B \ (\\<^sub>E i\I. S i)" shows "A = B" proof (intro set_eqI) fix x from ne obtain y where y: "\i. i \ I \ y i \ S i" by auto have "J \ (I - J) = {}" by auto show "x \ A \ x \ B" proof cases assume x: "x \ (\\<^sub>E i\J. S i)" have "x \ A \ merge J (I - J) (x,y) \ (\x. restrict x J) -` A \ (\\<^sub>E i\I. S i)" using y x \J \ I\ PiE_cancel_merge[of "J" "I - J" x y S] by (auto simp del: PiE_cancel_merge simp add: Un_absorb1) then show "x \ A \ x \ B" using y x \J \ I\ PiE_cancel_merge[of "J" "I - J" x y S] by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq) qed (insert sets, auto) qed lemma restrict_vimage: "I \ J = {} \ (\x. (restrict x I, restrict x J)) -` (Pi\<^sub>E I E \ Pi\<^sub>E J F) = Pi (I \ J) (merge I J (E, F))" by (auto simp: restrict_Pi_cancel PiE_def) lemma merge_vimage: "I \ J = {} \ merge I J -` Pi\<^sub>E (I \ J) E = Pi I E \ Pi J E" by (auto simp: restrict_Pi_cancel PiE_def) subsection \Finite product spaces\ definition\<^marker>\tag important\ prod_emb where "prod_emb I M K X = (\x. restrict x K) -` X \ (\\<^sub>E i\I. space (M i))" lemma prod_emb_iff: "f \ prod_emb I M K X \ f \ extensional I \ (restrict f K \ X) \ (\i\I. f i \ space (M i))" unfolding prod_emb_def PiE_def by auto lemma shows prod_emb_empty[simp]: "prod_emb M L K {} = {}" and prod_emb_Un[simp]: "prod_emb M L K (A \ B) = prod_emb M L K A \ prod_emb M L K B" and prod_emb_Int: "prod_emb M L K (A \ B) = prod_emb M L K A \ prod_emb M L K B" and prod_emb_UN[simp]: "prod_emb M L K (\i\I. F i) = (\i\I. prod_emb M L K (F i))" and prod_emb_INT[simp]: "I \ {} \ prod_emb M L K (\i\I. F i) = (\i\I. prod_emb M L K (F i))" and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B" by (auto simp: prod_emb_def) lemma prod_emb_PiE: "J \ I \ (\i. i \ J \ E i \ space (M i)) \ prod_emb I M J (\\<^sub>E i\J. E i) = (\\<^sub>E i\I. if i \ J then E i else space (M i))" by (force simp: prod_emb_def PiE_iff if_split_mem2) lemma prod_emb_PiE_same_index[simp]: "(\i. i \ I \ E i \ space (M i)) \ prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E" by (auto simp: prod_emb_def PiE_iff) lemma prod_emb_trans[simp]: "J \ K \ K \ L \ prod_emb L M K (prod_emb K M J X) = prod_emb L M J X" by (auto simp add: Int_absorb1 prod_emb_def PiE_def) lemma prod_emb_Pi: assumes "X \ (\ j\J. sets (M j))" "J \ K" shows "prod_emb K M J (Pi\<^sub>E J X) = (\\<^sub>E i\K. if i \ J then X i else space (M i))" using assms sets.space_closed by (auto simp: prod_emb_def PiE_iff split: if_split_asm) blast+ lemma prod_emb_id: "B \ (\\<^sub>E i\L. space (M i)) \ prod_emb L M L B = B" by (auto simp: prod_emb_def subset_eq extensional_restrict) lemma prod_emb_mono: "F \ G \ prod_emb A M B F \ prod_emb A M B G" by (auto simp: prod_emb_def) definition\<^marker>\tag important\ PiM :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) measure" where "PiM I M = extend_measure (\\<^sub>E i\I. space (M i)) {(J, X). (J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))} (\(J, X). prod_emb I M J (\\<^sub>E j\J. X j)) (\(J, X). \j\J \ {i\I. emeasure (M i) (space (M i)) \ 1}. if j \ J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))" definition\<^marker>\tag important\ prod_algebra :: "'i set \ ('i \ 'a measure) \ ('i \ 'a) set set" where "prod_algebra I M = (\(J, X). prod_emb I M J (\\<^sub>E j\J. X j)) ` {(J, X). (J \ {} \ I = {}) \ finite J \ J \ I \ X \ (\ j\J. sets (M j))}" abbreviation "Pi\<^sub>M I M \ PiM I M" syntax "_PiM" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" ("(3\\<^sub>M _\_./ _)" 10) translations "\\<^sub>M x\I. M" == "CONST PiM I (%x. M)" lemma extend_measure_cong: assumes "\ = \'" "I = I'" "G = G'" "\i. i \ I' \ \ i = \' i" shows "extend_measure \ I G \ = extend_measure \' I' G' \'" unfolding extend_measure_def by (auto simp add: assms) lemma Pi_cong_sets: "\I = J; \x. x \ I \ M x = N x\ \ Pi I M = Pi J N" unfolding Pi_def by auto lemma PiM_cong: assumes "I = J" "\x. x \ I \ M x = N x" shows "PiM I M = PiM J N" unfolding PiM_def proof (rule extend_measure_cong, goal_cases) case 1 show ?case using assms by (subst assms(1), intro PiE_cong[of J "\i. space (M i)" "\i. space (N i)"]) simp_all next case 2 have "\K. K \ J \ (\ j\K. sets (M j)) = (\ j\K. sets (N j))" using assms by (intro Pi_cong_sets) auto thus ?case by (auto simp: assms) next case 3 show ?case using assms by (intro ext) (auto simp: prod_emb_def dest: PiE_mem) next case (4 x) thus ?case using assms by (auto intro!: prod.cong split: if_split_asm) qed lemma prod_algebra_sets_into_space: "prod_algebra I M \ Pow (\\<^sub>E i\I. space (M i))" by (auto simp: prod_emb_def prod_algebra_def) lemma prod_algebra_eq_finite: assumes I: "finite I" shows "prod_algebra I M = {(\\<^sub>E i\I. X i) |X. X \ (\ j\I. sets (M j))}" (is "?L = ?R") proof (intro iffI set_eqI) fix A assume "A \ ?L" then obtain J E where J: "J \ {} \ I = {}" "finite J" "J \ I" "\i\J. E i \ sets (M i)" and A: "A = prod_emb I M J (\\<^sub>E j\J. E j)" by (auto simp: prod_algebra_def) let ?A = "\\<^sub>E i\I. if i \ J then E i else space (M i)" have A: "A = ?A" unfolding A using J by (intro prod_emb_PiE sets.sets_into_space) auto show "A \ ?R" unfolding A using J sets.top by (intro CollectI exI[of _ "\i. if i \ J then E i else space (M i)"]) simp next fix A assume "A \ ?R" then obtain X where A: "A = (\\<^sub>E i\I. X i)" and X: "X \ (\ j\I. sets (M j))" by auto then have A: "A = prod_emb I M I (\\<^sub>E i\I. X i)" by (simp add: prod_emb_PiE_same_index[OF sets.sets_into_space] Pi_iff) from X I show "A \ ?L" unfolding A by (auto simp: prod_algebra_def) qed lemma prod_algebraI: "finite J \ (J \ {} \ I = {}) \ J \ I \ (\i. i \ J \ E i \ sets (M i)) \ prod_emb I M J (\\<^sub>E j\J. E j) \ prod_algebra I M" by (auto simp: prod_algebra_def) lemma prod_algebraI_finite: "finite I \ (\i\I. E i \ sets (M i)) \ (Pi\<^sub>E I E) \ prod_algebra I M" using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \i\I. E i \ sets (M i)}" proof (safe intro!: Int_stableI) fix E F assume "\i\I. E i \ sets (M i)" "\i\I. F i \ sets (M i)" then show "\G. Pi\<^sub>E J E \ Pi\<^sub>E J F = Pi\<^sub>E J G \ (\i\I. G i \ sets (M i))" by (auto intro!: exI[of _ "\i. E i \ F i"] simp: PiE_Int) qed lemma prod_algebraE: assumes A: "A \ prod_algebra I M" obtains J E where "A = prod_emb I M J (\\<^sub>E j\J. E j)" "finite J" "J \ {} \ I = {}" "J \ I" "\i. i \ J \ E i \ sets (M i)" using A by (auto simp: prod_algebra_def) lemma prod_algebraE_all: assumes A: "A \ prod_algebra I M" obtains E where "A = Pi\<^sub>E I E" "E \ (\ i\I. sets (M i))" proof - from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)" and J: "J \ I" and E: "E \ (\ i\J. sets (M i))" by (auto simp: prod_algebra_def) from E have "\i. i \ J \ E i \ space (M i)" using sets.sets_into_space by auto then have "A = (\\<^sub>E i\I. if i\J then E i else space (M i))" using A J by (auto simp: prod_emb_PiE) moreover have "(\i. if i\J then E i else space (M i)) \ (\ i\I. sets (M i))" using sets.top E by auto ultimately show ?thesis using that by auto 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) B(5)[THEN sets.sets_into_space] apply (subst (1 2 3) prod_emb_PiE) apply (simp_all add: subset_eq PiE_Int) apply blast apply (intro PiE_cong) apply auto done also have "\ \ prod_algebra I M" using A B by (auto intro!: prod_algebraI) finally show "A \ B \ prod_algebra I M" . qed proposition prod_algebra_mono: assumes space: "\i. i \ I \ space (E i) = space (F i)" assumes sets: "\i. i \ I \ sets (E i) \ sets (F i)" shows "prod_algebra I E \ prod_algebra I F" proof fix A assume "A \ prod_algebra I E" then obtain J G where J: "J \ {} \ I = {}" "finite J" "J \ I" and A: "A = prod_emb I E J (\\<^sub>E i\J. G i)" and G: "\i. i \ J \ G i \ sets (E i)" by (auto simp: prod_algebra_def) moreover from space have "(\\<^sub>E i\I. space (E i)) = (\\<^sub>E i\I. space (F i))" by (rule PiE_cong) with A have "A = prod_emb I F J (\\<^sub>E i\J. G i)" by (simp add: prod_emb_def) moreover from sets G J have "\i. i \ J \ G i \ sets (F i)" by auto ultimately show "A \ prod_algebra I F" apply (simp add: prod_algebra_def image_iff) apply (intro exI[of _ J] exI[of _ G] conjI) apply auto done qed proposition prod_algebra_cong: assumes "I = J" and sets: "(\i. i \ I \ sets (M i) = sets (N i))" shows "prod_algebra I M = prod_algebra J N" proof - have space: "\i. i \ I \ space (M i) = space (N i)" using sets_eq_imp_space_eq[OF sets] by auto with sets show ?thesis unfolding \I = J\ by (intro antisym prod_algebra_mono) auto qed lemma space_in_prod_algebra: "(\\<^sub>E i\I. space (M i)) \ prod_algebra I M" proof cases assume "I = {}" then show ?thesis by (auto simp add: prod_algebra_def image_iff prod_emb_def) next assume "I \ {}" then obtain i where "i \ I" by auto then have "(\\<^sub>E i\I. space (M i)) = prod_emb I M {i} (\\<^sub>E i\{i}. space (M i))" by (auto simp: prod_emb_def) also have "\ \ prod_algebra I M" using \i \ I\ by (intro prod_algebraI) auto finally show ?thesis . qed lemma space_PiM: "space (\\<^sub>M i\I. M i) = (\\<^sub>E i\I. space (M i))" using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp lemma prod_emb_subset_PiM[simp]: "prod_emb I M K X \ space (PiM I M)" by (auto simp: prod_emb_def space_PiM) lemma space_PiM_empty_iff[simp]: "space (PiM I M) = {} \ (\i\I. space (M i) = {})" by (auto simp: space_PiM PiE_eq_empty_iff) lemma undefined_in_PiM_empty[simp]: "(\x. undefined) \ space (PiM {} M)" by (auto simp: space_PiM) lemma sets_PiM: "sets (\\<^sub>M i\I. M i) = sigma_sets (\\<^sub>E i\I. space (M i)) (prod_algebra I M)" using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp proposition sets_PiM_single: "sets (PiM I M) = sigma_sets (\\<^sub>E i\I. space (M i)) {{f\\\<^sub>E i\I. space (M i). f i \ A} | i A. i \ I \ A \ sets (M i)}" (is "_ = sigma_sets ?\ ?R") unfolding sets_PiM 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 = {}" with X have "A = {\x. undefined}" by (auto simp: prod_emb_def) with \I = {}\ show ?thesis by (auto intro!: sigma_sets_top) next assume "I \ {}" with X have "A = (\j\J. {f\(\\<^sub>E i\I. space (M i)). f j \ X j})" by (auto simp: prod_emb_def) also have "\ \ sigma_sets ?\ ?R" using X \I \ {}\ by (intro R.finite_INT sigma_sets.Basic) auto finally show "A \ sigma_sets ?\ ?R" . qed next fix A assume "A \ ?R" then obtain i B where A: "A = {f\\\<^sub>E i\I. space (M i). f i \ B}" "i \ I" "B \ sets (M i)" by auto then have "A = prod_emb I M {i} (\\<^sub>E i\{i}. B)" by (auto simp: prod_emb_def) also have "\ \ sigma_sets ?\ (prod_algebra I M)" using A by (intro sigma_sets.Basic prod_algebraI) auto finally show "A \ sigma_sets ?\ (prod_algebra I M)" . qed lemma sets_PiM_eq_proj: "I \ {} \ sets (PiM I M) = sets (SUP i\I. vimage_algebra (\\<^sub>E i\I. space (M i)) (\x. x i) (M i))" apply (simp add: sets_PiM_single) apply (subst sets_Sup_eq[where X="\\<^sub>E i\I. space (M i)"]) apply auto [] apply auto [] apply simp apply (subst arg_cong [of _ _ Sup, OF image_cong, OF refl]) apply (rule sets_vimage_algebra2) apply (auto intro!: arg_cong2[where f=sigma_sets]) done lemma shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\k. undefined}" and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\k. undefined} }" by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq) proposition sets_PiM_sigma: assumes \_cover: "\i. i \ I \ \S\E i. countable S \ \ i = \S" assumes E: "\i. i \ I \ E i \ Pow (\ i)" assumes J: "\j. j \ J \ finite j" "\J = I" defines "P \ {{f\(\\<^sub>E i\I. \ i). \i\j. f i \ A i} | A j. j \ J \ A \ Pi j E}" shows "sets (\\<^sub>M i\I. sigma (\ i) (E i)) = sets (sigma (\\<^sub>E i\I. \ i) P)" proof cases assume "I = {}" with \\J = I\ have "P = {{\_. undefined}} \ P = {}" by (auto simp: P_def) with \I = {}\ show ?thesis by (auto simp add: sets_PiM_empty sigma_sets_empty_eq) next let ?F = "\i. {(\x. x i) -` A \ Pi\<^sub>E I \ |A. A \ E i}" assume "I \ {}" then have "sets (Pi\<^sub>M I (\i. sigma (\ i) (E i))) = sets (SUP i\I. vimage_algebra (\\<^sub>E i\I. \ i) (\x. x i) (sigma (\ i) (E i)))" by (subst sets_PiM_eq_proj) (auto simp: space_measure_of_conv) also have "\ = sets (SUP i\I. sigma (Pi\<^sub>E I \) (?F i))" using E by (intro sets_SUP_cong arg_cong[where f=sets] vimage_algebra_sigma) auto also have "\ = sets (sigma (Pi\<^sub>E I \) (\i\I. ?F i))" using \I \ {}\ by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto also have "\ = sets (sigma (Pi\<^sub>E I \) P)" proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI) show "(\i\I. ?F i) \ Pow (Pi\<^sub>E I \)" "P \ Pow (Pi\<^sub>E I \)" by (auto simp: P_def) next interpret P: sigma_algebra "\\<^sub>E i\I. \ i" "sigma_sets (\\<^sub>E i\I. \ i) P" by (auto intro!: sigma_algebra_sigma_sets simp: P_def) fix Z assume "Z \ (\i\I. ?F i)" then obtain i A where i: "i \ I" "A \ E i" and Z_def: "Z = (\x. x i) -` A \ Pi\<^sub>E I \" by auto from \i \ I\ J obtain j where j: "i \ j" "j \ J" "j \ I" "finite j" by auto obtain S where S: "\i. i \ j \ S i \ E i" "\i. i \ j \ countable (S i)" "\i. i \ j \ \ i = \(S i)" by (metis subset_eq \_cover \j \ I\) define A' where "A' n = n(i := A)" for n then have A'_i: "\n. A' n i = A" by simp { fix n assume "n \ Pi\<^sub>E (j - {i}) S" then have "A' n \ Pi j E" unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def \A \ E i\ ) with \j \ J\ have "{f \ Pi\<^sub>E I \. \i\j. f i \ A' n i} \ P" by (auto simp: P_def) } note A'_in_P = this { fix x assume "x i \ A" "x \ Pi\<^sub>E I \" with S(3) \j \ I\ have "\i\j. \s\S i. x i \ s" by (auto simp: PiE_def Pi_def) then obtain s where s: "\i. i \ j \ s i \ S i" "\i. i \ j \ x i \ s i" by metis with \x i \ A\ have "\n\Pi\<^sub>E (j-{i}) S. \i\j. x i \ A' n i" by (intro bexI[of _ "restrict (s(i := A)) (j-{i})"]) (auto simp: A'_def split: if_splits) } then have "Z = (\n\Pi\<^sub>E (j-{i}) S. {f\(\\<^sub>E i\I. \ i). \i\j. f i \ A' n i})" unfolding Z_def by (auto simp add: set_eq_iff ball_conj_distrib \i\j\ A'_i dest: bspec[OF _ \i\j\] cong: conj_cong) also have "\ \ sigma_sets (\\<^sub>E i\I. \ i) P" using \finite j\ S(2) by (intro P.countable_UN' countable_PiE) (simp_all add: image_subset_iff A'_in_P) finally show "Z \ sigma_sets (\\<^sub>E i\I. \ i) P" . next interpret F: sigma_algebra "\\<^sub>E i\I. \ i" "sigma_sets (\\<^sub>E i\I. \ i) (\i\I. ?F i)" by (auto intro!: sigma_algebra_sigma_sets) fix b assume "b \ P" then obtain A j where b: "b = {f\(\\<^sub>E i\I. \ i). \i\j. f i \ A i}" "j \ J" "A \ Pi j E" by (auto simp: P_def) show "b \ sigma_sets (Pi\<^sub>E I \) (\i\I. ?F i)" proof cases assume "j = {}" with b have "b = (\\<^sub>E i\I. \ i)" by auto then show ?thesis by blast next assume "j \ {}" with J b(2,3) have eq: "b = (\i\j. ((\x. x i) -` A i \ Pi\<^sub>E I \))" unfolding b(1) by (auto simp: PiE_def Pi_def) show ?thesis unfolding eq using \A \ Pi j E\ \j \ J\ J(2) by (intro F.finite_INT J \j \ J\ \j \ {}\ sigma_sets.Basic) blast qed qed finally show "?thesis" . qed lemma sets_PiM_in_sets: assumes space: "space N = (\\<^sub>E i\I. space (M i))" assumes sets: "\i A. i \ I \ A \ sets (M i) \ {x\space N. x i \ A} \ sets N" shows "sets (\\<^sub>M i \ I. M i) \ sets N" unfolding sets_PiM_single space[symmetric] by (intro sets.sigma_sets_subset subsetI) (auto intro: sets) lemma sets_PiM_cong[measurable_cong]: assumes "I = J" "\i. i \ J \ sets (M i) = sets (N i)" shows "sets (PiM I M) = sets (PiM J N)" using assms sets_eq_imp_space_eq[OF assms(2)] by (simp add: sets_PiM_single cong: PiE_cong conj_cong) lemma sets_PiM_I: assumes "finite J" "J \ I" "\i\J. E i \ sets (M i)" shows "prod_emb I M J (\\<^sub>E j\J. E j) \ sets (\\<^sub>M i\I. M i)" proof cases assume "J = {}" then have "prod_emb I M J (\\<^sub>E j\J. E j) = (\\<^sub>E j\I. space (M j))" by (auto simp: prod_emb_def) then show ?thesis by (auto simp add: sets_PiM intro!: sigma_sets_top) next assume "J \ {}" with assms show ?thesis by (force simp add: sets_PiM prod_algebra_def) qed proposition measurable_PiM: assumes space: "f \ space N \ (\\<^sub>E i\I. space (M i))" assumes sets: "\X J. J \ {} \ I = {} \ finite J \ J \ I \ (\i. i \ J \ X i \ sets (M i)) \ f -` prod_emb I M J (Pi\<^sub>E J X) \ space N \ sets N" shows "f \ measurable N (PiM I M)" 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 lemma measurable_PiM_Collect: assumes space: "f \ space N \ (\\<^sub>E i\I. space (M i))" assumes sets: "\X J. J \ {} \ I = {} \ finite J \ J \ I \ (\i. i \ J \ X i \ sets (M i)) \ {\\space N. \i\J. f \ i \ X i} \ sets N" shows "f \ measurable N (PiM I M)" 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) finally show "f -` A \ space N \ sets N" . qed lemma measurable_PiM_single: assumes space: "f \ space N \ (\\<^sub>E i\I. space (M i))" assumes sets: "\A i. i \ I \ A \ sets (M i) \ {\ \ space N. f \ i \ A} \ sets N" shows "f \ measurable N (PiM I M)" using sets_PiM_single proof (rule measurable_sigma_sets) fix A assume "A \ {{f \ \\<^sub>E i\I. space (M i). f i \ A} |i A. i \ I \ A \ sets (M i)}" then obtain B i where "A = {f \ \\<^sub>E i\I. space (M i). f i \ B}" and B: "i \ I" "B \ sets (M i)" by auto with space have "f -` A \ space N = {\ \ space N. f \ i \ B}" by auto also have "\ \ sets N" using B by (rule sets) finally show "f -` A \ space N \ sets N" . qed (auto simp: space) lemma measurable_PiM_single': assumes f: "\i. i \ I \ f i \ measurable N (M i)" and "(\\ i. f i \) \ space N \ (\\<^sub>E i\I. space (M i))" shows "(\\ i. f i \) \ measurable N (Pi\<^sub>M I M)" proof (rule measurable_PiM_single) fix A i assume A: "i \ I" "A \ sets (M i)" then have "{\ \ space N. f i \ \ A} = f i -` A \ space N" by auto then show "{\ \ space N. f i \ \ A} \ sets N" using A f by (auto intro!: measurable_sets) qed fact lemma sets_PiM_I_finite[measurable]: assumes "finite I" and sets: "(\i. i \ I \ E i \ sets (M i))" shows "(\\<^sub>E j\I. E j) \ sets (\\<^sub>M i\I. M i)" using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] \finite I\ sets by auto lemma measurable_component_singleton[measurable (raw)]: assumes "i \ I" shows "(\x. x i) \ measurable (Pi\<^sub>M I M) (M i)" proof (unfold measurable_def, intro CollectI conjI ballI) fix A assume "A \ sets (M i)" then have "(\x. x i) -` A \ space (Pi\<^sub>M I M) = prod_emb I M {i} (\\<^sub>E j\{i}. A)" using sets.sets_into_space \i \ I\ by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: if_split_asm) then show "(\x. x i) -` A \ space (Pi\<^sub>M I M) \ sets (Pi\<^sub>M I M)" using \A \ sets (M i)\ \i \ I\ by (auto intro!: sets_PiM_I) qed (insert \i \ I\, auto simp: space_PiM) lemma measurable_component_singleton'[measurable_dest]: assumes f: "f \ measurable N (Pi\<^sub>M I M)" assumes g: "g \ measurable L N" assumes i: "i \ I" shows "(\x. (f (g x)) i) \ measurable L (M i)" using measurable_compose[OF measurable_compose[OF g f] measurable_component_singleton, OF i] . lemma measurable_PiM_component_rev: "i \ I \ f \ measurable (M i) N \ (\x. f (x i)) \ measurable (PiM I M) N" by simp lemma measurable_case_nat[measurable (raw)]: assumes [measurable (raw)]: "i = 0 \ f \ measurable M N" "\j. i = Suc j \ (\x. g x j) \ measurable M N" shows "(\x. case_nat (f x) (g x) i) \ measurable M N" by (cases i) simp_all lemma measurable_case_nat'[measurable (raw)]: assumes fg[measurable]: "f \ measurable N M" "g \ measurable N (\\<^sub>M i\UNIV. M)" shows "(\x. case_nat (f x) (g x)) \ measurable N (\\<^sub>M i\UNIV. M)" using fg[THEN measurable_space] by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split) lemma measurable_add_dim[measurable]: "(\(f, y). f(i := y)) \ measurable (Pi\<^sub>M I M \\<^sub>M M i) (Pi\<^sub>M (insert i I) M)" (is "?f \ measurable ?P ?I") proof (rule measurable_PiM_single) fix j A assume j: "j \ insert i I" and A: "A \ sets (M j)" have "{\ \ space ?P. (\(f, x). fun_upd f i x) \ j \ A} = (if j = i then space (Pi\<^sub>M I M) \ A else ((\x. x j) \ fst) -` A \ space ?P)" using sets.sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM) also have "\ \ sets ?P" using A j by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) finally show "{\ \ space ?P. case_prod (\f. fun_upd f i) \ j \ A} \ sets ?P" . qed (auto simp: space_pair_measure space_PiM PiE_def) proposition measurable_fun_upd: assumes I: "I = J \ {i}" assumes f[measurable]: "f \ measurable N (PiM J M)" assumes h[measurable]: "h \ measurable N (M i)" shows "(\x. (f x) (i := h x)) \ measurable N (PiM I M)" proof (intro measurable_PiM_single') fix j assume "j \ I" then show "(\\. ((f \)(i := h \)) j) \ measurable N (M j)" unfolding I by (cases "j = i") auto next show "(\x. (f x)(i := h x)) \ space N \ (\\<^sub>E i\I. space (M i))" using I f[THEN measurable_space] h[THEN measurable_space] by (auto simp: space_PiM PiE_iff extensional_def) qed lemma measurable_component_update: "x \ space (Pi\<^sub>M I M) \ i \ I \ (\v. x(i := v)) \ measurable (M i) (Pi\<^sub>M (insert i I) M)" by simp lemma measurable_merge[measurable]: "merge I J \ measurable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \ J) M)" (is "?f \ measurable ?P ?U") proof (rule measurable_PiM_single) fix i A assume A: "A \ sets (M i)" "i \ I \ J" then have "{\ \ space ?P. merge I J \ i \ A} = (if i \ I then ((\x. x i) \ fst) -` A \ space ?P else ((\x. x i) \ snd) -` A \ space ?P)" by (auto simp: merge_def) also have "\ \ sets ?P" using A by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton]) finally show "{\ \ space ?P. merge I J \ i \ A} \ sets ?P" . qed (auto simp: space_pair_measure space_PiM PiE_iff merge_def extensional_def) lemma measurable_restrict[measurable (raw)]: assumes X: "\i. i \ I \ X i \ measurable N (M i)" shows "(\x. \i\I. X i x) \ measurable N (Pi\<^sub>M I M)" proof (rule measurable_PiM_single) fix A i assume A: "i \ I" "A \ sets (M i)" then have "{\ \ space N. (\i\I. X i \) i \ A} = X i -` A \ space N" by auto then show "{\ \ space N. (\i\I. X i \) i \ A} \ sets N" using A X by (auto intro!: measurable_sets) qed (insert X, auto simp add: PiE_def dest: measurable_space) lemma measurable_abs_UNIV: "(\n. (\\. f n \) \ measurable M (N n)) \ (\\ n. f n \) \ measurable M (PiM UNIV N)" by (intro measurable_PiM_single) (auto dest: measurable_space) lemma measurable_restrict_subset: "J \ L \ (\f. restrict f J) \ measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)" by (intro measurable_restrict measurable_component_singleton) auto lemma measurable_restrict_subset': assumes "J \ L" "\x. x \ J \ sets (M x) = sets (N x)" shows "(\f. restrict f J) \ measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)" proof- from assms(1) have "(\f. restrict f J) \ measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)" by (rule measurable_restrict_subset) also from assms(2) have "measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M) = measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)" by (intro sets_PiM_cong measurable_cong_sets) simp_all finally show ?thesis . qed lemma measurable_prod_emb[intro, simp]: "J \ L \ X \ sets (Pi\<^sub>M J M) \ prod_emb L M J X \ sets (Pi\<^sub>M L M)" unfolding prod_emb_def space_PiM[symmetric] by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton) lemma merge_in_prod_emb: assumes "y \ space (PiM I M)" "x \ X" and X: "X \ sets (Pi\<^sub>M J M)" and "J \ I" shows "merge J I (x, y) \ prod_emb I M J X" using assms sets.sets_into_space[OF X] by (simp add: merge_def prod_emb_def subset_eq space_PiM PiE_def extensional_restrict Pi_iff cong: if_cong restrict_cong) (simp add: extensional_def) lemma prod_emb_eq_emptyD: assumes J: "J \ I" and ne: "space (PiM I M) \ {}" and X: "X \ sets (Pi\<^sub>M J M)" and *: "prod_emb I M J X = {}" shows "X = {}" proof safe fix x assume "x \ X" obtain \ where "\ \ space (PiM I M)" using ne by blast from merge_in_prod_emb[OF this \x\X\ X J] * show "x \ {}" by auto qed lemma sets_in_Pi_aux: "finite I \ (\j. j \ I \ {x\space (M j). x \ F j} \ sets (M j)) \ {x\space (PiM I M). x \ Pi I F} \ sets (PiM I M)" by (simp add: subset_eq Pi_iff) lemma sets_in_Pi[measurable (raw)]: "finite I \ f \ measurable N (PiM I M) \ (\j. j \ I \ {x\space (M j). x \ F j} \ sets (M j)) \ Measurable.pred N (\x. f x \ Pi I F)" unfolding pred_def by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_Pi_aux]) auto lemma sets_in_extensional_aux: "{x\space (PiM I M). x \ extensional I} \ sets (PiM I M)" proof - have "{x\space (PiM I M). x \ extensional I} = space (PiM I M)" by (auto simp add: extensional_def space_PiM) then show ?thesis by simp qed lemma sets_in_extensional[measurable (raw)]: "f \ measurable N (PiM I M) \ Measurable.pred N (\x. f x \ extensional I)" unfolding pred_def by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto lemma sets_PiM_I_countable: assumes I: "countable I" and E: "\i. i \ I \ E i \ sets (M i)" shows "Pi\<^sub>E I E \ sets (Pi\<^sub>M I M)" proof cases assume "I \ {}" then have "Pi\<^sub>E I E = (\i\I. prod_emb I M {i} (Pi\<^sub>E {i} E))" using E[THEN sets.sets_into_space] by (auto simp: PiE_iff prod_emb_def fun_eq_iff) also have "\ \ sets (PiM I M)" using I \I \ {}\ by (safe intro!: sets.countable_INT' measurable_prod_emb sets_PiM_I_finite E) finally show ?thesis . qed (simp add: sets_PiM_empty) lemma sets_PiM_D_countable: assumes A: "A \ PiM I M" shows "\J\I. \X\PiM J M. countable J \ A = prod_emb I M J X" using A[unfolded sets_PiM_single] proof induction case (Basic A) then obtain i X where *: "i \ I" "X \ sets (M i)" and "A = {f \ \\<^sub>E i\I. space (M i). f i \ X}" by auto then have A: "A = prod_emb I M {i} (\\<^sub>E _\{i}. X)" by (auto simp: prod_emb_def) then show ?case by (intro exI[of _ "{i}"] conjI bexI[of _ "\\<^sub>E _\{i}. X"]) (auto intro: countable_finite * sets_PiM_I_finite) next case Empty then show ?case by (intro exI[of _ "{}"] conjI bexI[of _ "{}"]) auto next case (Compl A) then obtain J X where "J \ I" "X \ sets (Pi\<^sub>M J M)" "countable J" "A = prod_emb I M J X" by auto then show ?case by (intro exI[of _ J] bexI[of _ "space (PiM J M) - X"] conjI) (auto simp add: space_PiM prod_emb_PiE intro!: sets_PiM_I_countable) next case (Union K) obtain J X where J: "\i. J i \ I" "\i. countable (J i)" and X: "\i. X i \ sets (Pi\<^sub>M (J i) M)" and K: "\i. K i = prod_emb I M (J i) (X i)" by (metis Union.IH) show ?case proof (intro exI[of _ "\i. J i"] bexI[of _ "\i. prod_emb (\i. J i) M (J i) (X i)"] conjI) show "(\i. J i) \ I" "countable (\i. J i)" using J by auto with J show "\(K ` UNIV) = prod_emb I M (\i. J i) (\i. prod_emb (\i. J i) M (J i) (X i))" by (simp add: K[abs_def] SUP_upper) qed(auto intro: X) qed proposition measure_eqI_PiM_finite: assumes [simp]: "finite I" "sets P = PiM I M" "sets Q = PiM I M" assumes eq: "\A. (\i. i \ I \ A i \ sets (M i)) \ P (Pi\<^sub>E I A) = Q (Pi\<^sub>E I A)" assumes A: "range A \ prod_algebra I M" "(\i. A i) = space (PiM I M)" "\i::nat. P (A i) \ \" shows "P = Q" proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space]) show "range A \ prod_algebra I M" "(\i. A i) = (\\<^sub>E i\I. space (M i))" "\i. P (A i) \ \" unfolding space_PiM[symmetric] by fact+ fix X assume "X \ prod_algebra I M" then obtain J E where X: "X = prod_emb I M J (\\<^sub>E j\J. E j)" and J: "finite J" "J \ I" "\j. j \ J \ E j \ sets (M j)" by (force elim!: prod_algebraE) then show "emeasure P X = emeasure Q X" unfolding X by (subst (1 2) prod_emb_Pi) (auto simp: eq) qed (simp_all add: sets_PiM) proposition measure_eqI_PiM_infinite: assumes [simp]: "sets P = PiM I M" "sets Q = PiM I M" assumes eq: "\A J. finite J \ J \ I \ (\i. i \ J \ A i \ sets (M i)) \ P (prod_emb I M J (Pi\<^sub>E J A)) = Q (prod_emb I M J (Pi\<^sub>E J A))" assumes A: "finite_measure P" shows "P = Q" proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space]) interpret finite_measure P by fact define i where "i = (SOME i. i \ I)" have i: "I \ {} \ i \ I" unfolding i_def by (rule someI_ex) auto define A where "A n = (if I = {} then prod_emb I M {} (\\<^sub>E i\{}. {}) else prod_emb I M {i} (\\<^sub>E i\{i}. space (M i)))" for n :: nat then show "range A \ prod_algebra I M" using prod_algebraI[of "{}" I "\i. space (M i)" M] by (auto intro!: prod_algebraI i) have "\i. A i = space (PiM I M)" by (auto simp: prod_emb_def space_PiM PiE_iff A_def i ex_in_conv[symmetric] exI) then show "(\i. A i) = (\\<^sub>E i\I. space (M i))" "\i. emeasure P (A i) \ \" by (auto simp: space_PiM) next fix X assume X: "X \ prod_algebra I M" then obtain J E where X: "X = prod_emb I M J (\\<^sub>E j\J. E j)" and J: "finite J" "J \ I" "\j. j \ J \ E j \ sets (M j)" by (force elim!: prod_algebraE) then show "emeasure P X = emeasure Q X" by (auto intro!: eq) qed (auto simp: sets_PiM) locale\<^marker>\tag unimportant\ product_sigma_finite = fixes M :: "'i \ 'a measure" assumes sigma_finite_measures: "\i. sigma_finite_measure (M i)" sublocale\<^marker>\tag unimportant\ product_sigma_finite \ M?: sigma_finite_measure "M i" for i by (rule sigma_finite_measures) locale\<^marker>\tag unimportant\ finite_product_sigma_finite = product_sigma_finite M for M :: "'i \ 'a measure" + fixes I :: "'i set" assumes finite_index: "finite I" proposition (in finite_product_sigma_finite) sigma_finite_pairs: "\F::'i \ nat \ 'a set. (\i\I. range (F i) \ sets (M i)) \ (\k. \i\I. emeasure (M i) (F i k) \ \) \ incseq (\k. \\<^sub>E i\I. F i k) \ (\k. \\<^sub>E i\I. F i k) = space (PiM I M)" 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" note space_PiM[simp] show ?thesis proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) fix i show "range (F i) \ sets (M i)" by fact next fix i k show "emeasure (M i) (F i k) \ \" by fact next fix x assume "x \ (\i. ?F i)" with F(1) show "x \ space (PiM I M)" by (auto simp: PiE_def dest!: sets.sets_into_space) next fix f assume "f \ space (PiM I M)" with Pi_UN[OF finite_index, of "\k i. F i k"] F show "f \ (\i. ?F i)" by (auto simp: incseq_def PiE_def) next fix i show "?F i \ ?F (Suc i)" using \\i. incseq (F i)\[THEN incseq_SucD] by auto qed qed lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\_. undefined} = 1" proof - let ?\ = "\A. if A = {} then 0 else (1::ennreal)" have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\\<^sub>E i\{}. {})) = 1" proof (subst emeasure_extend_measure_Pair[OF PiM_def]) show "positive (PiM {} M) ?\" by (auto simp: positive_def) show "countably_additive (PiM {} M) ?\" by (rule sets.countably_additiveI_finite) (auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: ) qed (auto simp: prod_emb_def) also have "(prod_emb {} M {} (\\<^sub>E i\{}. {})) = {\_. undefined}" by (auto simp: prod_emb_def) finally show ?thesis by simp qed lemma PiM_empty: "PiM {} M = count_space {\_. undefined}" by (rule measure_eqI) (auto simp add: sets_PiM_empty) lemma (in product_sigma_finite) emeasure_PiM: "finite I \ (\i. i\I \ A i \ sets (M i)) \ emeasure (PiM I M) (Pi\<^sub>E I A) = (\i\I. emeasure (M i) (A i))" proof (induct I arbitrary: A rule: finite_induct) case (insert i I) interpret finite_product_sigma_finite M I by standard fact have "finite (insert i I)" using \finite I\ by auto interpret I': finite_product_sigma_finite M "insert i I" by standard fact let ?h = "(\(f, y). f(i := y))" let ?P = "distr (Pi\<^sub>M I M \\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h" let ?\ = "emeasure ?P" let ?I = "{j \ insert i I. emeasure (M j) (space (M j)) \ 1}" let ?f = "\J E j. if j \ J then emeasure (M j) (E j) else emeasure (M j) (space (M j))" have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) = (\i\insert i I. emeasure (M i) (A i))" proof (subst emeasure_extend_measure_Pair[OF PiM_def]) fix J E assume "(J \ {} \ insert i I = {}) \ finite J \ J \ insert i I \ E \ (\ j\J. sets (M j))" then have J: "J \ {}" "finite J" "J \ insert i I" and E: "\j\J. E j \ sets (M j)" by auto let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)" let ?p' = "prod_emb I M (J - {i}) (\\<^sub>E j\J-{i}. E j)" have "?\ ?p = emeasure (Pi\<^sub>M I M \\<^sub>M (M i)) (?h -` ?p \ space (Pi\<^sub>M I M \\<^sub>M M i))" by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+ also have "?h -` ?p \ space (Pi\<^sub>M I M \\<^sub>M M i) = ?p' \ (if i \ J then E i else space (M i))" using J E[rule_format, THEN sets.sets_into_space] by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: if_split_asm) also have "emeasure (Pi\<^sub>M I M \\<^sub>M (M i)) (?p' \ (if i \ J then E i else space (M i))) = emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \ J then (E i) else space (M i))" using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto also have "?p' = (\\<^sub>E j\I. if j \ J-{i} then E j else space (M j))" using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def Pi_iff split: if_split_asm) blast+ also have "emeasure (Pi\<^sub>M I M) (\\<^sub>E j\I. if j \ J-{i} then E j else space (M j)) = (\ j\I. if j \ J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))" using E by (subst insert) (auto intro!: prod.cong) also have "(\j\I. if j \ J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) * emeasure (M i) (if i \ J then E i else space (M i)) = (\j\insert i I. ?f J E j)" using insert by (auto simp: mult.commute intro!: arg_cong2[where f="(*)"] prod.cong) also have "\ = (\j\J \ ?I. ?f J E j)" using insert(1,2) J E by (intro prod.mono_neutral_right) auto finally show "?\ ?p = \" . show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \ Pow (\\<^sub>E i\insert i I. space (M i))" using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def) next show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\" using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all next show "(insert i I \ {} \ insert i I = {}) \ finite (insert i I) \ insert i I \ insert i I \ A \ (\ j\insert i I. sets (M j))" using insert by auto qed (auto intro!: prod.cong) with insert show ?case by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space) qed simp lemma (in product_sigma_finite) PiM_eqI: assumes I[simp]: "finite I" and P: "sets P = PiM I M" assumes eq: "\A. (\i. i \ I \ A i \ sets (M i)) \ P (Pi\<^sub>E I A) = (\i\I. emeasure (M i) (A i))" 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 by (simp add: eq emeasure_PiM) define A where "A n = (\\<^sub>E i\I. C i n)" for n with C show "range A \ prod_algebra I M" "\i. emeasure (Pi\<^sub>M I M) (A i) \ \" "(\i. A i) = space (PiM I M)" by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_prod_eq_top) qed qed lemma (in product_sigma_finite) sigma_finite: assumes "finite I" shows "sigma_finite_measure (PiM I M)" proof interpret finite_product_sigma_finite M I by standard fact obtain F where F: "\j. countable (F j)" "\j f. f \ F j \ f \ sets (M j)" "\j f. f \ F j \ emeasure (M j) f \ \" and in_space: "\j. space (M j) = \(F j)" using sigma_finite_countable by (metis subset_eq) moreover have "(\(Pi\<^sub>E I ` Pi\<^sub>E I F)) = space (Pi\<^sub>M I M)" using in_space by (auto simp: space_PiM PiE_iff intro!: PiE_choice[THEN iffD2]) ultimately show "\A. countable A \ A \ sets (Pi\<^sub>M I M) \ \A = space (Pi\<^sub>M I M) \ (\a\A. emeasure (Pi\<^sub>M I M) a \ \)" by (intro exI[of _ "Pi\<^sub>E I ` Pi\<^sub>E I F"]) (auto intro!: countable_PiE sets_PiM_I_finite simp: PiE_iff emeasure_PiM finite_index ennreal_prod_eq_top) qed sublocale finite_product_sigma_finite \ sigma_finite_measure "Pi\<^sub>M I M" using sigma_finite[OF finite_index] . lemma (in finite_product_sigma_finite) measure_times: "(\i. i \ I \ A i \ sets (M i)) \ emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\i\I. emeasure (M i) (A i))" using emeasure_PiM[OF finite_index] by auto lemma (in product_sigma_finite) nn_integral_empty: "0 \ f (\k. undefined) \ integral\<^sup>N (Pi\<^sub>M {} M) f = f (\k. undefined)" by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2) lemma\<^marker>\tag important\ (in product_sigma_finite) distr_merge: assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" shows "distr (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \ J) M) (merge I J) = Pi\<^sub>M (I \ J) M" (is "?D = ?P") proof (rule PiM_eqI) interpret I: finite_product_sigma_finite M I by standard fact interpret J: finite_product_sigma_finite M J by standard fact fix A assume A: "\i. i \ I \ J \ A i \ sets (M i)" have *: "(merge I J -` Pi\<^sub>E (I \ J) A \ space (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)) = Pi\<^sub>E I A \ Pi\<^sub>E J A" using A[THEN sets.sets_into_space] by (auto simp: space_PiM space_pair_measure) from A fin show "emeasure (distr (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \ J) M) (merge I J)) (Pi\<^sub>E (I \ J) A) = (\i\I \ J. emeasure (M i) (A i))" by (subst emeasure_distr) (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times prod.union_disjoint) qed (insert fin, simp_all) proposition (in product_sigma_finite) product_nn_integral_fold: assumes IJ: "I \ J = {}" "finite I" "finite J" and f[measurable]: "f \ borel_measurable (Pi\<^sub>M (I \ J) M)" shows "integral\<^sup>N (Pi\<^sub>M (I \ J) M) f = (\\<^sup>+ x. (\\<^sup>+ y. f (merge I J (x, y)) \(Pi\<^sub>M J M)) \(Pi\<^sub>M I M))" proof - interpret I: finite_product_sigma_finite M I by standard fact interpret J: finite_product_sigma_finite M J by standard fact interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard have P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)" using measurable_comp[OF measurable_merge f] by (simp add: comp_def) show ?thesis apply (subst distr_merge[OF IJ, symmetric]) apply (subst nn_integral_distr[OF measurable_merge]) apply measurable [] apply (subst J.nn_integral_fst[symmetric, OF P_borel]) apply simp done qed lemma (in product_sigma_finite) distr_singleton: "distr (Pi\<^sub>M {i} M) (M i) (\x. x i) = M i" (is "?D = _") proof (intro measure_eqI[symmetric]) interpret I: finite_product_sigma_finite M "{i}" by standard simp fix A assume A: "A \ sets (M i)" then have "(\x. x i) -` A \ space (Pi\<^sub>M {i} M) = (\\<^sub>E i\{i}. A)" using sets.sets_into_space by (auto simp: space_PiM) then show "emeasure (M i) A = emeasure ?D A" using A I.measure_times[of "\_. A"] by (simp add: emeasure_distr measurable_component_singleton) qed simp lemma (in product_sigma_finite) product_nn_integral_singleton: assumes f: "f \ borel_measurable (M i)" shows "integral\<^sup>N (Pi\<^sub>M {i} M) (\x. f (x i)) = integral\<^sup>N (M i) f" proof - interpret I: finite_product_sigma_finite M "{i}" by standard simp from f show ?thesis apply (subst distr_singleton[symmetric]) apply (subst nn_integral_distr[OF measurable_component_singleton]) apply simp_all done qed proposition (in product_sigma_finite) product_nn_integral_insert: assumes I[simp]: "finite I" "i \ I" and f: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\\<^sup>+ x. (\\<^sup>+ y. f (x(i := y)) \(M i)) \(Pi\<^sub>M I M))" proof - interpret I: finite_product_sigma_finite M I by standard auto interpret i: finite_product_sigma_finite M "{i}" by standard auto have IJ: "I \ {i} = {}" and insert: "I \ {i} = insert i I" using f by auto show ?thesis unfolding product_nn_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f] proof (rule nn_integral_cong, subst product_nn_integral_singleton[symmetric]) fix x assume x: "x \ space (Pi\<^sub>M I M)" let ?f = "\y. f (x(i := y))" show "?f \ borel_measurable (M i)" using measurable_comp[OF measurable_component_update f, OF x \i \ I\] unfolding comp_def . show "(\\<^sup>+ y. f (merge I {i} (x, y)) \Pi\<^sub>M {i} M) = (\\<^sup>+ y. f (x(i := y i)) \Pi\<^sub>M {i} M)" using x by (auto intro!: nn_integral_cong arg_cong[where f=f] simp add: space_PiM extensional_def PiE_def) qed qed lemma (in product_sigma_finite) product_nn_integral_insert_rev: assumes I[simp]: "finite I" "i \ I" and [measurable]: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\\<^sup>+ y. (\\<^sup>+ x. f (x(i := y)) \(Pi\<^sub>M I M)) \(M i))" apply (subst product_nn_integral_insert[OF assms]) apply (rule pair_sigma_finite.Fubini') apply intro_locales [] apply (rule sigma_finite[OF I(1)]) apply measurable done lemma (in product_sigma_finite) product_nn_integral_prod: assumes "finite I" "\i. i \ I \ f i \ borel_measurable (M i)" shows "(\\<^sup>+ x. (\i\I. f i (x i)) \Pi\<^sub>M I M) = (\i\I. integral\<^sup>N (M i) (f i))" using assms proof (induction I) case (insert i I) note insert.prems[measurable] note \finite I\[intro, simp] interpret I: finite_product_sigma_finite M I by standard auto have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" using insert by (auto intro!: prod.cong) have prod: "\J. J \ insert i I \ (\x. (\i\J. f i (x i))) \ borel_measurable (Pi\<^sub>M J M)" using sets.sets_into_space insert by (intro borel_measurable_prod_ennreal measurable_comp[OF measurable_component_singleton, unfolded comp_def]) auto then show ?case apply (simp add: product_nn_integral_insert[OF insert(1,2)]) apply (simp add: insert(2-) * nn_integral_multc) apply (subst nn_integral_cmult) apply (auto simp add: insert(2-)) done qed (simp add: space_PiM) proposition (in product_sigma_finite) product_nn_integral_pair: assumes [measurable]: "case_prod f \ borel_measurable (M x \\<^sub>M M y)" assumes xy: "x \ y" shows "(\\<^sup>+\. f (\ x) (\ y) \PiM {x, y} M) = (\\<^sup>+z. f (fst z) (snd z) \(M x \\<^sub>M M y))" proof - interpret psm: pair_sigma_finite "M x" "M y" unfolding pair_sigma_finite_def using sigma_finite_measures by simp_all have "{x, y} = {y, x}" by auto also have "(\\<^sup>+\. f (\ x) (\ y) \PiM {y, x} M) = (\\<^sup>+y. \\<^sup>+\. f (\ x) y \PiM {x} M \M y)" using xy by (subst product_nn_integral_insert_rev) simp_all also have "... = (\\<^sup>+y. \\<^sup>+x. f x y \M x \M y)" by (intro nn_integral_cong, subst product_nn_integral_singleton) simp_all also have "... = (\\<^sup>+z. f (fst z) (snd z) \(M x \\<^sub>M M y))" by (subst psm.nn_integral_snd[symmetric]) simp_all finally show ?thesis . qed lemma (in product_sigma_finite) distr_component: "distr (M i) (Pi\<^sub>M {i} M) (\x. \i\{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P") proof (intro PiM_eqI) fix A assume A: "\ia. ia \ {i} \ A ia \ sets (M ia)" then have "(\x. \i\{i}. x) -` Pi\<^sub>E {i} A \ space (M i) = A i" by (fastforce dest: sets.sets_into_space) with A show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\x. \i\{i}. x)) (Pi\<^sub>E {i} A) = (\i\{i}. emeasure (M i) (A i))" by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict) qed simp_all lemma (in product_sigma_finite) assumes IJ: "I \ J = {}" "finite I" "finite J" and A: "A \ sets (Pi\<^sub>M (I \ J) M)" shows emeasure_fold_integral: "emeasure (Pi\<^sub>M (I \ J) M) A = (\\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\y. merge I J (x, y)) -` A \ space (Pi\<^sub>M J M)) \Pi\<^sub>M I M)" (is ?I) and emeasure_fold_measurable: "(\x. emeasure (Pi\<^sub>M J M) ((\y. merge I J (x, y)) -` A \ space (Pi\<^sub>M J M))) \ borel_measurable (Pi\<^sub>M I M)" (is ?B) proof - interpret I: finite_product_sigma_finite M I by standard fact interpret J: finite_product_sigma_finite M J by standard fact interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" .. have merge: "merge I J -` A \ space (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) \ sets (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)" by (intro measurable_sets[OF _ A] measurable_merge assms) show ?I apply (subst distr_merge[symmetric, OF IJ]) apply (subst emeasure_distr[OF measurable_merge A]) apply (subst J.emeasure_pair_measure_alt[OF merge]) apply (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure) done show ?B using IJ.measurable_emeasure_Pair1[OF merge] by (simp add: vimage_comp comp_def space_pair_measure cong: measurable_cong) qed lemma sets_Collect_single: "i \ I \ A \ sets (M i) \ { x \ space (Pi\<^sub>M I M). x i \ A } \ sets (Pi\<^sub>M I M)" by simp lemma pair_measure_eq_distr_PiM: fixes M1 :: "'a measure" and M2 :: "'a measure" assumes "sigma_finite_measure M1" "sigma_finite_measure M2" shows "(M1 \\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \\<^sub>M M2) (\x. (x True, x False))" (is "?P = ?D") proof (rule pair_measure_eqI[OF assms]) interpret B: product_sigma_finite "case_bool M1 M2" unfolding product_sigma_finite_def using assms by (auto split: bool.split) let ?B = "Pi\<^sub>M UNIV (case_bool M1 M2)" have [simp]: "fst \ (\x. (x True, x False)) = (\x. x True)" "snd \ (\x. (x True, x False)) = (\x. x False)" by auto fix A B assume A: "A \ sets M1" and B: "B \ sets M2" have "emeasure M1 A * emeasure M2 B = (\ i\UNIV. emeasure (case_bool M1 M2 i) (case_bool A B i))" by (simp add: UNIV_bool ac_simps) also have "\ = emeasure ?B (Pi\<^sub>E UNIV (case_bool A B))" using A B by (subst B.emeasure_PiM) (auto split: bool.split) also have "Pi\<^sub>E UNIV (case_bool A B) = (\x. (x True, x False)) -` (A \ B) \ space ?B" using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space] by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split) finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \ B)" using A B measurable_component_singleton[of True UNIV "case_bool M1 M2"] measurable_component_singleton[of False UNIV "case_bool M1 M2"] by (subst emeasure_distr) (auto simp: measurable_pair_iff) qed simp lemma infprod_in_sets[intro]: fixes E :: "nat \ 'a set" assumes E: "\i. E i \ sets (M i)" shows "Pi UNIV E \ sets (\\<^sub>M i\UNIV::nat set. M i)" proof - have "Pi UNIV E = (\i. prod_emb UNIV M {..i} (\\<^sub>E j\{..i}. E j))" using E E[THEN sets.sets_into_space] by (auto simp: prod_emb_def Pi_iff extensional_def) with E show ?thesis by auto qed subsection \Measurability\ text \There are two natural sigma-algebras on a product space: the borel sigma algebra, generated by open sets in the product, and the product sigma algebra, countably generated by products of measurable sets along finitely many coordinates. The second one is defined and studied in \<^file>\Finite_Product_Measure.thy\. These sigma-algebra share a lot of natural properties (measurability of coordinates, for instance), but there is a fundamental difference: open sets are generated by arbitrary unions, not only countable ones, so typically many open sets will not be measurable with respect to the product sigma algebra (while all sets in the product sigma algebra are borel). The two sigma algebras coincide only when everything is countable (i.e., the product is countable, and the borel sigma algebra in the factor is countably generated). In this paragraph, we develop basic measurability properties for the borel sigma algebra, and compare it with the product sigma algebra as explained above. \ lemma measurable_product_coordinates [measurable (raw)]: "(\x. x i) \ measurable borel borel" by (rule borel_measurable_continuous_onI[OF continuous_on_product_coordinates]) lemma measurable_product_then_coordinatewise: fixes f::"'a \ 'b \ ('c::topological_space)" assumes [measurable]: "f \ borel_measurable M" shows "(\x. f x i) \ borel_measurable M" proof - have "(\x. f x i) = (\y. y i) o f" unfolding comp_def by auto then show ?thesis by simp qed text \To compare the Borel sigma algebra with the product sigma algebra, we give a presentation of the product sigma algebra that is more similar to the one we used above for the product topology.\ lemma sets_PiM_finite: "sets (Pi\<^sub>M I M) = sigma_sets (\\<^sub>E i\I. space (M i)) {(\\<^sub>E i\I. X i) |X. (\i. X i \ sets (M i)) \ finite {i. X i \ space (M i)}}" proof have "{(\\<^sub>E i\I. X i) |X. (\i. X i \ sets (M i)) \ finite {i. X i \ space (M i)}} \ sets (Pi\<^sub>M I M)" proof (auto) fix X assume H: "\i. X i \ sets (M i)" "finite {i. X i \ space (M i)}" then have *: "X i \ sets (M i)" for i by simp define J where "J = {i \ I. X i \ space (M i)}" have "finite J" "J \ I" unfolding J_def using H by auto define Y where "Y = (\\<^sub>E j\J. X j)" have "prod_emb I M J Y \ sets (Pi\<^sub>M I M)" unfolding Y_def apply (rule sets_PiM_I) using \finite J\ \J \ I\ * by auto moreover have "prod_emb I M J Y = (\\<^sub>E i\I. X i)" unfolding prod_emb_def Y_def J_def using H sets.sets_into_space[OF *] by (auto simp add: PiE_iff, blast) ultimately show "Pi\<^sub>E I X \ sets (Pi\<^sub>M I M)" by simp qed then show "sigma_sets (\\<^sub>E i\I. space (M i)) {(\\<^sub>E i\I. X i) |X. (\i. X i \ sets (M i)) \ finite {i. X i \ space (M i)}} \ sets (Pi\<^sub>M I M)" by (metis (mono_tags, lifting) sets.sigma_sets_subset' sets.top space_PiM) have *: "\X. {f. (\i\I. f i \ space (M i)) \ f \ extensional I \ f i \ A} = Pi\<^sub>E I X \ (\i. X i \ sets (M i)) \ finite {i. X i \ space (M i)}" if "i \ I" "A \ sets (M i)" for i A proof - define X where "X = (\j. if j = i then A else space (M j))" have "{f. (\i\I. f i \ space (M i)) \ f \ extensional I \ f i \ A} = Pi\<^sub>E I X" unfolding X_def using sets.sets_into_space[OF \A \ sets (M i)\] \i \ I\ by (auto simp add: PiE_iff extensional_def, metis subsetCE, metis) moreover have "X j \ sets (M j)" for j unfolding X_def using \A \ sets (M i)\ by auto moreover have "finite {j. X j \ space (M j)}" unfolding X_def by simp ultimately show ?thesis by auto qed show "sets (Pi\<^sub>M I M) \ sigma_sets (\\<^sub>E i\I. space (M i)) {(\\<^sub>E i\I. X i) |X. (\i. X i \ sets (M i)) \ finite {i. X i \ space (M i)}}" unfolding sets_PiM_single apply (rule sigma_sets_mono') apply (auto simp add: PiE_iff *) done qed lemma sets_PiM_subset_borel: "sets (Pi\<^sub>M UNIV (\_. borel)) \ sets borel" proof - have *: "Pi\<^sub>E UNIV X \ sets borel" if [measurable]: "\i. X i \ sets borel" "finite {i. X i \ UNIV}" for X::"'a \ 'b set" proof - define I where "I = {i. X i \ UNIV}" have "finite I" unfolding I_def using that by simp have "Pi\<^sub>E UNIV X = (\i\I. (\x. x i)-`(X i) \ space borel) \ space borel" unfolding I_def by auto also have "... \ sets borel" using that \finite I\ by measurable finally show ?thesis by simp qed then have "{(\\<^sub>E i\UNIV. X i) |X::('a \ 'b set). (\i. X i \ sets borel) \ finite {i. X i \ space borel}} \ sets borel" by auto then show ?thesis unfolding sets_PiM_finite space_borel by (simp add: * sets.sigma_sets_subset') qed proposition sets_PiM_equal_borel: "sets (Pi\<^sub>M UNIV (\i::('a::countable). borel::('b::second_countable_topology measure))) = sets borel" proof obtain K::"('a \ 'b) set set" where K: "topological_basis K" "countable K" "\k. k \ K \ \X. (k = Pi\<^sub>E UNIV X) \ (\i. open (X i)) \ finite {i. X i \ UNIV}" using product_topology_countable_basis by fast have *: "k \ sets (Pi\<^sub>M UNIV (\_. borel))" if "k \ K" for k proof - obtain X where H: "k = Pi\<^sub>E UNIV X" "\i. open (X i)" "finite {i. X i \ UNIV}" using K(3)[OF \k \ K\] by blast show ?thesis unfolding H(1) sets_PiM_finite space_borel using borel_open[OF H(2)] H(3) by auto qed have **: "U \ sets (Pi\<^sub>M UNIV (\_. borel))" if "open U" for U::"('a \ 'b) set" proof - obtain B where "B \ K" "U = (\B)" using \open U\ \topological_basis K\ by (metis topological_basis_def) have "countable B" using \B \ K\ \countable K\ countable_subset by blast moreover have "k \ sets (Pi\<^sub>M UNIV (\_. borel))" if "k \ B" for k using \B \ K\ * that by auto ultimately show ?thesis unfolding \U = (\B)\ by auto qed have "sigma_sets UNIV (Collect open) \ sets (Pi\<^sub>M UNIV (\i::'a. (borel::('b measure))))" apply (rule sets.sigma_sets_subset') using ** by auto then show "sets (borel::('a \ 'b) measure) \ sets (Pi\<^sub>M UNIV (\_. borel))" unfolding borel_def by auto qed (simp add: sets_PiM_subset_borel) lemma measurable_coordinatewise_then_product: fixes f::"'a \ ('b::countable) \ ('c::second_countable_topology)" assumes [measurable]: "\i. (\x. f x i) \ borel_measurable M" shows "f \ borel_measurable M" proof - have "f \ measurable M (Pi\<^sub>M UNIV (\_. borel))" by (rule measurable_PiM_single', auto simp add: assms) then show ?thesis using sets_PiM_equal_borel measurable_cong_sets by blast qed end diff --git a/src/HOL/Analysis/Gamma_Function.thy b/src/HOL/Analysis/Gamma_Function.thy --- a/src/HOL/Analysis/Gamma_Function.thy +++ b/src/HOL/Analysis/Gamma_Function.thy @@ -1,3542 +1,3556 @@ (* Title: HOL/Analysis/Gamma_Function.thy Author: Manuel Eberl, TU München *) section \The Gamma Function\ theory Gamma_Function imports Equivalence_Lebesgue_Henstock_Integration Summation_Tests Harmonic_Numbers "HOL-Library.Nonpos_Ints" "HOL-Library.Periodic_Fun" begin text \ Several equivalent definitions of the Gamma function and its most important properties. Also contains the definition and some properties of the log-Gamma function and the Digamma function and the other Polygamma functions. Based on the Gamma function, we also prove the Weierstra{\ss} product form of the sin function and, based on this, the solution of the Basel problem (the sum over all \<^term>\1 / (n::nat)^2\. \ lemma pochhammer_eq_0_imp_nonpos_Int: "pochhammer (x::'a::field_char_0) n = 0 \ x \ \\<^sub>\\<^sub>0" by (auto simp: pochhammer_eq_0_iff) lemma closed_nonpos_Ints [simp]: "closed (\\<^sub>\\<^sub>0 :: 'a :: real_normed_algebra_1 set)" proof - have "\\<^sub>\\<^sub>0 = (of_int ` {n. n \ 0} :: 'a set)" by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int) also have "closed \" by (rule closed_of_int_image) finally show ?thesis . qed lemma plus_one_in_nonpos_Ints_imp: "z + 1 \ \\<^sub>\\<^sub>0 \ z \ \\<^sub>\\<^sub>0" using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all lemma of_int_in_nonpos_Ints_iff: "(of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ 0" by (auto simp: nonpos_Ints_def) lemma one_plus_of_int_in_nonpos_Ints_iff: "(1 + of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ -1" proof - have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp also have "\ \ \\<^sub>\\<^sub>0 \ n + 1 \ 0" by (subst of_int_in_nonpos_Ints_iff) simp_all also have "\ \ n \ -1" by presburger finally show ?thesis . qed lemma one_minus_of_nat_in_nonpos_Ints_iff: "(1 - of_nat n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n > 0" proof - have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp also have "\ \ \\<^sub>\\<^sub>0 \ n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger finally show ?thesis . qed lemma fraction_not_in_ints: assumes "\(n dvd m)" "n \ 0" shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" proof assume "of_int m / (of_int n :: 'a) \ \" then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases) with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: field_split_simps) hence "m = k * n" by (subst (asm) of_int_eq_iff) hence "n dvd m" by simp with assms(1) show False by contradiction qed lemma fraction_not_in_nats: assumes "\n dvd m" "n \ 0" shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" proof assume "of_int m / of_int n \ (\ :: 'a set)" also note Nats_subset_Ints finally have "of_int m / of_int n \ (\ :: 'a set)" . moreover have "of_int m / of_int n \ (\ :: 'a set)" using assms by (intro fraction_not_in_ints) ultimately show False by contradiction qed lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \ \ \ z \ \\<^sub>\\<^sub>0" by (auto simp: Ints_def nonpos_Ints_def) lemma double_in_nonpos_Ints_imp: assumes "2 * (z :: 'a :: field_char_0) \ \\<^sub>\\<^sub>0" shows "z \ \\<^sub>\\<^sub>0 \ z + 1/2 \ \\<^sub>\\<^sub>0" proof- from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases') thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps) qed lemma sin_series: "(\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" proof - from sin_converges[of z] have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z" . also have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z \ (\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" by (subst sums_mono_reindex[of "\n. 2*n+1", symmetric]) (auto simp: sin_coeff_def strict_mono_def ac_simps elim!: oddE) finally show ?thesis . qed lemma cos_series: "(\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" proof - from cos_converges[of z] have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z" . also have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z \ (\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" by (subst sums_mono_reindex[of "\n. 2*n", symmetric]) (auto simp: cos_coeff_def strict_mono_def ac_simps elim!: evenE) finally show ?thesis . qed lemma sin_z_over_z_series: fixes z :: "'a :: {real_normed_field,banach}" assumes "z \ 0" shows "(\n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)" proof - from sin_series[of z] have "(\n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z" by (simp add: field_simps scaleR_conv_of_real) from sums_mult[OF this, of "inverse z"] and assms show ?thesis by (simp add: field_simps) qed lemma sin_z_over_z_series': fixes z :: "'a :: {real_normed_field,banach}" assumes "z \ 0" shows "(\n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)" proof - from sums_split_initial_segment[OF sin_converges[of z], of 1] have "(\n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps) qed lemma has_field_derivative_sin_z_over_z: fixes A :: "'a :: {real_normed_field,banach} set" shows "((\z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)" (is "(?f has_field_derivative ?f') _") proof (rule has_field_derivative_at_within) have "((\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) has_field_derivative (\n. diffs (\n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)" proof (rule termdiffs_strong) from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1] show "summable (\n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def) qed simp also have "(\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) = ?f" proof fix z show "(\n. of_real (sin_coeff (n+1)) * z^n) = ?f z" by (cases "z = 0") (insert sin_z_over_z_series'[of z], simp_all add: scaleR_conv_of_real sums_iff sin_coeff_def) qed also have "(\n. diffs (\n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) = diffs (\n. of_real (sin_coeff (Suc n))) 0" by simp also have "\ = 0" by (simp add: sin_coeff_def diffs_def) finally show "((\z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" . qed lemma round_Re_minimises_norm: "norm ((z::complex) - of_int m) \ norm (z - of_int (round (Re z)))" proof - let ?n = "round (Re z)" have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)" by (simp add: cmod_def) also have "\Re z - of_int ?n\ \ \Re z - of_int m\" by (rule round_diff_minimal) hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \ sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)" by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff) also have "\ = norm (z - of_int m)" by (simp add: cmod_def) finally show ?thesis . qed lemma Re_pos_in_ball: assumes "Re z > 0" "t \ ball z (Re z/2)" shows "Re t > 0" proof - have "Re (z - t) \ norm (z - t)" by (rule complex_Re_le_cmod) also from assms have "\ < Re z / 2" by (simp add: dist_complex_def) finally show "Re t > 0" using assms by simp qed lemma no_nonpos_Int_in_ball_complex: assumes "Re z > 0" "t \ ball z (Re z/2)" shows "t \ \\<^sub>\\<^sub>0" using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases) lemma no_nonpos_Int_in_ball: assumes "t \ ball z (dist z (round (Re z)))" shows "t \ \\<^sub>\\<^sub>0" proof assume "t \ \\<^sub>\\<^sub>0" then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases) have "dist z (of_int n) \ dist z t + dist t (of_int n)" by (rule dist_triangle) also from assms have "dist z t < dist z (round (Re z))" by simp also have "\ \ dist z (of_int n)" using round_Re_minimises_norm[of z] by (simp add: dist_complex_def) finally have "dist t (of_int n) > 0" by simp with \t = of_int n\ show False by simp qed lemma no_nonpos_Int_in_ball': assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \ \\<^sub>\\<^sub>0" obtains d where "d > 0" "\t. t \ ball z d \ t \ \\<^sub>\\<^sub>0" proof (rule that) from assms show "setdist {z} \\<^sub>\\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto next fix t assume "t \ ball z (setdist {z} \\<^sub>\\<^sub>0)" thus "t \ \\<^sub>\\<^sub>0" using setdist_le_dist[of z "{z}" t "\\<^sub>\\<^sub>0"] by force qed lemma no_nonpos_Real_in_ball: assumes z: "z \ \\<^sub>\\<^sub>0" and t: "t \ ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" shows "t \ \\<^sub>\\<^sub>0" using z proof (cases "Im z = 0") assume A: "Im z = 0" with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff) with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff) next assume A: "Im z \ 0" have "abs (Im z) - abs (Im t) \ abs (Im z - Im t)" by linarith also have "\ = abs (Im (z - t))" by simp also have "\ \ norm (z - t)" by (rule abs_Im_le_cmod) also from A t have "\ \ abs (Im z) / 2" by (simp add: dist_complex_def) finally have "abs (Im t) > 0" using A by simp thus ?thesis by (force simp add: complex_nonpos_Reals_iff) qed subsection \The Euler form and the logarithmic Gamma function\ text \ We define the Gamma function by first defining its multiplicative inverse \rGamma\. This is more convenient because \rGamma\ is entire, which makes proofs of its properties more convenient because one does not have to watch out for discontinuities. (e.g. \rGamma\ fulfils \rGamma z = z * rGamma (z + 1)\ everywhere, whereas the \\\ function does not fulfil the analogous equation on the non-positive integers) We define the \\\ function (resp.\ its reciprocale) in the Euler form. This form has the advantage that it is a relatively simple limit that converges everywhere. The limit at the poles is 0 (due to division by 0). The functional equation \Gamma (z + 1) = z * Gamma z\ follows immediately from the definition. \ definition\<^marker>\tag important\ Gamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)" definition Gamma_series' :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n" definition rGamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))" lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)" and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)" unfolding Gamma_series_def rGamma_series_def by simp_all lemma rGamma_series_minus_of_nat: "eventually (\n. rGamma_series (- of_nat k) n = 0) sequentially" using eventually_ge_at_top[of k] by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff) lemma Gamma_series_minus_of_nat: "eventually (\n. Gamma_series (- of_nat k) n = 0) sequentially" using eventually_ge_at_top[of k] by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff) lemma Gamma_series'_minus_of_nat: "eventually (\n. Gamma_series' (- of_nat k) n = 0) sequentially" using eventually_gt_at_top[of k] by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff) lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ rGamma_series z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp) lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp) lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series' z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp) lemma Gamma_series_Gamma_series': assumes z: "z \ \\<^sub>\\<^sub>0" shows "(\n. Gamma_series' z n / Gamma_series z n) \ 1" proof (rule Lim_transform_eventually) from eventually_gt_at_top[of "0::nat"] show "eventually (\n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n" by (cases n, simp) (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec' dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp) also from n have "\ = z / of_nat n + 1" by (simp add: field_split_simps) finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" .. qed have "(\x. z / of_nat x) \ 0" by (rule tendsto_norm_zero_cancel) (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n], simp add: norm_divide inverse_eq_divide) from tendsto_add[OF this tendsto_const[of 1]] show "(\n. z / of_nat n + 1) \ 1" by simp qed text \ We now show that the series that defines the \\\ function in the Euler form converges and that the function defined by it is continuous on the complex halfspace with positive real part. We do this by showing that the logarithm of the Euler series is continuous and converges locally uniformly, which means that the log-Gamma function defined by its limit is also continuous. This will later allow us to lift holomorphicity and continuity from the log-Gamma function to the inverse of the Gamma function, and from that to the Gamma function itself. \ definition\<^marker>\tag important\ ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\k=1..n. ln (z / of_nat k + 1))" definition\<^marker>\tag unimportant\ ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series' z n = - euler_mascheroni*z - ln z + (\k=1..n. z / of_nat n - ln (z / of_nat k + 1))" definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \ 'a" where "ln_Gamma z = lim (ln_Gamma_series z)" text \ We now show that the log-Gamma series converges locally uniformly for all complex numbers except the non-positive integers. We do this by proving that the series is locally Cauchy. \ context begin private lemma ln_Gamma_series_complex_converges_aux: fixes z :: complex and k :: nat assumes z: "z \ 0" and k: "of_nat k \ 2*norm z" "k \ 2" shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \ 2*(norm z + norm z^2) / of_nat k^2" proof - let ?k = "of_nat k :: complex" and ?z = "norm z" have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)" by (simp add: algebra_simps) also have "norm ... \ ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)" by (subst norm_mult [symmetric], rule norm_triangle_ineq) also have "norm (Ln (1 + -1/?k) - (-1/?k)) \ (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))" using k by (intro Ln_approx_linear) (simp add: norm_divide) hence "?z * norm (ln (1-1/?k) + 1/?k) \ ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))" by (intro mult_left_mono) simp_all also have "... \ (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k by (simp add: field_simps power2_eq_square norm_divide) also have "... \ (?z * 2) / of_nat k^2" using k by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) also have "norm (ln (1+z/?k) - z/?k) \ norm (z/?k)^2 / (1 - norm (z/?k))" using k by (intro Ln_approx_linear) (simp add: norm_divide) hence "norm (ln (1+z/?k) - z/?k) \ ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)" by (simp add: field_simps norm_divide) also have "... \ (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k by (simp add: field_simps power2_eq_square) also have "... \ (?z^2 * 2) / of_nat k^2" using k by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) also note add_divide_distrib [symmetric] finally show ?thesis by (simp only: distrib_left mult.commute) qed lemma ln_Gamma_series_complex_converges: assumes z: "z \ \\<^sub>\\<^sub>0" assumes d: "d > 0" "\n. n \ \\<^sub>\\<^sub>0 \ norm (z - of_int n) > d" shows "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n :: complex)" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI') fix e :: real assume e: "e > 0" define e'' where "e'' = (SUP t\ball z d. norm t + norm t^2)" define e' where "e' = e / (2*e'')" have "bounded ((\t. norm t + norm t^2) ` cball z d)" by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) hence "bounded ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto hence bdd: "bdd_above ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above) with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z]) have e'': "norm t + norm t^2 \ e''" if "t \ ball z d" for t unfolding e''_def using that by (rule cSUP_upper[OF _ bdd]) from e z e''_pos have e': "e' > 0" unfolding e'_def by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps) have "summable (\k. inverse ((real_of_nat k)^2))" by (rule inverse_power_summable) simp - from summable_partial_sum_bound[OF this e'] guess M . note M = this + from summable_partial_sum_bound[OF this e'] + obtain M where M: "\m n. M \ m \ norm (\k = m..n. inverse ((real k)\<^sup>2)) < e'" + by auto define N where "N = max 2 (max (nat \2 * (norm z + d)\) M)" { from d have "\2 * (cmod z + d)\ \ \0::real\" by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all hence "2 * (norm z + d) \ of_nat (nat \2 * (norm z + d)\)" unfolding N_def by (simp_all) also have "... \ of_nat N" unfolding N_def by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1) finally have "of_nat N \ 2 * (norm z + d)" . moreover have "N \ 2" "N \ M" unfolding N_def by simp_all moreover have "(\k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \ N" for m n using M[OF order.trans[OF \N \ M\ that]] unfolding real_norm_def by (subst (asm) abs_of_nonneg) (auto intro: sum_nonneg simp: field_split_simps) moreover note calculation } note N = this show "\M. \t\ball z d. \m\M. \n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e" unfolding dist_complex_def proof (intro exI[of _ N] ballI allI impI) fix t m n assume t: "t \ ball z d" and mn: "m \ N" "n > m" from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def) also have "dist z 0 - dist z t \ dist 0 t" using dist_triangle[of 0 z t] by (simp add: dist_commute) finally have t_nz: "t \ 0" by auto have "norm t \ norm z + norm (t - z)" by (rule norm_triangle_sub) also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute) also have "2 * (norm z + d) \ of_nat N" by (rule N) also have "N \ m" by (rule mn) finally have norm_t: "2 * norm t < of_nat m" by simp have "ln_Gamma_series t m - ln_Gamma_series t n = (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) + ((\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)))" by (simp add: ln_Gamma_series_def algebra_simps) also have "(\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)) = (\k\{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn by (simp add: sum_diff) also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) = (\k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn by (subst sum_telescope'' [symmetric]) simp_all also have "... = (\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N by (intro sum_cong_Suc) (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat) also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \ {Suc m..n}" for k using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: field_split_simps) hence "(\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) = (\k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N by (intro sum.cong) simp_all also note sum.distrib [symmetric] also have "norm (\k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \ (\k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t by (intro order.trans[OF norm_sum sum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all also have "... \ 2 * (norm t + norm t^2) * (\k=Suc m..n. 1 / (of_nat k)\<^sup>2)" by (simp add: sum_distrib_left) also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')" by (simp add: e'_def field_simps power2_eq_square) also from e''[OF t] e''_pos e have "\ \ e * 1" by (intro mult_left_mono) (simp_all add: field_simps) finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp qed qed end lemma ln_Gamma_series_complex_converges': assumes z: "(z :: complex) \ \\<^sub>\\<^sub>0" shows "\d>0. uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" proof - define d' where "d' = Re z" define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)" have "of_int (round d') \ \\<^sub>\\<^sub>0" if "d' \ 0" using that by (intro nonpos_Ints_of_int) (simp_all add: round_def) with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less) have "d < cmod (z - of_int n)" if "n \ \\<^sub>\\<^sub>0" for n proof (cases "Re z > 0") case True from nonpos_Ints_nonpos[OF that] have n: "n \ 0" by simp from True have "d = Re z/2" by (simp add: d_def d'_def) also from n True have "\ < Re (z - of_int n)" by simp also have "\ \ norm (z - of_int n)" by (rule complex_Re_le_cmod) finally show ?thesis . next case False with assms nonpos_Ints_of_int[of "round (Re z)"] have "z \ of_int (round d')" by (auto simp: not_less) with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def) also have "\ \ norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm) finally show ?thesis . qed hence conv: "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" by (intro ln_Gamma_series_complex_converges d_pos z) simp_all from d_pos conv show ?thesis by blast qed lemma ln_Gamma_series_complex_converges'': "(z :: complex) \ \\<^sub>\\<^sub>0 \ convergent (ln_Gamma_series z)" by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent) theorem ln_Gamma_complex_LIMSEQ: "(z :: complex) \ \\<^sub>\\<^sub>0 \ ln_Gamma_series z \ ln_Gamma z" using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def) lemma exp_ln_Gamma_series_complex: assumes "n > 0" "z \ \\<^sub>\\<^sub>0" shows "exp (ln_Gamma_series z n :: complex) = Gamma_series z n" proof - from assms obtain m where m: "n = Suc m" by (cases n) blast from assms have "z \ 0" by (intro notI) auto with assms have "exp (ln_Gamma_series z n) = (of_nat n) powr z / (z * (\k=1..n. exp (Ln (z / of_nat k + 1))))" unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_sum) also from assms have "(\k=1..n. exp (Ln (z / of_nat k + 1))) = (\k=1..n. z / of_nat k + 1)" by (intro prod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp) also have "... = (\k=1..n. z + k) / fact n" by (simp add: fact_prod) (subst prod_dividef [symmetric], simp_all add: field_simps) also from m have "z * ... = (\k=0..n. z + k) / fact n" by (simp add: prod.atLeast0_atMost_Suc_shift prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) also have "(\k=0..n. z + k) = pochhammer z (Suc n)" unfolding pochhammer_prod by (simp add: prod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n" unfolding Gamma_series_def using assms by (simp add: field_split_simps powr_def) finally show ?thesis . qed lemma ln_Gamma_series'_aux: assumes "(z::complex) \ \\<^sub>\\<^sub>0" shows "(\k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s") unfolding sums_def proof (rule Lim_transform) show "(\n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \ ?s" (is "?g \ _") by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms) have A: "eventually (\n. (\k 0" have "(\kk=1..n. z / of_nat k - ln (1 + z / of_nat k))" by (subst atLeast0LessThan [symmetric], subst sum.shift_bounds_Suc_ivl [symmetric], subst atLeastLessThanSuc_atLeastAtMost) simp_all also have "\ = z * of_real (harm n) - (\k=1..n. ln (1 + z / of_nat k))" by (simp add: harm_def sum_subtractf sum_distrib_left divide_inverse) also from n have "\ - ?g n = 0" by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps) finally show "(\kn. (\k 0" by (subst tendsto_cong[OF A]) simp_all qed lemma uniformly_summable_deriv_ln_Gamma: assumes z: "(z :: 'a :: {real_normed_field,banach}) \ 0" and d: "d > 0" "d \ norm z/2" shows "uniformly_convergent_on (ball z d) (\k z. \ik z. \i ball z d" have "norm z = norm (t + (z - t))" by simp have "norm (t + (z - t)) \ norm t + norm (z - t)" by (rule norm_triangle_ineq) also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm) finally have A: "norm t > norm z / 2" using z by (simp add: field_simps) have "norm t = norm (z + (t - z))" by simp also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) also from t d have "norm (t - z) \ norm z / 2" by (simp add: dist_norm norm_minus_commute) also from z have "\ < norm z" by simp finally have B: "norm t < 2 * norm z" by simp note A B } note ball = this show "eventually (\n. \t\ball z d. norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially" using eventually_gt_at_top apply eventually_elim proof safe fix t :: 'a assume t: "t \ ball z d" from z ball[OF t] have t_nz: "t \ 0" by auto fix n :: nat assume n: "n > nat \4 * norm z\" from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp also from n have "\ < of_nat n" by linarith finally have n: "of_nat n > 2 * norm t" . hence "of_nat n > norm t" by simp hence t': "t \ -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc) with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))" by (simp add: field_split_simps eq_neg_iff_add_eq_0 del: of_nat_Suc) also have "norm \ = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))" by (simp add: norm_divide norm_mult field_split_simps del: of_nat_Suc) also { from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \ of_nat (Suc n) / (2 * norm t)" by (intro divide_left_mono mult_pos_pos) simp_all also have "\ < norm (of_nat (Suc n) / t) - norm (1 :: 'a)" using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc) also have "\ \ norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq) finally have "inverse (norm (of_nat (Suc n)/t + 1)) \ 4 * norm z / of_nat (Suc n)" using z by (simp add: field_split_simps norm_divide mult_ac del: of_nat_Suc) } also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) = 4 * norm z * inverse (of_nat (Suc n)^2)" by (simp add: field_split_simps power2_eq_square del: of_nat_Suc) finally show "norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)" by (simp del: of_nat_Suc) qed next show "summable (\n. 4 * norm z * inverse ((of_nat (Suc n))^2))" by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable) qed subsection \The Polygamma functions\ lemma summable_deriv_ln_Gamma: "z \ (0 :: 'a :: {real_normed_field,banach}) \ summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))" unfolding summable_iff_convergent by (rule uniformly_convergent_imp_convergent, rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all definition\<^marker>\tag important\ Polygamma :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" where "Polygamma n z = (if n = 0 then (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else (-1)^Suc n * fact n * (\k. inverse ((z + of_nat k)^Suc n)))" abbreviation\<^marker>\tag important\ Digamma :: "('a :: {real_normed_field,banach}) \ 'a" where "Digamma \ Polygamma 0" lemma Digamma_def: "Digamma z = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni" by (simp add: Polygamma_def) lemma summable_Digamma: assumes "(z :: 'a :: {real_normed_field,banach}) \ 0" shows "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" proof - have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums (0 - inverse (z + of_nat 0))" by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]] show "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp qed lemma summable_offset: assumes "summable (\n. f (n + k) :: 'a :: real_normed_vector)" shows "summable f" proof - from assms have "convergent (\m. \nm. (\nnm. (\nnm. \n {k..n\\. f n) = (\nn=k..n=k..n=0..nnna. sum f {.. lim (\m. sum f {.. 0" and n: "n \ 2" shows "uniformly_convergent_on (ball z d) (\k z. \inorm z * e\" { fix t assume t: "t \ ball z d" have "norm t = norm (z + (t - z))" by simp also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute) finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def) } note ball = this show "eventually (\k. \t\ball z d. norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)) sequentially" using eventually_gt_at_top[of m] apply eventually_elim proof (intro ballI) fix k :: nat and t :: 'a assume k: "k > m" and t: "t \ ball z d" from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff) also have "\ \ norm (of_nat k :: 'a) - norm z * e" unfolding m_def by (subst norm_of_nat) linarith also from ball[OF t] have "\ \ norm (of_nat k :: 'a) - norm t" by simp also have "\ \ norm (of_nat k + t)" by (rule norm_diff_ineq) finally have "inverse ((norm (t + of_nat k))^n) \ inverse (real_of_nat (k - m)^n)" using k n by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc) thus "norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)" by (simp add: norm_inverse norm_power power_inverse) qed have "summable (\k. inverse ((real_of_nat k)^n))" using inverse_power_summable[of n] n by simp hence "summable (\k. inverse ((real_of_nat (k + m - m))^n))" by simp thus "summable (\k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset) qed lemma Polygamma_converges': fixes z :: "'a :: {real_normed_field,banach}" assumes z: "z \ 0" and n: "n \ 2" shows "summable (\k. inverse ((z + of_nat k)^n))" using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z] by (simp add: summable_iff_convergent) theorem Digamma_LIMSEQ: fixes z :: "'a :: {banach,real_normed_field}" assumes z: "z \ 0" shows "(\m. of_real (ln (real m)) - (\n Digamma z" proof - have "(\n. of_real (ln (real n / (real (Suc n))))) \ (of_real (ln 1) :: 'a)" by (intro tendsto_intros LIMSEQ_n_over_Suc_n) simp_all hence "(\n. of_real (ln (real n / (real n + 1)))) \ (0 :: 'a)" by (simp add: add_ac) hence lim: "(\n. of_real (ln (real n)) - of_real (ln (real n + 1))) \ (0::'a)" proof (rule Lim_transform_eventually) show "eventually (\n. of_real (ln (real n / (real n + 1))) = of_real (ln (real n)) - (of_real (ln (real n + 1)) :: 'a)) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_div) qed from summable_Digamma[OF z] have "(\n. inverse (of_nat (n+1)) - inverse (z + of_nat n)) sums (Digamma z + euler_mascheroni)" by (simp add: Digamma_def summable_sums) from sums_diff[OF this euler_mascheroni_sum] have "(\n. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)) - inverse (z + of_nat n)) sums Digamma z" by (simp add: add_ac) hence "(\m. (\nn Digamma z" by (simp add: sums_def sum_subtractf) also have "(\m. (\nm. of_real (ln (m + 1)) :: 'a)" by (subst sum_lessThan_telescope) simp_all finally show ?thesis by (rule Lim_transform) (insert lim, simp) qed theorem Polygamma_LIMSEQ: fixes z :: "'a :: {banach,real_normed_field}" assumes "z \ 0" and "n > 0" shows "(\k. inverse ((z + of_nat k)^Suc n)) sums ((-1) ^ Suc n * Polygamma n z / fact n)" using Polygamma_converges'[OF assms(1), of "Suc n"] assms(2) by (simp add: sums_iff Polygamma_def) theorem has_field_derivative_ln_Gamma_complex [derivative_intros]: fixes z :: complex assumes z: "z \ \\<^sub>\\<^sub>0" shows "(ln_Gamma has_field_derivative Digamma z) (at z)" proof - have not_nonpos_Int [simp]: "t \ \\<^sub>\\<^sub>0" if "Re t > 0" for t using that by (auto elim!: nonpos_Ints_cases') from z have z': "z \ \\<^sub>\\<^sub>0" and z'': "z \ 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I by blast+ let ?f' = "\z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))" let ?f = "\z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\z. \n. ?f' z n" define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" from z have d: "d > 0" "norm z/2 \ d" by (auto simp add: complex_nonpos_Reals_iff d_def) have ball: "Im t = 0 \ Re t > 0" if "dist z t < d" for t using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff) have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums (0 - inverse (z + of_nat 0))" by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) have "((\z. \n. ?f z n) has_field_derivative ?F' z) (at z)" using d z ln_Gamma_series'_aux[OF z'] apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma) apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff simp del: of_nat_Suc) apply (auto simp add: complex_nonpos_Reals_iff) done with z have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative ?F' z - euler_mascheroni - inverse z) (at z)" by (force intro!: derivative_eq_intros simp: Digamma_def) also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp also from sums have "-inverse z = (\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))" by (simp add: sums_iff) also from sums summable_deriv_ln_Gamma[OF z''] have "?F' z + \ = (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by (subst suminf_add) (simp_all add: add_ac sums_iff) also have "\ - euler_mascheroni = Digamma z" by (simp add: Digamma_def) finally have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative Digamma z) (at z)" . moreover from eventually_nhds_ball[OF d(1), of z] have "eventually (\z. ln_Gamma z = (\k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)" proof eventually_elim fix t assume "t \ ball z d" hence "t \ \\<^sub>\\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases) from ln_Gamma_series'_aux[OF this] show "ln_Gamma t = (\k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff) qed ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) qed declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros] lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni" by (simp add: Digamma_def) lemma Digamma_plus1: assumes "z \ 0" shows "Digamma (z+1) = Digamma z + 1/z" proof - have sums: "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) sums (inverse (z + of_nat 0) - 0)" by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) have "Digamma (z+1) = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) - euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac) also have "suminf ?f = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + (\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))" using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff) also have "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z" using sums by (simp add: sums_iff inverse_eq_divide) finally show ?thesis by (simp add: Digamma_def[of z]) qed theorem Polygamma_plus1: assumes "z \ 0" shows "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" proof (cases "n = 0") assume n: "n \ 0" let ?f = "\k. inverse ((z + of_nat k) ^ Suc n)" have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\k. ?f (k+1))" using n by (simp add: Polygamma_def add_ac) also have "(\k. ?f (k+1)) + (\k<1. ?f k) = (\k. ?f k)" using Polygamma_converges'[OF assms, of "Suc n"] n by (subst suminf_split_initial_segment [symmetric]) simp_all hence "(\k. ?f (k+1)) = (\k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps) also have "(-1) ^ Suc n * fact n * ((\k. ?f k) - inverse (z ^ Suc n)) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n by (simp add: inverse_eq_divide algebra_simps Polygamma_def) finally show ?thesis . qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide) theorem Digamma_of_nat: "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni" proof (induction n) case (Suc n) have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp also have "\ = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))" by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc) also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc) also have "\ + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni" by (simp add: harm_Suc) finally show ?case . qed (simp add: harm_def) lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni" by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl) lemma Polygamma_of_real: "x \ 0 \ Polygamma n (of_real x) = of_real (Polygamma n x)" unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"] by (simp_all add: suminf_of_real) lemma Polygamma_Real: "z \ \ \ z \ 0 \ Polygamma n z \ \" by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all lemma Digamma_half_integer: "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) = (\kk. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni" by (simp add: Digamma_def add_ac) also have "(\k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) = (\k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))" by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide) also have "\ = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums'] by (subst suminf_mult) (simp_all add: algebra_simps sums_iff) finally show ?case by simp next case (Suc n) have nz: "2 * of_nat n + (1:: 'a) \ 0" using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac) hence nz': "of_nat n + (1/2::'a) \ 0" by (simp add: field_simps) have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp also from nz' have "\ = Digamma (of_nat n + 1/2) + 1 / (of_nat n + 1/2)" by (rule Digamma_plus1) also from nz nz' have "1 / (of_nat n + 1/2 :: 'a) = 2 / (2 * of_nat n + 1)" by (subst divide_eq_eq) simp_all also note Suc finally show ?case by (simp add: add_ac) qed lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)" using Digamma_half_integer[of 0] by simp lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0" proof - have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp also have "\ = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp also note euler_mascheroni_less_13_over_22 also note ln2_le_25_over_36 finally show ?thesis by simp qed theorem has_field_derivative_Polygamma [derivative_intros]: fixes z :: "'a :: {real_normed_field,euclidean_space}" assumes z: "z \ \\<^sub>\\<^sub>0" shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)" proof (rule has_field_derivative_at_within, cases "n = 0") assume n: "n = 0" let ?f = "\k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)" let ?F = "\z. \k. ?f k z" and ?f' = "\k z. inverse ((z + of_nat k)\<^sup>2)" - from no_nonpos_Int_in_ball'[OF z] guess d . note d = this + from no_nonpos_Int_in_ball'[OF z] obtain d where d: "0 < d" "\t. t \ ball z d \ t \ \\<^sub>\\<^sub>0" + by auto from z have summable: "summable (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))" by (intro summable_Digamma) force from z have conv: "uniformly_convergent_on (ball z d) (\k z. \i2))" by (intro Polygamma_converges) auto with d have "summable (\k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent ) have "(?F has_field_derivative (\k. ?f' k z)) (at z)" proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) fix k :: nat and t :: 'a assume t: "t \ ball z d" from t d(2)[of t] show "((\z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)" by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases) qed (insert d(1) summable conv, (assumption|simp)+) with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n by (force simp: power2_eq_square intro!: derivative_eq_intros) next assume n: "n \ 0" from z have z': "z \ 0" by auto - from no_nonpos_Int_in_ball'[OF z] guess d . note d = this + from no_nonpos_Int_in_ball'[OF z] obtain d where d: "0 < d" "\t. t \ ball z d \ t \ \\<^sub>\\<^sub>0" + by auto define n' where "n' = Suc n" from n have n': "n' \ 2" by (simp add: n'_def) have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative (\k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)" proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) fix k :: nat and t :: 'a assume t: "t \ ball z d" with d have t': "t \ \\<^sub>\\<^sub>0" "t \ 0" by auto show "((\a. inverse ((a + of_nat k) ^ n')) has_field_derivative - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t' by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp) next have "uniformly_convergent_on (ball z d) (\k z. (- of_nat n' :: 'a) * (\ik z. \ik. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) = (- of_nat n') * (\k. inverse ((z + of_nat k) ^ (n' + 1)))" using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all finally have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative - of_nat n' * (\k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" . from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"] show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps) qed declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros] lemma isCont_Polygamma [continuous_intros]: fixes f :: "_ \ 'a :: {real_normed_field,euclidean_space}" shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Polygamma n (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Polygamma]]) lemma continuous_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A (Polygamma n :: _ \ 'a :: {real_normed_field,euclidean_space})" by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast lemma isCont_ln_Gamma_complex [continuous_intros]: fixes f :: "'a::t2_space \ complex" shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\z. ln_Gamma (f z)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]]) lemma continuous_on_ln_Gamma_complex [continuous_intros]: fixes A :: "complex set" shows "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A ln_Gamma" by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident]) fastforce lemma deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "deriv (Polygamma m) z = Polygamma (Suc m) (z :: 'a :: {real_normed_field,euclidean_space})" by (intro DERIV_imp_deriv has_field_derivative_Polygamma assms) thm has_field_derivative_Polygamma lemma higher_deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "(deriv ^^ n) (Polygamma m) z = Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})" proof - have "eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)" proof (induction n) case (Suc n) from Suc.IH have "eventually (\z. eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)" by (simp add: eventually_eventually) hence "eventually (\z. deriv ((deriv ^^ n) (Polygamma m)) z = deriv (Polygamma (m + n)) z) (nhds z)" by eventually_elim (intro deriv_cong_ev refl) moreover have "eventually (\z. z \ UNIV - \\<^sub>\\<^sub>0) (nhds z)" using assms by (intro eventually_nhds_in_open open_Diff open_UNIV) auto ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma) qed simp_all thus ?thesis by (rule eventually_nhds_x_imp_x) qed lemma deriv_ln_Gamma_complex: assumes "z \ \\<^sub>\\<^sub>0" shows "deriv ln_Gamma z = Digamma (z :: complex)" by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_complex assms) lemma higher_deriv_ln_Gamma_complex: assumes "(x::complex) \ \\<^sub>\\<^sub>0" shows "(deriv ^^ j) ln_Gamma x = (if j = 0 then ln_Gamma x else Polygamma (j - 1) x)" proof (cases j) case (Suc j') have "(deriv ^^ j') (deriv ln_Gamma) x = (deriv ^^ j') Digamma x" using eventually_nhds_in_open[of "UNIV - \\<^sub>\\<^sub>0" x] assms by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_Diff deriv_ln_Gamma_complex) also have "\ = Polygamma j' x" using assms by (subst higher_deriv_Polygamma) (auto elim!: nonpos_Ints_cases simp: complex_nonpos_Reals_iff) finally show ?thesis using Suc by (simp del: funpow.simps add: funpow_Suc_right) qed simp_all text \ We define a type class that captures all the fundamental properties of the inverse of the Gamma function and defines the Gamma function upon that. This allows us to instantiate the type class both for the reals and for the complex numbers with a minimal amount of proof duplication. \ class\<^marker>\tag unimportant\ Gamma = real_normed_field + complete_space + fixes rGamma :: "'a \ 'a" assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \ (\n. z = - of_nat n)" assumes differentiable_rGamma_aux1: "(\n. z \ - of_nat n) \ let d = (THE d. (\n. \k d) - scaleR euler_mascheroni 1 in filterlim (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R norm (y - z)) (nhds 0) (at z)" assumes differentiable_rGamma_aux2: "let z = - of_nat n in filterlim (\y. (rGamma y - rGamma z - (-1)^n * (prod of_nat {1..n}) * (y - z)) /\<^sub>R norm (y - z)) (nhds 0) (at z)" assumes rGamma_series_aux: "(\n. z \ - of_nat n) \ let fact' = (\n. prod of_nat {1..n}); exp = (\x. THE e. (\n. \kR fact k) \ e); pochhammer' = (\a n. (\n = 0..n. a + of_nat n)) in filterlim (\n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1)))) (nhds (rGamma z)) sequentially" begin subclass banach .. end definition "Gamma z = inverse (rGamma z)" subsection \Basic properties\ lemma Gamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ Gamma z = 0" and rGamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ rGamma z = 0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma Gamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ Gamma z \ 0" and rGamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ rGamma z \ 0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma Gamma_eq_zero_iff: "Gamma z = 0 \ z \ \\<^sub>\\<^sub>0" and rGamma_eq_zero_iff: "rGamma z = 0 \ z \ \\<^sub>\\<^sub>0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)" unfolding Gamma_def by simp lemma rGamma_series_LIMSEQ [tendsto_intros]: "rGamma_series z \ rGamma z" proof (cases "z \ \\<^sub>\\<^sub>0") case False hence "z \ - of_nat n" for n by auto from rGamma_series_aux[OF this] show ?thesis by (simp add: rGamma_series_def[abs_def] fact_prod pochhammer_Suc_prod exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost) qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ) theorem Gamma_series_LIMSEQ [tendsto_intros]: "Gamma_series z \ Gamma z" proof (cases "z \ \\<^sub>\\<^sub>0") case False hence "(\n. inverse (rGamma_series z n)) \ inverse (rGamma z)" by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff) also have "(\n. inverse (rGamma_series z n)) = Gamma_series z" by (simp add: rGamma_series_def Gamma_series_def[abs_def]) finally show ?thesis by (simp add: Gamma_def) qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ) lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)" using Gamma_series_LIMSEQ[of z] by (simp add: limI) lemma rGamma_1 [simp]: "rGamma 1 = 1" proof - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact field_split_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) moreover have "rGamma_series 1 \ rGamma 1" by (rule tendsto_intros) ultimately show ?thesis by (intro LIMSEQ_unique) qed lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z" proof - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) also from n have "\ = ?f n * rGamma_series z n" by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def) finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. qed moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" by (intro tendsto_intros lim_inverse_n) hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" by (blast intro: Lim_transform_eventually) moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" by (intro tendsto_intros) ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast qed lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)" proof (induction n arbitrary: z) case (Suc n z) have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH) also note rGamma_plus1 [symmetric] finally show ?case by (simp add: add_ac pochhammer_rec') qed simp_all theorem Gamma_plus1: "z \ \\<^sub>\\<^sub>0 \ Gamma (z + 1) = z * Gamma z" using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff) theorem pochhammer_Gamma: "z \ \\<^sub>\\<^sub>0 \ pochhammer z n = Gamma (z + of_nat n) / Gamma z" using pochhammer_rGamma[of z] by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps) lemma Gamma_0 [simp]: "Gamma 0 = 0" and rGamma_0 [simp]: "rGamma 0 = 0" and Gamma_neg_1 [simp]: "Gamma (- 1) = 0" and rGamma_neg_1 [simp]: "rGamma (- 1) = 0" and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0" and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0" and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0" and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0" by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff) lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp theorem Gamma_fact: "Gamma (1 + of_nat n) = fact n" by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff flip: of_nat_Suc) lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)" by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst of_nat_Suc, subst Gamma_fact) (rule refl) lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)" proof (cases "n > 0") case True hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int) lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)" by (simp add: Gamma_of_int rGamma_inverse_Gamma) lemma Gamma_seriesI: assumes "(\n. g n / Gamma_series z n) \ 1" shows "g \ Gamma z" proof (rule Lim_transform_eventually) have "1/2 > (0::real)" by simp from tendstoD[OF assms, OF this] show "eventually (\n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially" by (force elim!: eventually_mono simp: dist_real_def) from assms have "(\n. g n / Gamma_series z n * Gamma_series z n) \ 1 * Gamma z" by (intro tendsto_intros) thus "(\n. g n / Gamma_series z n * Gamma_series z n) \ Gamma z" by simp qed lemma Gamma_seriesI': assumes "f \ rGamma z" assumes "(\n. g n * f n) \ 1" assumes "z \ \\<^sub>\\<^sub>0" shows "g \ Gamma z" proof (rule Lim_transform_eventually) have "1/2 > (0::real)" by simp from tendstoD[OF assms(2), OF this] show "eventually (\n. g n * f n / f n = g n) sequentially" by (force elim!: eventually_mono simp: dist_real_def) from assms have "(\n. g n * f n / f n) \ 1 / rGamma z" by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff) thus "(\n. g n * f n / f n) \ Gamma z" by (simp add: Gamma_def divide_inverse) qed lemma Gamma_series'_LIMSEQ: "Gamma_series' z \ Gamma z" by (cases "z \ \\<^sub>\\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series'] Gamma_series'_nonpos_Ints_LIMSEQ[of z]) subsection \Differentiability\ lemma has_field_derivative_rGamma_no_nonpos_int: assumes "z \ \\<^sub>\\<^sub>0" shows "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)" proof (rule has_field_derivative_at_within) from assms have "z \ - of_nat n" for n by auto from differentiable_rGamma_aux1[OF this] show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)" unfolding Digamma_def suminf_def sums_def[abs_def] has_field_derivative_def has_derivative_def netlimit_at by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric]) qed lemma has_field_derivative_rGamma_nonpos_int: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)" apply (rule has_field_derivative_at_within) using differentiable_rGamma_aux2[of n] unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_prod) simp lemma has_field_derivative_rGamma [derivative_intros]: "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \norm z\) * fact (nat \norm z\) else -rGamma z * Digamma z)) (at z within A)" using has_field_derivative_rGamma_no_nonpos_int[of z A] has_field_derivative_rGamma_nonpos_int[of "nat \norm z\" A] by (auto elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros] declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros] declare has_field_derivative_rGamma_nonpos_int [derivative_intros] declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros] declare has_field_derivative_rGamma [derivative_intros] theorem has_field_derivative_Gamma [derivative_intros]: "z \ \\<^sub>\\<^sub>0 \ (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)" unfolding Gamma_def [abs_def] by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff) declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros] (* TODO: Hide ugly facts properly *) hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2 differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma" by (rule DERIV_continuous_on has_field_derivative_rGamma)+ lemma continuous_on_Gamma [continuous_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A Gamma" by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast lemma isCont_rGamma [continuous_intros]: "isCont f z \ isCont (\x. rGamma (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_rGamma]]) lemma isCont_Gamma [continuous_intros]: "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Gamma (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]]) subsection\<^marker>\tag unimportant\ \The complex Gamma function\ instantiation\<^marker>\tag unimportant\ complex :: Gamma begin definition\<^marker>\tag unimportant\ rGamma_complex :: "complex \ complex" where "rGamma_complex z = lim (rGamma_series z)" lemma rGamma_series_complex_converges: "convergent (rGamma_series (z :: complex))" (is "?thesis1") and rGamma_complex_altdef: "rGamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2") proof - have "?thesis1 \ ?thesis2" proof (cases "z \ \\<^sub>\\<^sub>0") case False have "rGamma_series z \ exp (- ln_Gamma z)" proof (rule Lim_transform_eventually) - from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE) + from ln_Gamma_series_complex_converges'[OF False] + obtain d where "0 < d" "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" + by auto from this(1) uniformly_convergent_imp_convergent[OF this(2), of z] have "ln_Gamma_series z \ lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff) thus "(\n. exp (-ln_Gamma_series z n)) \ exp (- ln_Gamma z)" unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus) from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False show "eventually (\n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially" by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def) qed with False show ?thesis by (auto simp: convergent_def rGamma_complex_def intro!: limI) next case True then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases') also have "rGamma_series \ \ 0" by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const) finally show ?thesis using True by (auto simp: rGamma_complex_def convergent_def intro!: limI) qed thus "?thesis1" "?thesis2" by blast+ qed context\<^marker>\tag unimportant\ begin (* TODO: duplication *) private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)" proof - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) also from n have "\ = ?f n * rGamma_series z n" by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def add_ac) finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. qed moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" using rGamma_series_complex_converges by (intro tendsto_intros lim_inverse_n) (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def) hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" by (blast intro: Lim_transform_eventually) moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" using rGamma_series_complex_converges by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff) ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast qed private lemma has_field_derivative_rGamma_complex_no_nonpos_Int: assumes "(z :: complex) \ \\<^sub>\\<^sub>0" shows "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" proof - have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z proof (subst DERIV_cong_ev[OF refl _ refl]) from that have "eventually (\t. t \ ball z (Re z/2)) (nhds z)" by (intro eventually_nhds_in_nhd) simp_all thus "eventually (\t. rGamma t = exp (- ln_Gamma t)) (nhds z)" using no_nonpos_Int_in_ball_complex[OF that] by (auto elim!: eventually_mono simp: rGamma_complex_altdef) next have "z \ \\<^sub>\\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff) with that show "((\t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)" by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef) qed from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" proof (induction "nat \1 - Re z\" arbitrary: z) case (Suc n z) from Suc.prems have z: "z \ 0" by auto from Suc.hyps have "n = nat \- Re z\" by linarith hence A: "n = nat \1 - Re (z + 1)\" by simp from Suc.prems have B: "z + 1 \ \\<^sub>\\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp) have "((\z. z * (rGamma \ (\z. z + 1)) z) has_field_derivative -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)" by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps) also have "(\z. z * (rGamma \ (\z. z + 1 :: complex)) z) = rGamma" by (simp add: rGamma_complex_plus1) also from z have "Digamma (z + 1) * z - 1 = z * Digamma z" by (subst Digamma_plus1) (simp_all add: field_simps) also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z" by (simp add: rGamma_complex_plus1[of z, symmetric]) finally show ?case . qed (intro diff, simp) qed private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1" proof - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact field_split_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI) qed private lemma has_field_derivative_rGamma_complex_nonpos_Int: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))" proof (induction n) case 0 have A: "(0::complex) + 1 \ \\<^sub>\\<^sub>0" by simp have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)" by (rule derivative_eq_intros DERIV_chain refl has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1) thus ?case by (simp add: rGamma_complex_plus1) next case (Suc n) hence A: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat (Suc n) + 1 :: complex))" by simp have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))" by (rule derivative_eq_intros refl A DERIV_chain)+ (simp add: algebra_simps rGamma_complex_altdef) thus ?case by (simp add: rGamma_complex_plus1) qed instance proof fix z :: complex show "(rGamma z = 0) \ (\n. z = - of_nat n)" by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases') next fix z :: complex assume "\n. z \ - of_nat n" hence "z \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases') from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this] show "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] of_real_def[symmetric] suminf_def) next fix n :: nat from has_field_derivative_rGamma_complex_nonpos_Int[of n] show "let z = - of_nat n in (\y. (rGamma y - rGamma z - (- 1) ^ n * prod of_nat {1..n} * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" by (simp add: has_field_derivative_def has_derivative_def fact_prod Let_def) next fix z :: complex from rGamma_series_complex_converges[of z] have "rGamma_series z \ rGamma z" by (simp add: convergent_LIMSEQ_iff rGamma_complex_def) thus "let fact' = \n. prod of_nat {1..n}; exp = \x. THE e. (\n. \kR fact k) \ e; pochhammer' = \a n. \n = 0..n. a + of_nat n in (\n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma z" by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) qed end end lemma Gamma_complex_altdef: "Gamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))" unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus) lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)" proof - have "rGamma_series (cnj z) = (\n. cnj (rGamma_series z n))" by (intro ext) (simp_all add: rGamma_series_def exp_cnj) also have "... \ cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros) finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI]) qed lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)" unfolding Gamma_def by (simp add: cnj_rGamma) lemma Gamma_complex_real: "z \ \ \ Gamma z \ (\ :: complex set)" and rGamma_complex_real: "z \ \ \ rGamma z \ \" by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma) lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)" using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast lemma holomorphic_rGamma [holomorphic_intros]: "rGamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma) lemma holomorphic_rGamma' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. rGamma (f x)) holomorphic_on A" proof - have "rGamma \ f holomorphic_on A" using assms by (intro holomorphic_on_compose assms holomorphic_rGamma) thus ?thesis by (simp only: o_def) qed lemma analytic_rGamma: "rGamma analytic_on A" unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_rGamma) lemma field_differentiable_Gamma: "z \ \\<^sub>\\<^sub>0 \ Gamma field_differentiable (at z within A)" using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto lemma holomorphic_Gamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma) lemma holomorphic_Gamma' [holomorphic_intros]: assumes "f holomorphic_on A" and "\x. x \ A \ f x \ \\<^sub>\\<^sub>0" shows "(\x. Gamma (f x)) holomorphic_on A" proof - have "Gamma \ f holomorphic_on A" using assms by (intro holomorphic_on_compose assms holomorphic_Gamma) auto thus ?thesis by (simp only: o_def) qed lemma analytic_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_Gamma) lemma field_differentiable_ln_Gamma_complex: "z \ \\<^sub>\\<^sub>0 \ ln_Gamma field_differentiable (at (z::complex) within A)" by (rule field_differentiable_within_subset[of _ _ UNIV]) (force simp: field_differentiable_def intro!: derivative_intros)+ lemma holomorphic_ln_Gamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ ln_Gamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_ln_Gamma_complex) lemma analytic_ln_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ ln_Gamma analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_ln_Gamma) lemma has_field_derivative_rGamma_complex' [derivative_intros]: "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \-Re z\) * fact (nat \-Re z\) else -rGamma z * Digamma z)) (at z within A)" using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros] lemma field_differentiable_Polygamma: fixes z :: complex shows "z \ \\<^sub>\\<^sub>0 \ Polygamma n field_differentiable (at z within A)" using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto lemma holomorphic_on_Polygamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma) lemma analytic_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_on_Polygamma) subsection\<^marker>\tag unimportant\ \The real Gamma function\ lemma rGamma_series_real: "eventually (\n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially" using eventually_gt_at_top[of "0 :: nat"] proof eventually_elim fix n :: nat assume n: "n > 0" have "Re (rGamma_series (of_real x) n) = Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))" using n by (simp add: rGamma_series_def powr_def pochhammer_of_real) also from n have "\ = Re (of_real ((pochhammer x (Suc n)) / (fact n * (exp (x * ln (real_of_nat n))))))" by (subst exp_of_real) simp also from n have "\ = rGamma_series x n" by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def) finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" .. qed instantiation\<^marker>\tag unimportant\ real :: Gamma begin definition "rGamma_real x = Re (rGamma (of_real x :: complex))" instance proof fix x :: real have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def) also have "of_real \ = rGamma (of_real x :: complex)" by (intro of_real_Re rGamma_complex_real) simp_all also have "\ = 0 \ x \ \\<^sub>\\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff) also have "\ \ (\n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases') finally show "(rGamma x) = 0 \ (\n. x = - real_of_nat n)" by simp next fix x :: real assume "\n. x \ - of_nat n" hence x: "complex_of_real x \ \\<^sub>\\<^sub>0" by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases') then have "x \ 0" by auto with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)" by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real rGamma_real_def [abs_def]) thus "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma x + rGamma x * d * (y - x)) /\<^sub>R norm (y - x)) \x\ 0" by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] of_real_def[symmetric] suminf_def) next fix n :: nat have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))" by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real rGamma_real_def [abs_def]) thus "let x = - of_nat n in (\y. (rGamma y - rGamma x - (- 1) ^ n * prod of_nat {1..n} * (y - x)) /\<^sub>R norm (y - x)) \x::real\ 0" by (simp add: has_field_derivative_def has_derivative_def fact_prod Let_def) next fix x :: real have "rGamma_series x \ rGamma x" proof (rule Lim_transform_eventually) show "(\n. Re (rGamma_series (of_real x) n)) \ rGamma x" unfolding rGamma_real_def by (intro tendsto_intros) qed (insert rGamma_series_real, simp add: eq_commute) thus "let fact' = \n. prod of_nat {1..n}; exp = \x. THE e. (\n. \kR fact k) \ e; pochhammer' = \a n. \n = 0..n. a + of_nat n in (\n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma x" by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) qed end lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)" unfolding rGamma_real_def using rGamma_complex_real by simp lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)" unfolding Gamma_def by (simp add: rGamma_complex_of_real) lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))" by (rule sym, rule limI, rule tendsto_intros) lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))" by (rule sym, rule limI, rule tendsto_intros) lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))" using rGamma_complex_real[OF Reals_of_real[of x]] by (elim Reals_cases) (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real) lemma ln_Gamma_series_complex_of_real: "x > 0 \ n > 0 \ ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)" proof - assume xn: "x > 0" "n > 0" have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \ 1" for k using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps) with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_real) qed lemma ln_Gamma_real_converges: assumes "(x::real) > 0" shows "convergent (ln_Gamma_series x)" proof - have "(\n. ln_Gamma_series (complex_of_real x) n) \ ln_Gamma (of_real x)" using assms by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff) moreover from eventually_gt_at_top[of "0::nat"] have "eventually (\n. complex_of_real (ln_Gamma_series x n) = ln_Gamma_series (complex_of_real x) n) sequentially" by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms) ultimately have "(\n. complex_of_real (ln_Gamma_series x n)) \ ln_Gamma (of_real x)" by (subst tendsto_cong) assumption+ from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def) qed lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \ ln_Gamma_series x \ ln_Gamma x" using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff) lemma ln_Gamma_complex_of_real: "x > 0 \ ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)" proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually) assume x: "x > 0" show "eventually (\n. of_real (ln_Gamma_series x n) = ln_Gamma_series (complex_of_real x) n) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_Gamma_series_complex_of_real x) qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def) lemma Gamma_real_pos_exp: "x > (0 :: real) \ Gamma x = exp (ln_Gamma x)" by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff ln_Gamma_complex_of_real exp_of_real) lemma ln_Gamma_real_pos: "x > 0 \ ln_Gamma x = ln (Gamma x :: real)" unfolding Gamma_real_pos_exp by simp lemma ln_Gamma_complex_conv_fact: "n > 0 \ ln_Gamma (of_nat n :: complex) = ln (fact (n - 1))" using ln_Gamma_complex_of_real[of "real n"] Gamma_fact[of "n - 1", where 'a = real] by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric]) lemma ln_Gamma_real_conv_fact: "n > 0 \ ln_Gamma (real n) = ln (fact (n - 1))" using Gamma_fact[of "n - 1", where 'a = real] by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric]) lemma Gamma_real_pos [simp, intro]: "x > (0::real) \ Gamma x > 0" by (simp add: Gamma_real_pos_exp) lemma Gamma_real_nonneg [simp, intro]: "x > (0::real) \ Gamma x \ 0" by (simp add: Gamma_real_pos_exp) lemma has_field_derivative_ln_Gamma_real [derivative_intros]: assumes x: "x > (0::real)" shows "(ln_Gamma has_field_derivative Digamma x) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms show "((Re \ ln_Gamma \ complex_of_real) has_field_derivative Digamma x) (at x)" by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real o_def) from eventually_nhds_in_nhd[of x "{0<..}"] assms show "eventually (\y. ln_Gamma y = (Re \ ln_Gamma \ of_real) y) (nhds x)" by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open) qed lemma field_differentiable_ln_Gamma_real: "x > 0 \ ln_Gamma field_differentiable (at (x::real) within A)" by (rule field_differentiable_within_subset[of _ _ UNIV]) (auto simp: field_differentiable_def intro!: derivative_intros)+ declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros] lemma deriv_ln_Gamma_real: assumes "z > 0" shows "deriv ln_Gamma z = Digamma (z :: real)" by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_real assms) lemma higher_deriv_ln_Gamma_real: assumes "(x::real) > 0" shows "(deriv ^^ j) ln_Gamma x = (if j = 0 then ln_Gamma x else Polygamma (j - 1) x)" proof (cases j) case (Suc j') have "(deriv ^^ j') (deriv ln_Gamma) x = (deriv ^^ j') Digamma x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_Diff deriv_ln_Gamma_real) also have "\ = Polygamma j' x" using assms by (subst higher_deriv_Polygamma) (auto elim!: nonpos_Ints_cases simp: complex_nonpos_Reals_iff) finally show ?thesis using Suc by (simp del: funpow.simps add: funpow_Suc_right) qed simp_all lemma higher_deriv_ln_Gamma_complex_of_real: assumes "(x :: real) > 0" shows "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using assms by (auto simp: higher_deriv_ln_Gamma_real higher_deriv_ln_Gamma_complex ln_Gamma_complex_of_real Polygamma_of_real) lemma has_field_derivative_rGamma_real' [derivative_intros]: "(rGamma has_field_derivative (if x \ \\<^sub>\\<^sub>0 then (-1)^(nat \-x\) * fact (nat \-x\) else -rGamma x * Digamma x)) (at x within A)" using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros] lemma Polygamma_real_odd_pos: assumes "(x::real) \ \\<^sub>\\<^sub>0" "odd n" shows "Polygamma n x > 0" proof - from assms have "x \ 0" by auto with assms show ?thesis unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] by (auto simp: zero_less_power_eq simp del: power_Suc dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos) qed lemma Polygamma_real_even_neg: assumes "(x::real) > 0" "n > 0" "even n" shows "Polygamma n x < 0" using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] by (auto intro!: mult_pos_pos suminf_pos) lemma Polygamma_real_strict_mono: assumes "x > 0" "x < (y::real)" "even n" shows "Polygamma n x < Polygamma n y" proof - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) - then guess \ by (elim exE conjE) note \ = this - note \(3) - also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ > 0" + then obtain \ + where \: "x < \" "\ < y" + and Polygamma: "Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" + by auto + note Polygamma + also from \ assms have "(y - x) * Polygamma (Suc n) \ > 0" by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases) finally show ?thesis by simp qed lemma Polygamma_real_strict_antimono: assumes "x > 0" "x < (y::real)" "odd n" shows "Polygamma n x > Polygamma n y" proof - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) - then guess \ by (elim exE conjE) note \ = this - note \(3) - also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ < 0" + then obtain \ + where \: "x < \" "\ < y" + and Polygamma: "Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" + by auto + note Polygamma + also from \ assms have "(y - x) * Polygamma (Suc n) \ < 0" by (intro mult_pos_neg Polygamma_real_even_neg) simp_all finally show ?thesis by simp qed lemma Polygamma_real_mono: assumes "x > 0" "x \ (y::real)" "even n" shows "Polygamma n x \ Polygamma n y" using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2) by (cases "x = y") simp_all lemma Digamma_real_strict_mono: "(0::real) < x \ x < y \ Digamma x < Digamma y" by (rule Polygamma_real_strict_mono) simp_all lemma Digamma_real_mono: "(0::real) < x \ x \ y \ Digamma x \ Digamma y" by (rule Polygamma_real_mono) simp_all lemma Digamma_real_ge_three_halves_pos: assumes "x \ 3/2" shows "Digamma (x :: real) > 0" proof - have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos) also from assms have "\ \ Digamma x" by (intro Polygamma_real_mono) simp_all finally show ?thesis . qed lemma ln_Gamma_real_strict_mono: assumes "x \ 3/2" "x < y" shows "ln_Gamma (x :: real) < ln_Gamma y" proof - have "\\. x < \ \ \ < y \ ln_Gamma y - ln_Gamma x = (y - x) * Digamma \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) - then guess \ by (elim exE conjE) note \ = this - note \(3) - also from \(1,2) assms have "(y - x) * Digamma \ > 0" + then obtain \ where \: "x < \" "\ < y" + and ln_Gamma: "ln_Gamma y - ln_Gamma x = (y - x) * Digamma \" + by auto + note ln_Gamma + also from \ assms have "(y - x) * Digamma \ > 0" by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all finally show ?thesis by simp qed lemma Gamma_real_strict_mono: assumes "x \ 3/2" "x < y" shows "Gamma (x :: real) < Gamma y" proof - from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp also have "\ < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms) also from Gamma_real_pos_exp[of y] assms have "\ = Gamma y" by simp finally show ?thesis . qed theorem log_convex_Gamma_real: "convex_on {0<..} (ln \ Gamma :: real \ real)" by (rule convex_on_realI[of _ _ Digamma]) (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases') subsection \The uniqueness of the real Gamma function\ text \ The following is a proof of the Bohr--Mollerup theorem, which states that any log-convex function \G\ on the positive reals that fulfils \G(1) = 1\ and satisfies the functional equation \G(x + 1) = x G(x)\ must be equal to the Gamma function. In principle, if \G\ is a holomorphic complex function, one could then extend this from the positive reals to the entire complex plane (minus the non-positive integers, where the Gamma function is not defined). \ context\<^marker>\tag unimportant\ fixes G :: "real \ real" assumes G_1: "G 1 = 1" assumes G_plus1: "x > 0 \ G (x + 1) = x * G x" assumes G_pos: "x > 0 \ G x > 0" assumes log_convex_G: "convex_on {0<..} (ln \ G)" begin private lemma G_fact: "G (of_nat n + 1) = fact n" using G_plus1[of "real n + 1" for n] by (induction n) (simp_all add: G_1 G_plus1) private definition S :: "real \ real \ real" where "S x y = (ln (G y) - ln (G x)) / (y - x)" private lemma S_eq: "n \ 2 \ S (of_nat n) (of_nat n + x) = (ln (G (real n + x)) - ln (fact (n - 1))) / x" by (subst G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff) private lemma G_lower: assumes x: "x > 0" and n: "n \ 1" shows "Gamma_series x n \ G x" proof - have "(ln \ G) (real (Suc n)) \ ((ln \ G) (real (Suc n) + x) - (ln \ G) (real (Suc n) - 1)) / (real (Suc n) + x - (real (Suc n) - 1)) * (real (Suc n) - (real (Suc n) - 1)) + (ln \ G) (real (Suc n) - 1)" using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto hence "S (of_nat n) (of_nat (Suc n)) \ S (of_nat (Suc n)) (of_nat (Suc n) + x)" unfolding S_def using x by (simp add: field_simps) also have "S (of_nat n) (of_nat (Suc n)) = ln (fact n) - ln (fact (n-1))" unfolding S_def using n by (subst (1 2) G_fact [symmetric]) (simp_all add: add_ac of_nat_diff) also have "\ = ln (fact n / fact (n-1))" by (subst ln_div) simp_all also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all finally have "x * ln (real n) + ln (fact n) \ ln (G (real (Suc n) + x))" using x n by (subst (asm) S_eq) (simp_all add: field_simps) also have "x * ln (real n) + ln (fact n) = ln (exp (x * ln (real n)) * fact n)" using x by (simp add: ln_mult) finally have "exp (x * ln (real n)) * fact n \ G (real (Suc n) + x)" using x by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos) also have "G (real (Suc n) + x) = pochhammer x (Suc n) * G x" using G_plus1[of "real (Suc n) + x" for n] G_plus1[of x] x by (induction n) (simp_all add: pochhammer_Suc add_ac) finally show "Gamma_series x n \ G x" using x by (simp add: field_simps pochhammer_pos Gamma_series_def) qed private lemma G_upper: assumes x: "x > 0" "x \ 1" and n: "n \ 2" shows "G x \ Gamma_series x n * (1 + x / real n)" proof - have "(ln \ G) (real n + x) \ ((ln \ G) (real n + 1) - (ln \ G) (real n)) / (real n + 1 - (real n)) * ((real n + x) - real n) + (ln \ G) (real n)" using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto hence "S (of_nat n) (of_nat n + x) \ S (of_nat n) (of_nat n + 1)" unfolding S_def using x by (simp add: field_simps) also from n have "S (of_nat n) (of_nat n + 1) = ln (fact n) - ln (fact (n-1))" by (subst (1 2) G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff) also have "\ = ln (fact n / (fact (n-1)))" using n by (subst ln_div) simp_all also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all finally have "ln (G (real n + x)) \ x * ln (real n) + ln (fact (n - 1))" using x n by (subst (asm) S_eq) (simp_all add: field_simps) also have "\ = ln (exp (x * ln (real n)) * fact (n - 1))" using x by (simp add: ln_mult) finally have "G (real n + x) \ exp (x * ln (real n)) * fact (n - 1)" using x by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos) also have "G (real n + x) = pochhammer x n * G x" using G_plus1[of "real n + x" for n] x by (induction n) (simp_all add: pochhammer_Suc add_ac) finally have "G x \ exp (x * ln (real n)) * fact (n- 1) / pochhammer x n" using x by (simp add: field_simps pochhammer_pos) also from n have "fact (n - 1) = fact n / n" by (cases n) simp_all also have "exp (x * ln (real n)) * \ / pochhammer x n = Gamma_series x n * (1 + x / real n)" using n x by (simp add: Gamma_series_def divide_simps pochhammer_Suc) finally show ?thesis . qed private lemma G_eq_Gamma_aux: assumes x: "x > 0" "x \ 1" shows "G x = Gamma x" proof (rule antisym) show "G x \ Gamma x" proof (rule tendsto_upperbound) from G_lower[of x] show "eventually (\n. Gamma_series x n \ G x) sequentially" using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "1::nat"]]) qed (simp_all add: Gamma_series_LIMSEQ) next show "G x \ Gamma x" proof (rule tendsto_lowerbound) have "(\n. Gamma_series x n * (1 + x / real n)) \ Gamma x * (1 + 0)" by (rule tendsto_intros real_tendsto_divide_at_top Gamma_series_LIMSEQ filterlim_real_sequentially)+ thus "(\n. Gamma_series x n * (1 + x / real n)) \ Gamma x" by simp next from G_upper[of x] show "eventually (\n. Gamma_series x n * (1 + x / real n) \ G x) sequentially" using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "2::nat"]]) qed simp_all qed theorem Gamma_pos_real_unique: assumes x: "x > 0" shows "G x = Gamma x" proof - have G_eq: "G (real n + x) = Gamma (real n + x)" if "x \ {0<..1}" for n x using that proof (induction n) case (Suc n) from Suc have "x + real n > 0" by simp hence "x + real n \ \\<^sub>\\<^sub>0" by auto with Suc show ?case using G_plus1[of "real n + x"] Gamma_plus1[of "real n + x"] by (auto simp: add_ac) qed (simp_all add: G_eq_Gamma_aux) show ?thesis proof (cases "frac x = 0") case True hence "x = of_int (floor x)" by (simp add: frac_def) with x have x_eq: "x = of_nat (nat (floor x) - 1) + 1" by simp show ?thesis by (subst (1 2) x_eq, rule G_eq) simp_all next case False from assms have x_eq: "x = of_nat (nat (floor x)) + frac x" by (simp add: frac_def) have frac_le_1: "frac x \ 1" unfolding frac_def by linarith show ?thesis by (subst (1 2) x_eq, rule G_eq, insert False frac_le_1) simp_all qed qed end subsection \The Beta function\ definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)" lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)" by (simp add: inverse_eq_divide Beta_def Gamma_def) lemma Beta_commute: "Beta a b = Beta b a" unfolding Beta_def by (simp add: ac_simps) lemma has_field_derivative_Beta1 [derivative_intros]: assumes "x \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" shows "((\x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y)))) (at x within A)" unfolding Beta_altdef by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps) lemma Beta_pole1: "x \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma Beta_pole2: "y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma Beta_zero: "x + y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma has_field_derivative_Beta2 [derivative_intros]: assumes "y \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" shows "((\y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y)))) (at y within A)" using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac) theorem Beta_plus1_plus1: assumes "x \ \\<^sub>\\<^sub>0" "y \ \\<^sub>\\<^sub>0" shows "Beta (x + 1) y + Beta x (y + 1) = Beta x y" proof - have "Beta (x + 1) y + Beta x (y + 1) = (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)" by (simp add: Beta_altdef add_divide_distrib algebra_simps) also have "\ = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))" by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps) also from assms have "\ = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp finally show ?thesis . qed theorem Beta_plus1_left: assumes "x \ \\<^sub>\\<^sub>0" shows "(x + y) * Beta (x + 1) y = x * Beta x y" proof - have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))" unfolding Beta_altdef by (simp only: ac_simps) also have "\ = x * Beta x y" unfolding Beta_altdef by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps) finally show ?thesis . qed theorem Beta_plus1_right: assumes "y \ \\<^sub>\\<^sub>0" shows "(x + y) * Beta x (y + 1) = y * Beta x y" using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute) lemma Gamma_Gamma_Beta: assumes "x + y \ \\<^sub>\\<^sub>0" shows "Gamma x * Gamma y = Beta x y * Gamma (x + y)" unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"] by (simp add: rGamma_inverse_Gamma) subsection \Legendre duplication theorem\ context begin private lemma Gamma_legendre_duplication_aux: fixes z :: "'a :: Gamma" assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)" proof - let ?powr = "\b a. exp (a * of_real (ln (of_nat b)))" let ?h = "\n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) * exp (1/2 * of_real (ln (real_of_nat n)))" { fix z :: 'a assume z: "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" let ?g = "\n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n / Gamma_series' (2*z) (2*n)" have "eventually (\n. ?g n = ?h n) sequentially" using eventually_gt_at_top proof eventually_elim fix n :: nat assume n: "n > 0" let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a" have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) / (pochhammer z n * pochhammer (z + 1/2) n)" by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac) have B: "Gamma_series' (2*z) (2*n) = ?f' * ?powr 2 (2*z) * ?powr n (2*z) / (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double) from z have "pochhammer z n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) moreover from z have "pochhammer (z + 1/2) n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) = ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))" using n unfolding A B by (simp add: field_split_simps exp_minus) also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)" by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib) finally show "?g n = ?h n" by (simp only: mult_ac) qed moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \ \\<^sub>\\<^sub>0" by auto hence "?g \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "(*)2" "2*z"] by (intro tendsto_intros Gamma_series'_LIMSEQ) (simp_all add: o_def strict_mono_def Gamma_eq_zero_iff) ultimately have "?h \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" by (blast intro: Lim_transform_eventually) } note lim = this from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \ \\<^sub>\\<^sub>0" by auto from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \ \\<^sub>\\<^sub>0" by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all with lim[of "1/2 :: 'a"] have "?h \ 2 * Gamma (1/2 :: 'a)" by (simp add: exp_of_real) from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis by (simp add: field_split_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real) qed text \ The following lemma is somewhat annoying. With a little bit of complex analysis (Cauchy's integral theorem, to be exact), this would be completely trivial. However, we want to avoid depending on the complex analysis session at this point, so we prove it the hard way. \ private lemma Gamma_reflection_aux: defines "h \ \z::complex. if z \ \ then 0 else (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" defines "a \ complex_of_real pi" obtains h' where "continuous_on UNIV h'" "\z. (h has_field_derivative (h' z)) (at z)" proof - define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z define g where "g n = complex_of_real (sin_coeff (n+1))" for n define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z have a_nz: "a \ 0" unfolding a_def by simp have "(\n. f n * (a*z)^n) sums (F z) \ (\n. g n * (a*z)^n) sums (G z)" if "abs (Re z) < 1" for z proof (cases "z = 0"; rule conjI) assume "z \ 0" note z = this that from z have sin_nz: "sin (a*z) \ 0" unfolding a_def by (auto simp: sin_eq_0) have "(\n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"] by (simp add: scaleR_conv_of_real) from sums_split_initial_segment[OF this, of 1] have "(\n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac) from sums_mult[OF this, of "inverse (a*z)"] z a_nz have A: "(\n. g n * (a*z)^n) sums (sin (a*z)/(a*z))" by (simp add: field_simps g_def) with z show "(\n. g n * (a*z)^n) sums (G z)" by (simp add: G_def) from A z a_nz sin_nz have g_nz: "(\n. g n * (a*z)^n) \ 0" by (simp add: sums_iff g_def) have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def) from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1] have "(\n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))" by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def) from sums_mult[OF this, of "inverse z"] z assms show "(\n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def) next assume z: "z = 0" have "(\n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp with z show "(\n. f n * (a * z) ^ n) sums (F z)" by (simp add: f_def F_def sin_coeff_def cos_coeff_def) have "(\n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp with z show "(\n. g n * (a * z) ^ n) sums (G z)" by (simp add: g_def G_def sin_coeff_def cos_coeff_def) qed note sums = conjunct1[OF this] conjunct2[OF this] define h2 where [abs_def]: "h2 z = (\n. f n * (a*z)^n) / (\n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z define POWSER where [abs_def]: "POWSER f z = (\n. f n * (z^n :: complex))" for f z define POWSER' where [abs_def]: "POWSER' f z = (\n. diffs f n * (z^n))" for f and z :: complex define h2' where [abs_def]: "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t proof - from that have t: "t \ \ \ t = 0" by (auto elim!: Ints_cases) hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)" unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def) also have "a*cot (a*t) - 1/t = (F t) / (G t)" using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def) also have "\ = (\n. f n * (a*t)^n) / (\n. g n * (a*t)^n)" using sums[of t] that by (simp add: sums_iff) finally show "h t = h2 t" by (simp only: h2_def) qed let ?A = "{z. abs (Re z) < 1}" have "open ({z. Re z < 1} \ {z. Re z > -1})" using open_halfspace_Re_gt open_halfspace_Re_lt by auto also have "({z. Re z < 1} \ {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto finally have open_A: "open ?A" . hence [simp]: "interior ?A = ?A" by (simp add: interior_open) have summable_f: "summable (\n. f n * z^n)" for z by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) (simp_all add: norm_mult a_def del: of_real_add) have summable_g: "summable (\n. g n * z^n)" for z by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) (simp_all add: norm_mult a_def del: of_real_add) have summable_fg': "summable (\n. diffs f n * z^n)" "summable (\n. diffs g n * z^n)" for z by (intro termdiff_converges_all summable_f summable_g)+ have "(POWSER f has_field_derivative (POWSER' f z)) (at z)" "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z unfolding POWSER_def POWSER'_def by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+ note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def] have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z" for z unfolding POWSER_def POWSER'_def by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+ note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def] { fix z :: complex assume z: "abs (Re z) < 1" define d where "d = \ * of_real (norm z + 1)" have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add) have "eventually (\z. h z = h2 z) (nhds z)" using eventually_nhds_in_nhd[of z ?A] using h_eq z by (auto elim!: eventually_mono) moreover from sums(2)[OF z] z have nz: "(\n. g n * (a * z) ^ n) \ 0" unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def) have A: "z \ \ \ z = 0" using z by (auto elim!: Ints_cases) have no_int: "1 + z \ \ \ z = 0" using z Ints_diff[of "1+z" 1] A by (auto elim!: nonpos_Ints_cases) have no_int': "1 - z \ \ \ z = 0" using z Ints_diff[of 1 "1-z"] A by (auto elim!: nonpos_Ints_cases) from no_int no_int' have no_int: "1 - z \ \\<^sub>\\<^sub>0" "1 + z \ \\<^sub>\\<^sub>0" by auto have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+) (auto simp: h2'_def POWSER_def field_simps power2_eq_square) ultimately have deriv: "(h has_field_derivative h2' z) (at z)" by (subst DERIV_cong_ev[OF refl _ refl]) from sums(2)[OF z] z have "(\n. g n * (a * z) ^ n) \ 0" unfolding G_def by (auto simp: sums_iff a_def sin_eq_0) hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def by (intro continuous_intros cont continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto note deriv and this } note A = this interpret h: periodic_fun_simple' h proof fix z :: complex show "h (z + 1) = h z" proof (cases "z \ \") assume z: "z \ \" hence A: "z + 1 \ \" "z \ 0" using Ints_diff[of "z+1" 1] by auto hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)" by (subst (1 2) Digamma_plus1) simp_all with A z show "h (z + 1) = h z" by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def) qed (simp add: h_def) qed have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z proof - have "((\z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)" by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) (insert z, auto intro!: derivative_eq_intros) hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1) moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique) qed define h2'' where "h2'' z = h2' (z - of_int \Re z\)" for z have deriv: "(h has_field_derivative h2'' z) (at z)" for z proof - fix z :: complex have B: "\Re z - real_of_int \Re z\\ < 1" by linarith have "((\t. h (t - of_int \Re z\)) has_field_derivative h2'' z) (at z)" unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) (insert B, auto intro!: derivative_intros) thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int) qed have cont: "continuous_on UNIV h2''" proof (intro continuous_at_imp_continuous_on ballI) fix z :: complex define r where "r = \Re z\" define A where "A = {t. of_int r - 1 < Re t \ Re t < of_int r + 1}" have "continuous_on A (\t. h2' (t - of_int r))" unfolding A_def by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros) (simp_all add: abs_real_def) moreover have "h2'' t = h2' (t - of_int r)" if t: "t \ A" for t proof (cases "Re t \ of_int r") case True from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) with True have "\Re t\ = \Re z\" unfolding r_def by linarith thus ?thesis by (auto simp: r_def h2''_def) next case False from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) with False have t': "\Re t\ = \Re z\ - 1" unfolding r_def by linarith moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)" by (intro h2'_eq) simp_all ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t') qed ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl]) moreover { have "open ({t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t})" by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt) also have "{t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t} = A" unfolding A_def by blast finally have "open A" . } ultimately have C: "isCont h2'' t" if "t \ A" for t using that by (subst (asm) continuous_on_eq_continuous_at) auto have "of_int r - 1 < Re z" "Re z < of_int r + 1" unfolding r_def by linarith+ thus "isCont h2'' z" by (intro C) (simp_all add: A_def) qed from that[OF cont deriv] show ?thesis . qed lemma Gamma_reflection_complex: fixes z :: complex shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)" proof - let ?g = "\z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)" define g where [abs_def]: "g z = (if z \ \ then of_real pi else ?g z)" for z :: complex let ?h = "\z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" define h where [abs_def]: "h z = (if z \ \ then 0 else ?h z)" for z :: complex \ \@{term g} is periodic with period 1.\ interpret g: periodic_fun_simple' g proof fix z :: complex show "g (z + 1) = g z" proof (cases "z \ \") case False hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def) also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)" using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints by (subst Beta_plus1_left [symmetric]) auto also have "\ * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))" using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi) also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)" using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def) finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto qed (simp add: g_def) qed \ \@{term g} is entire.\ have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex proof (cases "z \ \") let ?h' = "\z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) + of_real pi * cos (z * of_real pi))" case False from False have "eventually (\t. t \ UNIV - \) (nhds z)" by (intro eventually_nhds_in_open) (auto simp: open_Diff) hence "eventually (\t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def) moreover { from False Ints_diff[of 1 "1-z"] have "1 - z \ \" by auto hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def) also from False have "sin (of_real pi * z) \ 0" by (subst sin_eq_0) auto hence "?h' z = h z * g z" using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def) finally have "(?g has_field_derivative (h z * g z)) (at z)" . } ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) next case True then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases) let ?t = "(\z::complex. if z = 0 then 1 else sin z / z) \ (\z. of_real pi * z)" have deriv_0: "(g has_field_derivative 0) (at 0)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "eventually (\z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)" using eventually_nhds_ball[OF zero_less_one, of "0::complex"] proof eventually_elim fix z :: complex assume z: "z \ ball 0 1" show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z" proof (cases "z = 0") assume z': "z \ 0" with z have z'': "z \ \\<^sub>\\<^sub>0" "z \ \" by (auto elim!: Ints_cases) from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp with z'' z' show ?thesis by (simp add: g_def ac_simps) qed (simp add: g_def) qed have "(?t has_field_derivative (0 * of_real pi)) (at 0)" using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"] by (intro DERIV_chain) simp_all thus "((\z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)" by (auto intro!: derivative_eq_intros simp: o_def) qed have "((g \ (\x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))" using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros) also have "g \ (\x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int) finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def) qed have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z proof (cases "z \ \") case True with that have "z = 0 \ z = 1" by (force elim!: Ints_cases) moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square) moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square Beta_def algebra_simps) ultimately show ?thesis by force next case False hence z: "z/2 \ \" "(z+1)/2 \ \" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases) hence z': "z/2 \ \\<^sub>\\<^sub>0" "(z+1)/2 \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) from z have "1-z/2 \ \" "1-((z+1)/2) \ \" using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto hence z'': "1-z/2 \ \\<^sub>\\<^sub>0" "1-((z+1)/2) \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) from z have "g (z/2) * g ((z+1)/2) = (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) * (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))" by (simp add: g_def) also from z' Gamma_legendre_duplication_aux[of "z/2"] have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z" by (simp add: add_divide_distrib) also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"] have "Gamma (1-z/2) * Gamma (1-(z+1)/2) = Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))" by (simp add: add_divide_distrib ac_simps) finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) * (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))" by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real) also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)" using cos_sin_eq[of "- of_real pi * z/2", symmetric] by (simp add: ring_distribs add_divide_distrib ac_simps) also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)" by (subst sin_times_cos) (simp add: field_simps) also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z" using \z \ \\ by (simp add: g_def) finally show ?thesis . qed have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z proof - define r where "r = \Re z / 2\" have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int) also have "of_int (2*r) = 2 * of_int r" by simp also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+ hence "Gamma (1/2)^2 * g (z - 2 * of_int r) = g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)" unfolding r_def by (intro g_eq[symmetric]) simp_all also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp also have "g \ = g (z/2)" by (rule g.minus_of_int) also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp also have "g \ = g ((z+1)/2)" by (rule g.minus_of_int) finally show ?thesis .. qed have g_nz [simp]: "g z \ 0" for z :: complex unfolding g_def using Ints_diff[of 1 "1 - z"] by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int) have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z proof - have "((\t. g (t/2) * g ((t+1)/2)) has_field_derivative (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps) hence "((\t. Gamma (1/2)^2 * g t) has_field_derivative Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" by (subst (1 2) g_eq[symmetric]) simp from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"] have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints) moreover have "(g has_field_derivative (g z * h z)) (at z)" using g_g'[of z] by (simp add: ac_simps) ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)" by (intro DERIV_unique) thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp qed obtain h' where h'_cont: "continuous_on UNIV h'" and h_h': "\z. (h has_field_derivative h' z) (at z)" unfolding h_def by (erule Gamma_reflection_aux) have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z proof - have "((\t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)" by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2]) hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)" by (subst (asm) h_eq[symmetric]) from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique) qed have h'_zero: "h' z = 0" for z proof - define m where "m = max 1 \Re z\" define B where "B = {t. abs (Re t) \ m \ abs (Im t) \ abs (Im z)}" have "closed ({t. Re t \ -m} \ {t. Re t \ m} \ {t. Im t \ -\Im z\} \ {t. Im t \ \Im z\})" (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le closed_halfspace_Im_ge closed_halfspace_Im_le) also have "?B = B" unfolding B_def by fastforce finally have "closed B" . moreover have "bounded B" unfolding bounded_iff proof (intro ballI exI) fix t assume t: "t \ B" have "norm t \ \Re t\ + \Im t\" by (rule cmod_le) also from t have "\Re t\ \ m" unfolding B_def by blast also from t have "\Im t\ \ \Im z\" unfolding B_def by blast finally show "norm t \ m + \Im z\" by - simp qed ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast define M where "M = (SUP z\B. norm (h' z))" have "compact (h' ` B)" by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+ hence bdd: "bdd_above ((\z. norm (h' z)) ` B)" using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded) have "norm (h' z) \ M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def) also have "M \ M/2" proof (subst M_def, subst cSUP_le_iff) have "z \ B" unfolding B_def m_def by simp thus "B \ {}" by auto next show "\z\B. norm (h' z) \ M/2" proof fix t :: complex assume t: "t \ B" from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp) also have "norm \ = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp also have "norm (h' (t/2) + h' ((t+1)/2)) \ norm (h' (t/2)) + norm (h' ((t+1)/2))" by (rule norm_triangle_ineq) also from t have "abs (Re ((t + 1)/2)) \ m" unfolding m_def B_def by auto with t have "t/2 \ B" "(t+1)/2 \ B" unfolding B_def by auto hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \ M + M" unfolding M_def by (intro add_mono cSUP_upper bdd) (auto simp: B_def) also have "(M + M) / 4 = M / 2" by simp finally show "norm (h' t) \ M/2" by - simp_all qed qed (insert bdd, auto) hence "M \ 0" by simp finally show "h' z = 0" by simp qed have h_h'_2: "(h has_field_derivative 0) (at z)" for z using h_h'[of z] h'_zero[of z] by simp have g_real: "g z \ \" if "z \ \" for z unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real) have h_real: "h z \ \" if "z \ \" for z unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real) have g_nz: "g z \ 0" for z unfolding g_def using Ints_diff[of 1 "1-z"] by (auto simp: Gamma_eq_zero_iff sin_eq_0) from h'_zero h_h'_2 have "\c. \z\UNIV. h z = c" by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm) then obtain c where c: "\z. h z = c" by auto have "\u. u \ closed_segment 0 1 \ Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))" by (intro complex_mvt_line g_g') then obtain u where u: "u \ closed_segment 0 1" "Re (g 1) - Re (g 0) = Re (h u * g u)" by auto from u(1) have u': "u \ \" unfolding closed_segment_def by (auto simp: scaleR_conv_of_real) from u' g_real[of u] g_nz[of u] have "Re (g u) \ 0" by (auto elim!: Reals_cases) with u(2) c[of u] g_real[of u] g_nz[of u] u' have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1) with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases) with c have A: "h z * g z = 0" for z by simp hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp hence "\c'. \z\UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all then obtain c' where c: "\z. g z = c'" by (force) from this[of 0] have "c' = pi" unfolding g_def by simp with c have "g z = pi" by simp show ?thesis proof (cases "z \ \") case False with \g z = pi\ show ?thesis by (auto simp: g_def divide_simps) next case True then obtain n where n: "z = of_int n" by (elim Ints_cases) with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force moreover have "of_int (1 - n) \ \\<^sub>\\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp ultimately show ?thesis using n by (cases "n \ 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int) qed qed lemma rGamma_reflection_complex: "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi" using Gamma_reflection_complex[of z] by (simp add: Gamma_def field_split_simps split: if_split_asm) lemma rGamma_reflection_complex': "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi" proof - have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))" using rGamma_plus1[of "-z", symmetric] by simp also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi" by (rule rGamma_reflection_complex) finally show ?thesis by simp qed lemma Gamma_reflection_complex': "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))" using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def field_split_simps) lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi" proof - from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1] have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square) hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp also have "\ = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all moreover have "Gamma (1/2 :: real) \ 0" using Gamma_real_pos[of "1/2"] by simp ultimately show ?thesis by (rule real_sqrt_unique [symmetric]) qed lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)" proof - have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp also have "\ = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real) finally show ?thesis . qed theorem Gamma_legendre_duplication: fixes z :: complex assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)" using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex) end subsection\<^marker>\tag unimportant\ \Limits and residues\ text \ The inverse of the Gamma function has simple zeros: \ lemma rGamma_zeros: "(\z. rGamma z / (z + of_nat n)) \ (- of_nat n) \ ((-1)^n * fact n :: 'a :: Gamma)" proof (subst tendsto_cong) let ?f = "\z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a" from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] show "eventually (\z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))" by (subst pochhammer_rGamma[of _ "Suc n"]) (auto elim!: eventually_mono simp: field_split_simps pochhammer_rec' eq_neg_iff_add_eq_0) have "isCont ?f (- of_nat n)" by (intro continuous_intros) thus "?f \ (- of_nat n) \ (- 1) ^ n * fact n" unfolding isCont_def by (simp add: pochhammer_same) qed text \ The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function, and their residues can easily be computed from the limit we have just proven: \ lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))" proof - from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] have "eventually (\z. rGamma z \ (0 :: 'a)) (at (- of_nat n))" by (auto elim!: eventually_mono nonpos_Ints_cases' simp: rGamma_eq_zero_iff dist_of_nat dist_minus) with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident] have "filterlim (\z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))" unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity]) (simp_all add: filterlim_at) moreover have "(\z. inverse (rGamma z) :: 'a) = Gamma" by (intro ext) (simp add: rGamma_inverse_Gamma) ultimately show ?thesis by (simp only: ) qed lemma Gamma_residues: "(\z. Gamma z * (z + of_nat n)) \ (- of_nat n) \ ((-1)^n / fact n :: 'a :: Gamma)" proof (subst tendsto_cong) let ?c = "(- 1) ^ n / fact n :: 'a" from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] show "eventually (\z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n))) (at (- of_nat n))" by (auto elim!: eventually_mono simp: field_split_simps rGamma_inverse_Gamma) have "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ inverse ((- 1) ^ n * fact n :: 'a)" by (intro tendsto_intros rGamma_zeros) simp_all also have "inverse ((- 1) ^ n * fact n) = ?c" by (simp_all add: field_simps flip: power_mult_distrib) finally show "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ ?c" . qed subsection \Alternative definitions\ subsubsection \Variant of the Euler form\ definition Gamma_series_euler' where "Gamma_series_euler' z n = inverse z * (\k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))" context begin private lemma Gamma_euler'_aux1: fixes z :: "'a :: {real_normed_field,banach}" assumes n: "n > 0" shows "exp (z * of_real (ln (of_nat n + 1))) = (\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))" proof - have "(\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) = exp (z * of_real (\k = 1..n. ln (1 + 1 / real_of_nat k)))" by (subst exp_sum [symmetric]) (simp_all add: sum_distrib_left) also have "(\k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\k=1..n. 1 + 1 / real_of_nat k)" by (subst ln_prod [symmetric]) (auto intro!: add_pos_nonneg) also have "(\k=1..n. 1 + 1 / of_nat k :: real) = (\k=1..n. (of_nat k + 1) / of_nat k)" by (intro prod.cong) (simp_all add: field_split_simps) also have "(\k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1" by (induction n) (simp_all add: prod.nat_ivl_Suc' field_split_simps) finally show ?thesis .. qed theorem Gamma_series_euler': assumes z: "(z :: 'a :: Gamma) \ \\<^sub>\\<^sub>0" shows "(\n. Gamma_series_euler' z n) \ Gamma z" proof (rule Gamma_seriesI, rule Lim_transform_eventually) let ?f = "\n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)" let ?r = "\n. ?f n / Gamma_series z n" let ?r' = "\n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))" from z have z': "z \ 0" by auto have "eventually (\n. ?r' n = ?r n) sequentially" using z by (auto simp: field_split_simps Gamma_series_def ring_distribs exp_diff ln_div intro: eventually_mono eventually_gt_at_top[of "0::nat"] dest: pochhammer_eq_0_imp_nonpos_Int) moreover have "?r' \ exp (z * of_real (ln 1))" by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all ultimately show "?r \ 1" by (force intro: Lim_transform_eventually) from eventually_gt_at_top[of "0::nat"] show "eventually (\n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" from n z' have "Gamma_series_euler' z n = exp (z * of_real (ln (of_nat n + 1))) / (z * (\k=1..n. (1 + z / of_nat k)))" by (subst Gamma_euler'_aux1) (simp_all add: Gamma_series_euler'_def prod.distrib prod_inversef[symmetric] divide_inverse) also have "(\k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n" proof (cases n) case (Suc n') then show ?thesis unfolding pochhammer_prod fact_prod by (simp add: atLeastLessThanSuc_atLeastAtMost field_simps prod_dividef prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) qed auto also have "z * \ = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec) finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp qed qed end subsubsection \Weierstrass form\ definition Gamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where "Gamma_series_Weierstrass z n = exp (-euler_mascheroni * z) / z * (\k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))" definition\<^marker>\tag unimportant\ rGamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where "rGamma_series_Weierstrass z n = exp (euler_mascheroni * z) * z * (\k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))" lemma Gamma_series_Weierstrass_nonpos_Ints: "eventually (\k. Gamma_series_Weierstrass (- of_nat n) k = 0) sequentially" using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_Weierstrass_def) lemma rGamma_series_Weierstrass_nonpos_Ints: "eventually (\k. rGamma_series_Weierstrass (- of_nat n) k = 0) sequentially" using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_Weierstrass_def) theorem Gamma_Weierstrass_complex: "Gamma_series_Weierstrass z \ Gamma (z :: complex)" proof (cases "z \ \\<^sub>\\<^sub>0") case True then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') also from True have "Gamma_series_Weierstrass \ \ Gamma z" by (simp add: tendsto_cong[OF Gamma_series_Weierstrass_nonpos_Ints] Gamma_nonpos_Int) finally show ?thesis . next case False hence z: "z \ 0" by auto let ?f = "(\x. \x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))" have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \ 1" for n :: nat using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp) have "(\n. \k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \ ln_Gamma z + euler_mascheroni * z + ln z" using ln_Gamma_series'_aux[OF False] by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def sum.shift_bounds_Suc_ivl sums_def atLeast0LessThan) from tendsto_exp[OF this] False z have "?f \ z * exp (euler_mascheroni * z) * Gamma z" by (simp add: exp_add exp_sum exp_diff mult_ac Gamma_complex_altdef A) from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z show "Gamma_series_Weierstrass z \ Gamma z" by (simp add: exp_minus field_split_simps Gamma_series_Weierstrass_def [abs_def]) qed lemma tendsto_complex_of_real_iff: "((\x. complex_of_real (f x)) \ of_real c) F = (f \ c) F" by (rule tendsto_of_real_iff) lemma Gamma_Weierstrass_real: "Gamma_series_Weierstrass x \ Gamma (x :: real)" using Gamma_Weierstrass_complex[of "of_real x"] unfolding Gamma_series_Weierstrass_def[abs_def] by (subst tendsto_complex_of_real_iff [symmetric]) (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real) lemma rGamma_Weierstrass_complex: "rGamma_series_Weierstrass z \ rGamma (z :: complex)" proof (cases "z \ \\<^sub>\\<^sub>0") case True then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') also from True have "rGamma_series_Weierstrass \ \ rGamma z" by (simp add: tendsto_cong[OF rGamma_series_Weierstrass_nonpos_Ints] rGamma_nonpos_Int) finally show ?thesis . next case False have "rGamma_series_Weierstrass z = (\n. inverse (Gamma_series_Weierstrass z n))" by (simp add: rGamma_series_Weierstrass_def[abs_def] Gamma_series_Weierstrass_def exp_minus divide_inverse prod_inversef[symmetric] mult_ac) also from False have "\ \ inverse (Gamma z)" by (intro tendsto_intros Gamma_Weierstrass_complex) (simp add: Gamma_eq_zero_iff) finally show ?thesis by (simp add: Gamma_def) qed subsubsection \Binomial coefficient form\ lemma Gamma_gbinomial: "(\n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \ rGamma (z+1)" proof (cases "z = 0") case False show ?thesis proof (rule Lim_transform_eventually) let ?powr = "\a b. exp (b * of_real (ln (of_nat a)))" show "eventually (\n. rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially" proof (intro always_eventually allI) fix n :: nat from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n" by (simp add: gbinomial_pochhammer' pochhammer_rec) also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z" by (simp add: rGamma_series_def field_split_simps exp_minus) finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" .. qed from False have "(\n. rGamma_series z n / z) \ rGamma z / z" by (intro tendsto_intros) also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z] by (simp add: field_simps) finally show "(\n. rGamma_series z n / z) \ rGamma (z+1)" . qed qed (simp_all add: binomial_gbinomial [symmetric]) lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)" by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric]) lemma gbinomial_asymptotic: fixes z :: "'a :: Gamma" shows "(\n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \ inverse (Gamma (- z))" unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"] by (subst (asm) gbinomial_minus') (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric]) lemma fact_binomial_limit: "(\n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \ 1 / fact k" proof (rule Lim_transform_eventually) have "(\n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n)))) \ 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \ _") using Gamma_gbinomial[of "of_nat k :: 'a"] by (simp add: binomial_gbinomial Gamma_def field_split_simps exp_of_real [symmetric] exp_minus) also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact) finally show "?f \ 1 / fact k" . show "eventually (\n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)" by (simp add: exp_of_nat_mult) thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp qed qed lemma binomial_asymptotic': "(\n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \ 1" using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp lemma gbinomial_Beta: assumes "z + 1 \ \\<^sub>\\<^sub>0" shows "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))" using assms proof (induction n arbitrary: z) case 0 hence "z + 2 \ \\<^sub>\\<^sub>0" using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute) with 0 show ?case by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute) next case (Suc n z) show ?case proof (cases "z \ \\<^sub>\\<^sub>0") case True with Suc.prems have "z = 0" by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff) show ?thesis proof (cases "n = 0") case True with \z = 0\ show ?thesis by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric]) next case False with \z = 0\ show ?thesis by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff) qed next case False have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp also have "\ = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)" by (subst gbinomial_factors) (simp add: field_simps) also from False have "\ = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))" (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1) also have "of_nat (Suc n) \ (\\<^sub>\\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)" by (subst Beta_plus1_right [symmetric]) simp_all finally show ?thesis . qed qed theorem gbinomial_Gamma: assumes "z + 1 \ \\<^sub>\\<^sub>0" shows "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))" proof - have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))" by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac) also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)" using Gamma_plus1[of "z+1"] by (auto simp add: field_split_simps) finally show ?thesis . qed subsubsection \Integral form\ lemma integrable_on_powr_from_0': assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0<..c}" proof - from c have *: "{0<..c} - {0..c} = {}" "{0..c} - {0<..c} = {0}" by auto show ?thesis by (rule integrable_spike_set [OF integrable_on_powr_from_0[OF a c]]) (simp_all add: *) qed lemma absolutely_integrable_Gamma_integral: assumes "Re z > 0" "a > 0" shows "(\t. complex_of_real t powr (z - 1) / of_real (exp (a * t))) absolutely_integrable_on {0<..}" (is "?f absolutely_integrable_on _") proof - have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" by (intro tendsto_intros ln_x_over_x_tendsto_0) hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp from order_tendstoD(2)[OF this, of "a/2"] and \a > 0\ have "eventually (\x. (Re z - 1) * ln x / x < a/2) at_top" by simp from eventually_conj[OF this eventually_gt_at_top[of 0]] obtain x0 where "\x\x0. (Re z - 1) * ln x / x < a/2 \ x > 0" by (auto simp: eventually_at_top_linorder) hence "x0 > 0" by simp have "x powr (Re z - 1) / exp (a * x) < exp (-(a/2) * x)" if "x \ x0" for x proof - from that and \\x\x0. _\ have x: "(Re z - 1) * ln x / x < a / 2" "x > 0" by auto have "x powr (Re z - 1) = exp ((Re z - 1) * ln x)" using \x > 0\ by (simp add: powr_def) also from x have "(Re z - 1) * ln x < (a * x) / 2" by (simp add: field_simps) finally show ?thesis by (simp add: field_simps exp_add [symmetric]) qed note x0 = \x0 > 0\ this have "?f absolutely_integrable_on ({0<..x0} \ {x0..})" proof (rule set_integrable_Un) show "?f absolutely_integrable_on {0<..x0}" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2]) show "integrable lebesgue (\x. indicat_real {0<..x0} x *\<^sub>R x powr (Re z - 1))" using x0(1) assms by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_powr_from_0') auto show "(\x. indicat_real {0<..x0} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \ borel_measurable lebesgue" by (intro measurable_completion) (auto intro!: borel_measurable_continuous_on_indicator continuous_intros) fix x :: real have "x powr (Re z - 1) / exp (a * x) \ x powr (Re z - 1) / 1" if "x \ 0" using that assms by (intro divide_left_mono) auto thus "norm (indicator {0<..x0} x *\<^sub>R ?f x) \ norm (indicator {0<..x0} x *\<^sub>R x powr (Re z - 1))" by (simp_all add: norm_divide norm_powr_real_powr indicator_def) qed next show "?f absolutely_integrable_on {x0..}" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2]) show "integrable lebesgue (\x. indicat_real {x0..} x *\<^sub>R exp (- (a / 2) * x))" using assms by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_exp_minus_to_infinity) auto show "(\x. indicat_real {x0..} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \ borel_measurable lebesgue" using x0(1) by (intro measurable_completion) (auto intro!: borel_measurable_continuous_on_indicator continuous_intros) fix x :: real show "norm (indicator {x0..} x *\<^sub>R ?f x) \ norm (indicator {x0..} x *\<^sub>R exp (-(a/2) * x))" using x0 by (auto simp: norm_divide norm_powr_real_powr indicator_def less_imp_le) qed qed auto also have "{0<..x0} \ {x0..} = {0<..}" using x0(1) by auto finally show ?thesis . qed lemma integrable_Gamma_integral_bound: fixes a c :: real assumes a: "a > -1" and c: "c \ 0" defines "f \ \x. if x \ {0..c} then x powr a else exp (-x/2)" shows "f integrable_on {0..}" proof - have "f integrable_on {0..c}" by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]]) (insert a c, simp_all add: f_def) moreover have A: "(\x. exp (-x/2)) integrable_on {c..}" using integrable_on_exp_minus_to_infinity[of "1/2"] by simp have "f integrable_on {c..}" by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def) ultimately show "f integrable_on {0..}" by (rule integrable_Un') (insert c, auto simp: max_def) qed theorem Gamma_integral_complex: assumes z: "Re z > 0" shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" proof - have A: "((\t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n)) has_integral (fact n / pochhammer z (n+1))) {0..1}" if "Re z > 0" for n z using that proof (induction n arbitrary: z) case 0 have "((\t. complex_of_real t powr (z - 1)) has_integral (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0 by (intro fundamental_theorem_of_calculus_interior) (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_field) thus ?case by simp next case (Suc n) let ?f = "\t. complex_of_real t powr z / z" let ?f' = "\t. complex_of_real t powr (z - 1)" let ?g = "\t. (1 - complex_of_real t) ^ Suc n" let ?g' = "\t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)" have "((\t. ?f' t * ?g t) has_integral (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}" (is "(_ has_integral ?I) _") proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g]) from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g" by (auto intro!: continuous_intros) next fix t :: real assume t: "t \ {0<..<1}" show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems by (auto intro!: derivative_eq_intros has_vector_derivative_real_field) show "(?g has_vector_derivative ?g' t) (at t)" by (rule has_vector_derivative_real_field derivative_eq_intros refl)+ simp_all next from Suc.prems have [simp]: "z \ 0" by auto from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp have [simp]: "z + of_nat n \ 0" "z + 1 + of_nat n \ 0" for n using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel) have "((\x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}" (is "(?A has_integral ?B) _") using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec) also have "?A = (\t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps) also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))" by (simp add: field_split_simps pochhammer_rec prod.shift_bounds_cl_Suc_ivl del: of_nat_Suc) finally show "((\t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}" by simp qed (simp_all add: bounded_bilinear_mult) thus ?case by simp qed have B: "((\t. if t \ {0..of_nat n} then of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0) has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n proof (cases "n > 0") case [simp]: True hence [simp]: "n \ 0" by auto with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0] have "((\x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n) has_integral fact n * of_nat n / pochhammer z (n+1)) ((\x. real n * x)`{0..1})" (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real) also from True have "((\x. real n*x)`{0..1}) = {0..real n}" by (subst image_mult_atLeastAtMost) simp_all also have "?f = (\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)" using True by (intro ext) (simp add: field_simps) finally have "((\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" (is ?P) . also have "?P \ ((\x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric]) also have "\ \ ((\x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div) finally have \ . note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"] have "((\x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n) has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P) by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric]) (simp add: algebra_simps) also have "?P \ ((\x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n) has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real) also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) = (of_nat n powr z * fact n / pochhammer z (n+1))" by (auto simp add: powr_def algebra_simps exp_diff exp_of_real) finally show ?thesis by (subst has_integral_restrict) simp_all next case False thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl) qed have "eventually (\n. Gamma_series z n = of_nat n powr z * fact n / pochhammer z (n+1)) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: powr_def algebra_simps Gamma_series_def) from this and Gamma_series_LIMSEQ[of z] have C: "(\k. of_nat k powr z * fact k / pochhammer z (k+1)) \ Gamma z" by (blast intro: Lim_transform_eventually) { fix x :: real assume x: "x \ 0" have lim_exp: "(\k. (1 - x / real k) ^ k) \ exp (-x)" using tendsto_exp_limit_sequentially[of "-x"] by simp have "(\k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k)) \ of_real x powr (z - 1) * of_real (exp (-x))" (is ?P) by (intro tendsto_intros lim_exp) also from eventually_gt_at_top[of "nat \x\"] have "eventually (\k. of_nat k > x) sequentially" by eventually_elim linarith hence "?P \ (\k. if x \ of_nat k then of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0) \ of_real x powr (z - 1) * of_real (exp (-x))" by (intro tendsto_cong) (auto elim!: eventually_mono) finally have \ . } hence D: "\x\{0..}. (\k. if x \ {0..real k} then of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0) \ of_real x powr (z - 1) / of_real (exp x)" by (simp add: exp_minus field_simps cong: if_cong) have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" by (intro tendsto_intros ln_x_over_x_tendsto_0) hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp from order_tendstoD(2)[OF this, of "1/2"] have "eventually (\x. (Re z - 1) * ln x / x < 1/2) at_top" by simp from eventually_conj[OF this eventually_gt_at_top[of 0]] obtain x0 where "\x\x0. (Re z - 1) * ln x / x < 1/2 \ x > 0" by (auto simp: eventually_at_top_linorder) hence x0: "x0 > 0" "\x. x \ x0 \ (Re z - 1) * ln x < x / 2" by auto define h where "h = (\x. if x \ {0..x0} then x powr (Re z - 1) else exp (-x/2))" have le_h: "x powr (Re z - 1) * exp (-x) \ h x" if x: "x \ 0" for x proof (cases "x > x0") case True from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)" by (simp add: powr_def exp_diff exp_minus field_simps exp_add) also from x0(2)[of x] True have "\ < exp (-x/2)" by (simp add: field_simps) finally show ?thesis using True by (auto simp add: h_def) next case False from x have "x powr (Re z - 1) * exp (- x) \ x powr (Re z - 1) * 1" by (intro mult_left_mono) simp_all with False show ?thesis by (auto simp add: h_def) qed have E: "\x\{0..}. cmod (if x \ {0..real k} then of_real x powr (z - 1) * (1 - complex_of_real x / of_nat k) ^ k else 0) \ h x" (is "\x\_. ?f x \ _") for k proof safe fix x :: real assume x: "x \ 0" { fix x :: real and n :: nat assume x: "x \ of_nat n" have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp also have "norm \ = \(1 - x / real n)\" by (subst norm_of_real) (rule refl) also from x have "\ = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: field_split_simps) finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" . } note D = this from D[of x k] x have "?f x \ (if of_nat k \ x \ k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)" by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg) also have "\ \ x powr (Re z - 1) * exp (-x)" by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n) also from x have "\ \ h x" by (rule le_h) finally show "?f x \ h x" . qed have F: "h integrable_on {0..}" unfolding h_def by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all) show ?thesis by (rule has_integral_dominated_convergence[OF B F E D C]) qed lemma Gamma_integral_real: assumes x: "x > (0 :: real)" shows "((\t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}" proof - have A: "((\t. complex_of_real t powr (complex_of_real x - 1) / complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}" using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real) have "((\t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}" by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric]) from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def) qed lemma absolutely_integrable_Gamma_integral': assumes "Re z > 0" shows "(\t. complex_of_real t powr (z - 1) / of_real (exp t)) absolutely_integrable_on {0<..}" using absolutely_integrable_Gamma_integral [OF assms zero_less_one] by simp lemma Gamma_integral_complex': assumes z: "Re z > 0" shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0<..}" proof - have "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" by (rule Gamma_integral_complex) fact+ hence "((\t. if t \ {0<..} then of_real t powr (z - 1) / of_real (exp t) else 0) has_integral Gamma z) {0..}" by (rule has_integral_spike [of "{0}", rotated 2]) auto also have "?this = ?thesis" by (subst has_integral_restrict) auto finally show ?thesis . qed lemma Gamma_conv_nn_integral_real: assumes "s > (0::real)" shows "Gamma s = nn_integral lborel (\t. ennreal (indicator {0..} t * t powr (s - 1) / exp t))" using nn_integral_has_integral_lebesgue[OF _ Gamma_integral_real[OF assms]] by simp lemma integrable_Beta: assumes "a > 0" "b > (0::real)" shows "set_integrable lborel {0..1} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" proof - define C where "C = max 1 ((1/2) powr (b - 1))" define D where "D = max 1 ((1/2) powr (a - 1))" have C: "(1 - x) powr (b - 1) \ C" if "x \ {0..1/2}" for x proof (cases "b < 1") case False with that have "(1 - x) powr (b - 1) \ (1 powr (b - 1))" by (intro powr_mono2) auto thus ?thesis by (auto simp: C_def) qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 C_def) have D: "x powr (a - 1) \ D" if "x \ {1/2..1}" for x proof (cases "a < 1") case False with that have "x powr (a - 1) \ (1 powr (a - 1))" by (intro powr_mono2) auto thus ?thesis by (auto simp: D_def) next case True qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 D_def) have [simp]: "C \ 0" "D \ 0" by (simp_all add: C_def D_def) have I1: "set_integrable lborel {0..1/2} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "(\t. t powr (a - 1)) integrable_on {0..1/2}" by (rule integrable_on_powr_from_0) (use assms in auto) hence "(\t. t powr (a - 1)) absolutely_integrable_on {0..1/2}" by (subst absolutely_integrable_on_iff_nonneg) auto from integrable_mult_right[OF this [unfolded set_integrable_def], of C] show "integrable lborel (\x. indicat_real {0..1/2} x *\<^sub>R (C * x powr (a - 1)))" by (subst (asm) integrable_completion) (auto simp: mult_ac) next fix x :: real have "x powr (a - 1) * (1 - x) powr (b - 1) \ x powr (a - 1) * C" if "x \ {0..1/2}" using that by (intro mult_left_mono powr_mono2 C) auto thus "norm (indicator {0..1/2} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \ norm (indicator {0..1/2} x *\<^sub>R (C * x powr (a - 1)))" by (auto simp: indicator_def abs_mult mult_ac) qed (auto intro!: AE_I2 simp: indicator_def) have I2: "set_integrable lborel {1/2..1} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "(\t. t powr (b - 1)) integrable_on {0..1/2}" by (rule integrable_on_powr_from_0) (use assms in auto) hence "(\t. t powr (b - 1)) integrable_on (cbox 0 (1/2))" by simp from integrable_affinity[OF this, of "-1" 1] have "(\t. (1 - t) powr (b - 1)) integrable_on {1/2..1}" by simp hence "(\t. (1 - t) powr (b - 1)) absolutely_integrable_on {1/2..1}" by (subst absolutely_integrable_on_iff_nonneg) auto from integrable_mult_right[OF this [unfolded set_integrable_def], of D] show "integrable lborel (\x. indicat_real {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))" by (subst (asm) integrable_completion) (auto simp: mult_ac) next fix x :: real have "x powr (a - 1) * (1 - x) powr (b - 1) \ D * (1 - x) powr (b - 1)" if "x \ {1/2..1}" using that by (intro mult_right_mono powr_mono2 D) auto thus "norm (indicator {1/2..1} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \ norm (indicator {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))" by (auto simp: indicator_def abs_mult mult_ac) qed (auto intro!: AE_I2 simp: indicator_def) have "set_integrable lborel ({0..1/2} \ {1/2..1}) (\t. t powr (a - 1) * (1 - t) powr (b - 1))" by (intro set_integrable_Un I1 I2) auto also have "{0..1/2} \ {1/2..1} = {0..(1::real)}" by auto finally show ?thesis . qed lemma integrable_Beta': assumes "a > 0" "b > (0::real)" shows "(\t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}" using integrable_Beta[OF assms] by (rule set_borel_integral_eq_integral) theorem has_integral_Beta_real: assumes a: "a > 0" and b: "b > (0 :: real)" shows "((\t. t powr (a - 1) * (1 - t) powr (b - 1)) has_integral Beta a b) {0..1}" proof - define B where "B = integral {0..1} (\x. x powr (a - 1) * (1 - x) powr (b - 1))" have [simp]: "B \ 0" unfolding B_def using a b by (intro integral_nonneg integrable_Beta') auto from a b have "ennreal (Gamma a * Gamma b) = (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \lborel) * (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \lborel)" by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) * ennreal (indicator {0..} u * u powr (b - 1) / exp u) \lborel \lborel)" by (simp add: nn_integral_cmult nn_integral_multc) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator ({0..}\{0..}) (t,u) * t powr (a - 1) * u powr (b - 1) / exp (t + u)) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel \lborel)" proof (rule nn_integral_cong, goal_cases) case (1 t) have "(\\<^sup>+u. ennreal (indicator ({0..}\{0..}) (t,u) * t powr (a - 1) * u powr (b - 1) / exp (t + u)) \distr lborel borel ((+) (-t))) = (\\<^sup>+u. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel)" by (subst nn_integral_distr) (auto intro!: nn_integral_cong simp: indicator_def) thus ?case by (subst (asm) lborel_distr_plus) qed also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel \lborel)" by (subst lborel_pair.Fubini') (auto simp: case_prod_unfold indicator_def cong: measurable_cong_sets) also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) * ennreal (indicator {0..} u / exp u) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def ennreal_mult' [symmetric]) also have "\ = (\\<^sup>+u. (\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) * ennreal (indicator {0..} u / exp u) \lborel)" by (subst nn_integral_multc [symmetric]) auto also have "\ = (\\<^sup>+u. (\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) * ennreal (indicator {0<..} u / exp u) \lborel)" by (intro nn_integral_cong_AE eventually_mono[OF AE_lborel_singleton[of 0]]) (auto simp: indicator_def) also have "\ = (\\<^sup>+u. ennreal B * ennreal (indicator {0..} u / exp u * u powr (a + b - 1)) \lborel)" proof (intro nn_integral_cong, goal_cases) case (1 u) show ?case proof (cases "u > 0") case True have "(\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) = (\\<^sup>+t. ennreal (indicator {0..1} t * (u * t) powr (a - 1) * (u - u * t) powr (b - 1)) \distr lborel borel ((*) (1 / u)))" (is "_ = nn_integral _ ?f") using True by (subst nn_integral_distr) (auto simp: indicator_def field_simps intro!: nn_integral_cong) also have "distr lborel borel ((*) (1 / u)) = density lborel (\_. u)" using \u > 0\ by (subst lborel_distr_mult) auto also have "nn_integral \ ?f = (\\<^sup>+x. ennreal (indicator {0..1} x * (u * (u * x) powr (a - 1) * (u * (1 - x)) powr (b - 1))) \lborel)" using \u > 0\ by (subst nn_integral_density) (auto simp: ennreal_mult' [symmetric] algebra_simps) also have "\ = (\\<^sup>+x. ennreal (u powr (a + b - 1)) * ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel)" using \u > 0\ a b by (intro nn_integral_cong) (auto simp: indicator_def powr_mult powr_add powr_diff mult_ac ennreal_mult' [symmetric]) also have "\ = ennreal (u powr (a + b - 1)) * (\\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel)" by (subst nn_integral_cmult) auto also have "((\x. x powr (a - 1) * (1 - x) powr (b - 1)) has_integral integral {0..1} (\x. x powr (a - 1) * (1 - x) powr (b - 1))) {0..1}" using a b by (intro integrable_integral integrable_Beta') from nn_integral_has_integral_lebesgue[OF _ this] a b have "(\\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel) = B" by (simp add: mult_ac B_def) finally show ?thesis using \u > 0\ by (simp add: ennreal_mult' [symmetric] mult_ac) qed auto qed also have "\ = ennreal B * ennreal (Gamma (a + b))" using a b by (subst nn_integral_cmult) (auto simp: Gamma_conv_nn_integral_real) also have "\ = ennreal (B * Gamma (a + b))" by (subst (1 2) mult.commute, intro ennreal_mult' [symmetric]) (use a b in auto) finally have "B = Beta a b" using a b Gamma_real_pos[of "a + b"] by (subst (asm) ennreal_inj) (auto simp: field_simps Beta_def Gamma_eq_zero_iff) moreover have "(\t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}" by (intro integrable_Beta' a b) ultimately show ?thesis by (simp add: has_integral_iff B_def) qed subsection \The Weierstra{\ss} product formula for the sine\ theorem sin_product_formula_complex: fixes z :: complex shows "(\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k^2)) \ sin (of_real pi * z)" proof - let ?f = "rGamma_series_Weierstrass" have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n)) \ (- of_real pi * inverse z) * (rGamma z * rGamma (- z))" by (intro tendsto_intros rGamma_Weierstrass_complex) also have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) = (\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k ^ 2))" proof fix n :: nat have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * (\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)" by (simp add: rGamma_series_Weierstrass_def mult_ac exp_minus divide_simps prod.distrib[symmetric] power2_eq_square) also have "(\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) = (\k=1..n. 1 - z^2 / of_nat k ^ 2)" by (intro prod.cong) (simp_all add: power2_eq_square field_simps) finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \" by (simp add: field_split_simps) qed also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)" by (subst rGamma_reflection_complex') (simp add: field_split_simps) finally show ?thesis . qed lemma sin_product_formula_real: "(\n. pi * (x::real) * (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x)" proof - from sin_product_formula_complex[of "of_real x"] have "(\n. of_real pi * of_real x * (\k=1..n. 1 - (of_real x)^2 / (of_nat k)^2)) \ sin (of_real pi * of_real x :: complex)" (is "?f \ ?y") . also have "?f = (\n. of_real (pi * x * (\k=1..n. 1 - x^2 / (of_nat k^2))))" by simp also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult) finally show ?thesis by (subst (asm) tendsto_of_real_iff) qed lemma sin_product_formula_real': assumes "x \ (0::real)" shows "(\n. (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x) / (pi * x)" using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms by simp theorem wallis: "(\n. \k=1..n. (4*real k^2) / (4*real k^2 - 1)) \ pi / 2" proof - from tendsto_inverse[OF tendsto_mult[OF sin_product_formula_real[of "1/2"] tendsto_const[of "2/pi"]]] have "(\n. (\k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) \ pi/2" by (simp add: prod_inversef [symmetric]) also have "(\n. (\k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) = (\n. (\k=1..n. (4*real k^2)/(4*real k^2 - 1)))" by (intro ext prod.cong refl) (simp add: field_split_simps) finally show ?thesis . qed subsection \The Solution to the Basel problem\ theorem inverse_squares_sums: "(\n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)" proof - define P where "P x n = (\k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n define K where "K = (\n. inverse (real_of_nat (Suc n))^2)" define f where [abs_def]: "f x = (\n. P x n / of_nat (Suc n)^2)" for x define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x have sums: "(\n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x proof (cases "x = 0") assume x: "x = 0" have "summable (\n. inverse ((real_of_nat (Suc n))\<^sup>2))" using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums) next assume x: "x \ 0" have "(\n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))" unfolding P_def using x by (intro telescope_sums' sin_product_formula_real') also have "(\n. P x n - P x (Suc n)) = (\n. (x^2 / of_nat (Suc n)^2) * P x n)" unfolding P_def by (simp add: prod.nat_ivl_Suc' algebra_simps) also have "P x 0 = 1" by (simp add: P_def) finally have "(\n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" . from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp qed have "continuous_on (ball 0 1) f" proof (rule uniform_limit_theorem; (intro always_eventually allI)?) show "uniform_limit (ball 0 1) (\n x. \k ball 0 1" { fix k :: nat assume k: "k \ 1" from x have "x^2 < 1" by (auto simp: abs_square_less_1) also from k have "\ \ of_nat k^2" by simp finally have "(1 - x^2 / of_nat k^2) \ {0..1}" using k by (simp_all add: field_simps del: of_nat_Suc) } hence "(\k=1..n. abs (1 - x^2 / of_nat k^2)) \ (\k=1..n. 1)" by (intro prod_mono) simp thus "norm (P x n / (of_nat (Suc n)^2)) \ 1 / of_nat (Suc n)^2" unfolding P_def by (simp add: field_simps abs_prod del: of_nat_Suc) qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide) qed (auto simp: P_def intro!: continuous_intros) hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all hence "(f \ 0 \ f 0)" by (simp add: isCont_def) also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide) finally have "f \ 0 \ K" . moreover have "f \ 0 \ pi^2 / 6" proof (rule Lim_transform_eventually) define f' where [abs_def]: "f' x = (\n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x have "eventually (\x. x \ (0::real)) (at 0)" by (auto simp add: eventually_at intro!: exI[of _ 1]) thus "eventually (\x. f' x = f x) (at 0)" proof eventually_elim fix x :: real assume x: "x \ 0" have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def) with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"] have "(\n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))" by (simp add: eval_nat_numeral) from sums_divide[OF this, of "x^3 * pi"] x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)" by (simp add: field_split_simps eval_nat_numeral) with x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)" by (simp add: g_def) hence "f' x = g x / x^2" by (simp add: sums_iff f'_def) also have "\ = f x" using sums[of x] x by (simp add: sums_iff g_def f_def) finally show "f' x = f x" . qed have "isCont f' 0" unfolding f'_def proof (intro isCont_powser_converges_everywhere) fix x :: real show "summable (\n. -sin_coeff (n+3) * pi^(n+2) * x^n)" proof (cases "x = 0") assume x: "x \ 0" from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x show ?thesis by (simp add: field_split_simps eval_nat_numeral) qed (simp only: summable_0_powser) qed hence "f' \ 0 \ f' 0" by (simp add: isCont_def) also have "f' 0 = pi * pi / fact 3" unfolding f'_def by (subst powser_zero) (simp add: sin_coeff_def) finally show "f' \ 0 \ pi^2 / 6" by (simp add: eval_nat_numeral) qed ultimately have "K = pi^2 / 6" by (rule LIM_unique) moreover from inverse_power_summable[of 2] have "summable (\n. (inverse (real_of_nat (Suc n)))\<^sup>2)" by (subst summable_Suc_iff) (simp add: power_inverse) ultimately show ?thesis unfolding K_def by (auto simp add: sums_iff power_divide inverse_eq_divide) qed end diff --git a/src/HOL/Analysis/Interval_Integral.thy b/src/HOL/Analysis/Interval_Integral.thy --- a/src/HOL/Analysis/Interval_Integral.thy +++ b/src/HOL/Analysis/Interval_Integral.thy @@ -1,1107 +1,1129 @@ (* Title: HOL/Analysis/Interval_Integral.thy Author: Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU) Lebesgue integral over an interval (with endpoints possibly +-\) *) theory Interval_Integral (*FIX ME rename? Lebesgue *) imports Equivalence_Lebesgue_Henstock_Integration begin definition "einterval a b = {x. a < ereal x \ ereal x < b}" lemma einterval_eq[simp]: shows einterval_eq_Icc: "einterval (ereal a) (ereal b) = {a <..< b}" and einterval_eq_Ici: "einterval (ereal a) \ = {a <..}" and einterval_eq_Iic: "einterval (- \) (ereal b) = {..< b}" and einterval_eq_UNIV: "einterval (- \) \ = UNIV" by (auto simp: einterval_def) lemma einterval_same: "einterval a a = {}" by (auto simp: einterval_def) lemma einterval_iff: "x \ einterval a b \ a < ereal x \ ereal x < b" by (simp add: einterval_def) lemma einterval_nonempty: "a < b \ \c. c \ einterval a b" by (cases a b rule: ereal2_cases, auto simp: einterval_def intro!: dense gt_ex lt_ex) lemma open_einterval[simp]: "open (einterval a b)" by (cases a b rule: ereal2_cases) (auto simp: einterval_def intro!: open_Collect_conj open_Collect_less continuous_intros) lemma borel_einterval[measurable]: "einterval a b \ sets borel" unfolding einterval_def by measurable subsection \Approximating a (possibly infinite) interval\ lemma filterlim_sup1: "(LIM x F. f x :> G1) \ (LIM x F. f x :> (sup G1 G2))" unfolding filterlim_def by (auto intro: le_supI1) lemma ereal_incseq_approx: fixes a b :: ereal assumes "a < b" obtains X :: "nat \ real" where "incseq X" "\i. a < X i" "\i. X i < b" "X \ b" proof (cases b) case PInf with \a < b\ have "a = -\ \ (\r. a = ereal r)" by (cases a) auto moreover have "(\x. ereal (real (Suc x))) \ \" by (simp add: Lim_PInfty filterlim_sequentially_Suc) (metis le_SucI of_nat_Suc of_nat_mono order_trans real_arch_simple) moreover have "\r. (\x. ereal (r + real (Suc x))) \ \" by (simp add: filterlim_sequentially_Suc Lim_PInfty) (metis add.commute diff_le_eq nat_ceiling_le_eq) ultimately show thesis by (intro that[of "\i. real_of_ereal a + Suc i"]) (auto simp: incseq_def PInf) next case (real b') define d where "d = b' - (if a = -\ then b' - 1 else real_of_ereal a)" with \a < b\ have a': "0 < d" by (cases a) (auto simp: real) moreover have "\i r. r < b' \ (b' - r) * 1 < (b' - r) * real (Suc (Suc i))" by (intro mult_strict_left_mono) auto with \a < b\ a' have "\i. a < ereal (b' - d / real (Suc (Suc i)))" by (cases a) (auto simp: real d_def field_simps) moreover have "(\i. b' - d / real i) \ b'" by (force intro: tendsto_eq_intros tendsto_divide_0[OF tendsto_const] filterlim_sup1 simp: at_infinity_eq_at_top_bot filterlim_real_sequentially) then have "(\i. b' - d / Suc (Suc i)) \ b'" by (blast intro: dest: filterlim_sequentially_Suc [THEN iffD2]) ultimately show thesis by (intro that[of "\i. b' - d / Suc (Suc i)"]) (auto simp: real incseq_def intro!: divide_left_mono) -qed (insert \a < b\, auto) +qed (use \a < b\ in auto) lemma ereal_decseq_approx: fixes a b :: ereal assumes "a < b" obtains X :: "nat \ real" where "decseq X" "\i. a < X i" "\i. X i < b" "X \ a" proof - have "-b < -a" using \a < b\ by simp - from ereal_incseq_approx[OF this] guess X . + from ereal_incseq_approx[OF this] obtain X where + "incseq X" + "\i. - b < ereal (X i)" + "\i. ereal (X i) < - a" + "(\x. ereal (X x)) \ - a" + by auto then show thesis apply (intro that[of "\i. - X i"]) apply (auto simp: decseq_def incseq_def simp flip: uminus_ereal.simps) apply (metis ereal_minus_less_minus ereal_uminus_uminus ereal_Lim_uminus)+ done qed proposition einterval_Icc_approximation: fixes a b :: ereal assumes "a < b" obtains u l :: "nat \ real" where "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" proof - from dense[OF \a < b\] obtain c where "a < c" "c < b" by safe - from ereal_incseq_approx[OF \c < b\] guess u . note u = this - from ereal_decseq_approx[OF \a < c\] guess l . note l = this + from ereal_incseq_approx[OF \c < b\] obtain u where u: + "incseq u" + "\i. c < ereal (u i)" + "\i. ereal (u i) < b" + "(\x. ereal (u x)) \ b" + by auto + from ereal_decseq_approx[OF \a < c\] obtain l where l: + "decseq l" + "\i. a < ereal (l i)" + "\i. ereal (l i) < c" + "(\x. ereal (l x)) \ a" + by auto { fix i from less_trans[OF \l i < c\ \c < u i\] have "l i < u i" by simp } have "einterval a b = (\i. {l i .. u i})" proof (auto simp: einterval_iff) fix x assume "a < ereal x" "ereal x < b" have "eventually (\i. ereal (l i) < ereal x) sequentially" using l(4) \a < ereal x\ by (rule order_tendstoD) moreover have "eventually (\i. ereal x < ereal (u i)) sequentially" using u(4) \ereal x< b\ by (rule order_tendstoD) ultimately have "eventually (\i. l i < x \ x < u i) sequentially" by eventually_elim auto then show "\i. l i \ x \ x \ u i" by (auto intro: less_imp_le simp: eventually_sequentially) next fix x i assume "l i \ x" "x \ u i" with \a < ereal (l i)\ \ereal (u i) < b\ show "a < ereal x" "ereal x < b" by (auto simp flip: ereal_less_eq(3)) qed show thesis by (intro that) fact+ qed (* TODO: in this definition, it would be more natural if einterval a b included a and b when they are real. *) definition\<^marker>\tag important\ interval_lebesgue_integral :: "real measure \ ereal \ ereal \ (real \ 'a) \ 'a::{banach, second_countable_topology}" where "interval_lebesgue_integral M a b f = (if a \ b then (LINT x:einterval a b|M. f x) else - (LINT x:einterval b a|M. f x))" syntax "_ascii_interval_lebesgue_integral" :: "pttrn \ real \ real \ real measure \ real \ real" ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60) translations "LINT x=a..b|M. f" == "CONST interval_lebesgue_integral M a b (\x. f)" definition\<^marker>\tag important\ interval_lebesgue_integrable :: "real measure \ ereal \ ereal \ (real \ 'a::{banach, second_countable_topology}) \ bool" where "interval_lebesgue_integrable M a b f = (if a \ b then set_integrable M (einterval a b) f else set_integrable M (einterval b a) f)" syntax "_ascii_interval_lebesgue_borel_integral" :: "pttrn \ real \ real \ real \ real" ("(4LBINT _=_.._. _)" [0,60,60,61] 60) translations "LBINT x=a..b. f" == "CONST interval_lebesgue_integral CONST lborel a b (\x. f)" subsection\Basic properties of integration over an interval\ lemma interval_lebesgue_integral_cong: "a \ b \ (\x. x \ einterval a b \ f x = g x) \ einterval a b \ sets M \ interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g" by (auto intro: set_lebesgue_integral_cong simp: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ a \ b \ AE x \ einterval a b in M. f x = g x \ einterval a b \ sets M \ interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g" by (auto intro: set_lebesgue_integral_cong_AE simp: interval_lebesgue_integral_def) lemma interval_integrable_mirror: shows "interval_lebesgue_integrable lborel a b (\x. f (-x)) \ interval_lebesgue_integrable lborel (-b) (-a) f" proof - have *: "indicator (einterval a b) (- x) = (indicator (einterval (-b) (-a)) x :: real)" for a b :: ereal and x :: real by (cases a b rule: ereal2_cases) (auto simp: einterval_def split: split_indicator) show ?thesis unfolding interval_lebesgue_integrable_def using lborel_integrable_real_affine_iff[symmetric, of "-1" "\x. indicator (einterval _ _) x *\<^sub>R f x" 0] by (simp add: * set_integrable_def) qed lemma interval_lebesgue_integral_add [intro, simp]: fixes M a b f assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g" shows "interval_lebesgue_integrable M a b (\x. f x + g x)" and "interval_lebesgue_integral M a b (\x. f x + g x) = interval_lebesgue_integral M a b f + interval_lebesgue_integral M a b g" using assms by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def field_simps) lemma interval_lebesgue_integral_diff [intro, simp]: fixes M a b f assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g" shows "interval_lebesgue_integrable M a b (\x. f x - g x)" and "interval_lebesgue_integral M a b (\x. f x - g x) = interval_lebesgue_integral M a b f - interval_lebesgue_integral M a b g" using assms by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def field_simps) lemma interval_lebesgue_integrable_mult_right [intro, simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "(c \ 0 \ interval_lebesgue_integrable M a b f) \ interval_lebesgue_integrable M a b (\x. c * f x)" by (simp add: interval_lebesgue_integrable_def) lemma interval_lebesgue_integrable_mult_left [intro, simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "(c \ 0 \ interval_lebesgue_integrable M a b f) \ interval_lebesgue_integrable M a b (\x. f x * c)" by (simp add: interval_lebesgue_integrable_def) lemma interval_lebesgue_integrable_divide [intro, simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, field, second_countable_topology}" shows "(c \ 0 \ interval_lebesgue_integrable M a b f) \ interval_lebesgue_integrable M a b (\x. f x / c)" by (simp add: interval_lebesgue_integrable_def) lemma interval_lebesgue_integral_mult_right [simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "interval_lebesgue_integral M a b (\x. c * f x) = c * interval_lebesgue_integral M a b f" by (simp add: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_mult_left [simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, second_countable_topology}" shows "interval_lebesgue_integral M a b (\x. f x * c) = interval_lebesgue_integral M a b f * c" by (simp add: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_divide [simp]: fixes M a b c and f :: "real \ 'a::{banach, real_normed_field, field, second_countable_topology}" shows "interval_lebesgue_integral M a b (\x. f x / c) = interval_lebesgue_integral M a b f / c" by (simp add: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_uminus: "interval_lebesgue_integral M a b (\x. - f x) = - interval_lebesgue_integral M a b f" by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def) lemma interval_lebesgue_integral_of_real: "interval_lebesgue_integral M a b (\x. complex_of_real (f x)) = of_real (interval_lebesgue_integral M a b f)" unfolding interval_lebesgue_integral_def by (auto simp: interval_lebesgue_integral_def set_integral_complex_of_real) lemma interval_lebesgue_integral_le_eq: fixes a b f assumes "a \ b" shows "interval_lebesgue_integral M a b f = (LINT x : einterval a b | M. f x)" using assms by (auto simp: interval_lebesgue_integral_def) lemma interval_lebesgue_integral_gt_eq: fixes a b f assumes "a > b" shows "interval_lebesgue_integral M a b f = -(LINT x : einterval b a | M. f x)" using assms by (auto simp: interval_lebesgue_integral_def less_imp_le einterval_def) lemma interval_lebesgue_integral_gt_eq': fixes a b f assumes "a > b" shows "interval_lebesgue_integral M a b f = - interval_lebesgue_integral M b a f" using assms by (auto simp: interval_lebesgue_integral_def less_imp_le einterval_def) lemma interval_integral_endpoints_same [simp]: "(LBINT x=a..a. f x) = 0" by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def einterval_same) lemma interval_integral_endpoints_reverse: "(LBINT x=a..b. f x) = -(LBINT x=b..a. f x)" by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integral_def set_lebesgue_integral_def einterval_same) lemma interval_integrable_endpoints_reverse: "interval_lebesgue_integrable lborel a b f \ interval_lebesgue_integrable lborel b a f" by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integrable_def einterval_same) lemma interval_integral_reflect: "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))" proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (auto simp: interval_lebesgue_integral_def interval_integrable_endpoints_reverse split: if_split_asm) next case (le a b) have "LBINT x:{x. - x \ einterval a b}. f (- x) = LBINT x:einterval (- b) (- a). f (- x)" unfolding interval_lebesgue_integrable_def set_lebesgue_integral_def apply (rule Bochner_Integration.integral_cong [OF refl]) by (auto simp: einterval_iff ereal_uminus_le_reorder ereal_uminus_less_reorder not_less simp flip: uminus_ereal.simps split: split_indicator) then show ?case unfolding interval_lebesgue_integral_def by (subst set_integral_reflect) (simp add: le) qed lemma interval_lebesgue_integral_0_infty: "interval_lebesgue_integrable M 0 \ f \ set_integrable M {0<..} f" "interval_lebesgue_integral M 0 \ f = (LINT x:{0<..}|M. f x)" unfolding zero_ereal_def by (auto simp: interval_lebesgue_integral_le_eq interval_lebesgue_integrable_def) lemma interval_integral_to_infinity_eq: "(LINT x=ereal a..\ | M. f x) = (LINT x : {a<..} | M. f x)" unfolding interval_lebesgue_integral_def by auto proposition interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \ f) = (set_integrable M {a<..} f)" unfolding interval_lebesgue_integrable_def by auto subsection\Basic properties of integration over an interval wrt lebesgue measure\ lemma interval_integral_zero [simp]: fixes a b :: ereal shows "LBINT x=a..b. 0 = 0" unfolding interval_lebesgue_integral_def set_lebesgue_integral_def einterval_eq by simp lemma interval_integral_const [intro, simp]: fixes a b c :: real shows "interval_lebesgue_integrable lborel a b (\x. c)" and "LBINT x=a..b. c = c * (b - a)" unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def einterval_eq by (auto simp: less_imp_le field_simps measure_def set_integrable_def set_lebesgue_integral_def) lemma interval_integral_cong_AE: assumes [measurable]: "f \ borel_measurable borel" "g \ borel_measurable borel" assumes "AE x \ einterval (min a b) (max a b) in lborel. f x = g x" shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g" using assms proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b]) next case (le a b) then show ?case by (auto simp: interval_lebesgue_integral_def max_def min_def intro!: set_lebesgue_integral_cong_AE) qed lemma interval_integral_cong: assumes "\x. x \ einterval (min a b) (max a b) \ f x = g x" shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g" using assms proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b]) next case (le a b) then show ?case by (auto simp: interval_lebesgue_integral_def max_def min_def intro!: set_lebesgue_integral_cong) qed lemma interval_lebesgue_integrable_cong_AE: "f \ borel_measurable lborel \ g \ borel_measurable lborel \ AE x \ einterval (min a b) (max a b) in lborel. f x = g x \ interval_lebesgue_integrable lborel a b f = interval_lebesgue_integrable lborel a b g" apply (simp add: interval_lebesgue_integrable_def) apply (intro conjI impI set_integrable_cong_AE) apply (auto simp: min_def max_def) done lemma interval_integrable_abs_iff: fixes f :: "real \ real" shows "f \ borel_measurable lborel \ interval_lebesgue_integrable lborel a b (\x. \f x\) = interval_lebesgue_integrable lborel a b f" unfolding interval_lebesgue_integrable_def by (subst (1 2) set_integrable_abs_iff') simp_all lemma interval_integral_Icc: fixes a b :: real shows "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {a..b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{a, b}"] simp add: interval_lebesgue_integral_def) lemma interval_integral_Icc': "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {x. a \ ereal x \ ereal x \ b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"] simp add: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ioc: "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {a<..b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{a, b}"] simp add: interval_lebesgue_integral_def einterval_iff) (* TODO: other versions as well? *) (* Yes: I need the Icc' version. *) lemma interval_integral_Ioc': "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {x. a < ereal x \ ereal x \ b}. f x)" by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"] simp add: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ico: "a \ b \ (LBINT x=a..b. f x) = (LBINT x : {a..a\ < \ \ (LBINT x=a..\. f x) = (LBINT x : {real_of_ereal a <..}. f x)" by (auto simp: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_Ioo: "a \ b \ \a\ < \ ==> \b\ < \ \ (LBINT x=a..b. f x) = (LBINT x : {real_of_ereal a <..< real_of_ereal b}. f x)" by (auto simp: interval_lebesgue_integral_def einterval_iff) lemma interval_integral_discrete_difference: fixes f :: "real \ 'b::{banach, second_countable_topology}" and a b :: ereal assumes "countable X" and eq: "\x. a \ b \ a < x \ x < b \ x \ X \ f x = g x" and anti_eq: "\x. b \ a \ b < x \ x < a \ x \ X \ f x = g x" assumes "\x. x \ X \ emeasure M {x} = 0" "\x. x \ X \ {x} \ sets M" shows "interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g" unfolding interval_lebesgue_integral_def set_lebesgue_integral_def apply (intro if_cong refl arg_cong[where f="\x. - x"] integral_discrete_difference[of X] assms) apply (auto simp: eq anti_eq einterval_iff split: split_indicator) done lemma interval_integral_sum: fixes a b c :: ereal assumes integrable: "interval_lebesgue_integrable lborel (min a (min b c)) (max a (max b c)) f" shows "(LBINT x=a..b. f x) + (LBINT x=b..c. f x) = (LBINT x=a..c. f x)" proof - let ?I = "\a b. LBINT x=a..b. f x" { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \ b" "b \ c" then have ord: "a \ b" "b \ c" "a \ c" and f': "set_integrable lborel (einterval a c) f" by (auto simp: interval_lebesgue_integrable_def) then have f: "set_borel_measurable borel (einterval a c) f" unfolding set_integrable_def set_borel_measurable_def by (drule_tac borel_measurable_integrable) simp have "(LBINT x:einterval a c. f x) = (LBINT x:einterval a b \ einterval b c. f x)" proof (rule set_integral_cong_set) show "AE x in lborel. (x \ einterval a b \ einterval b c) = (x \ einterval a c)" using AE_lborel_singleton[of "real_of_ereal b"] ord by (cases a b c rule: ereal3_cases) (auto simp: einterval_iff) show "set_borel_measurable lborel (einterval a c) f" "set_borel_measurable lborel (einterval a b \ einterval b c) f" unfolding set_borel_measurable_def using ord by (auto simp: einterval_iff intro!: set_borel_measurable_subset[OF f, unfolded set_borel_measurable_def]) qed also have "\ = (LBINT x:einterval a b. f x) + (LBINT x:einterval b c. f x)" using ord by (intro set_integral_Un_AE) (auto intro!: set_integrable_subset[OF f'] simp: einterval_iff not_less) finally have "?I a b + ?I b c = ?I a c" using ord by (simp add: interval_lebesgue_integral_def) } note 1 = this { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \ b" "b \ c" from 1[OF this] have "?I b c + ?I a b = ?I a c" by (metis add.commute) } note 2 = this have 3: "\a b. b \ a \ (LBINT x=a..b. f x) = - (LBINT x=b..a. f x)" by (rule interval_integral_endpoints_reverse) show ?thesis using integrable apply (cases a b b c a c rule: linorder_le_cases[case_product linorder_le_cases linorder_cases]) apply simp_all (* remove some looping cases *) by (simp_all add: min_absorb1 min_absorb2 max_absorb1 max_absorb2 field_simps 1 2 3) qed lemma interval_integrable_isCont: fixes a b and f :: "real \ 'a::{banach, second_countable_topology}" shows "(\x. min a b \ x \ x \ max a b \ isCont f x) \ interval_lebesgue_integrable lborel a b f" proof (induct a b rule: linorder_wlog) case (le a b) then show ?case unfolding interval_lebesgue_integrable_def set_integrable_def by (auto simp: interval_lebesgue_integrable_def intro!: set_integrable_subset[unfolded set_integrable_def, OF borel_integrable_compact[of "{a .. b}"]] continuous_at_imp_continuous_on) qed (auto intro: interval_integrable_endpoints_reverse[THEN iffD1]) lemma interval_integrable_continuous_on: fixes a b :: real and f assumes "a \ b" and "continuous_on {a..b} f" shows "interval_lebesgue_integrable lborel a b f" using assms unfolding interval_lebesgue_integrable_def apply simp by (rule set_integrable_subset, rule borel_integrable_atLeastAtMost' [of a b], auto) lemma interval_integral_eq_integral: fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ set_integrable lborel {a..b} f \ LBINT x=a..b. f x = integral {a..b} f" by (subst interval_integral_Icc, simp) (rule set_borel_integral_eq_integral) lemma interval_integral_eq_integral': fixes f :: "real \ 'a::euclidean_space" shows "a \ b \ set_integrable lborel (einterval a b) f \ LBINT x=a..b. f x = integral (einterval a b) f" by (subst interval_lebesgue_integral_le_eq, simp) (rule set_borel_integral_eq_integral) subsection\General limit approximation arguments\ proposition interval_integral_Icc_approx_nonneg: fixes a b :: ereal assumes "a < b" fixes u l :: "nat \ real" assumes approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" fixes f :: "real \ real" assumes f_integrable: "\i. set_integrable lborel {l i..u i} f" assumes f_nonneg: "AE x in lborel. a < ereal x \ ereal x < b \ 0 \ f x" assumes f_measurable: "set_borel_measurable lborel (einterval a b) f" assumes lbint_lim: "(\i. LBINT x=l i.. u i. f x) \ C" shows "set_integrable lborel (einterval a b) f" "(LBINT x=a..b. f x) = C" proof - have 1 [unfolded set_integrable_def]: "\i. set_integrable lborel {l i..u i} f" by (rule f_integrable) have 2: "AE x in lborel. mono (\n. indicator {l n..u n} x *\<^sub>R f x)" proof - from f_nonneg have "AE x in lborel. \i. l i \ x \ x \ u i \ 0 \ f x" by eventually_elim (metis approx(5) approx(6) dual_order.strict_trans1 ereal_less_eq(3) le_less_trans) then show ?thesis apply eventually_elim apply (auto simp: mono_def split: split_indicator) apply (metis approx(3) decseqD order_trans) apply (metis approx(2) incseqD order_trans) done qed have 3: "AE x in lborel. (\i. indicator {l i..u i} x *\<^sub>R f x) \ indicator (einterval a b) x *\<^sub>R f x" proof - { fix x i assume "l i \ x" "x \ u i" then have "eventually (\i. l i \ x \ x \ u i) sequentially" apply (auto simp: eventually_sequentially intro!: exI[of _ i]) apply (metis approx(3) decseqD order_trans) apply (metis approx(2) incseqD order_trans) done then have "eventually (\i. f x * indicator {l i..u i} x = f x) sequentially" by eventually_elim auto } then show ?thesis unfolding approx(1) by (auto intro!: AE_I2 tendsto_eventually split: split_indicator) qed have 4: "(\i. \ x. indicator {l i..u i} x *\<^sub>R f x \lborel) \ C" using lbint_lim by (simp add: interval_integral_Icc [unfolded set_lebesgue_integral_def] approx less_imp_le) have 5: "(\x. indicat_real (einterval a b) x *\<^sub>R f x) \ borel_measurable lborel" using f_measurable set_borel_measurable_def by blast have "(LBINT x=a..b. f x) = lebesgue_integral lborel (\x. indicator (einterval a b) x *\<^sub>R f x)" using assms by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def less_imp_le) also have "\ = C" by (rule integral_monotone_convergence [OF 1 2 3 4 5]) finally show "(LBINT x=a..b. f x) = C" . show "set_integrable lborel (einterval a b) f" unfolding set_integrable_def by (rule integrable_monotone_convergence[OF 1 2 3 4 5]) qed proposition interval_integral_Icc_approx_integrable: fixes u l :: "nat \ real" and a b :: ereal fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes "a < b" assumes approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" assumes f_integrable: "set_integrable lborel (einterval a b) f" shows "(\i. LBINT x=l i.. u i. f x) \ (LBINT x=a..b. f x)" proof - have "(\i. LBINT x:{l i.. u i}. f x) \ (LBINT x:einterval a b. f x)" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) show "integrable lborel (\x. norm (indicator (einterval a b) x *\<^sub>R f x))" using f_integrable integrable_norm set_integrable_def by blast show "(\x. indicat_real (einterval a b) x *\<^sub>R f x) \ borel_measurable lborel" using f_integrable by (simp add: set_integrable_def) then show "\i. (\x. indicat_real {l i..u i} x *\<^sub>R f x) \ borel_measurable lborel" by (rule set_borel_measurable_subset [unfolded set_borel_measurable_def]) (auto simp: approx) show "\i. AE x in lborel. norm (indicator {l i..u i} x *\<^sub>R f x) \ norm (indicator (einterval a b) x *\<^sub>R f x)" by (intro AE_I2) (auto simp: approx split: split_indicator) show "AE x in lborel. (\i. indicator {l i..u i} x *\<^sub>R f x) \ indicator (einterval a b) x *\<^sub>R f x" proof (intro AE_I2 tendsto_intros tendsto_eventually) fix x { fix i assume "l i \ x" "x \ u i" with \incseq u\[THEN incseqD, of i] \decseq l\[THEN decseqD, of i] have "eventually (\i. l i \ x \ x \ u i) sequentially" by (auto simp: eventually_sequentially decseq_def incseq_def intro: order_trans) } then show "eventually (\xa. indicator {l xa..u xa} x = (indicator (einterval a b) x::real)) sequentially" using approx order_tendstoD(2)[OF \l \ a\, of x] order_tendstoD(1)[OF \u \ b\, of x] by (auto split: split_indicator) qed qed with \a < b\ \\i. l i < u i\ show ?thesis by (simp add: interval_lebesgue_integral_le_eq[symmetric] interval_integral_Icc less_imp_le) qed subsection\A slightly stronger Fundamental Theorem of Calculus\ text\Three versions: first, for finite intervals, and then two versions for arbitrary intervals.\ (* TODO: make the older versions corollaries of these (using continuous_at_imp_continuous_on, etc.) *) lemma interval_integral_FTC_finite: fixes f F :: "real \ 'a::euclidean_space" and a b :: real assumes f: "continuous_on {min a b..max a b} f" assumes F: "\x. min a b \ x \ x \ max a b \ (F has_vector_derivative (f x)) (at x within {min a b..max a b})" shows "(LBINT x=a..b. f x) = F b - F a" proof (cases "a \ b") case True have "(LBINT x=a..b. f x) = (LBINT x. indicat_real {a..b} x *\<^sub>R f x)" by (simp add: True interval_integral_Icc set_lebesgue_integral_def) also have "\ = F b - F a" proof (rule integral_FTC_atLeastAtMost [OF True]) show "continuous_on {a..b} f" using True f by linarith show "\x. \a \ x; x \ b\ \ (F has_vector_derivative f x) (at x within {a..b})" by (metis F True max.commute max_absorb1 min_def) qed finally show ?thesis . next case False then have "b \ a" by simp have "- interval_lebesgue_integral lborel (ereal b) (ereal a) f = - (LBINT x. indicat_real {b..a} x *\<^sub>R f x)" by (simp add: \b \ a\ interval_integral_Icc set_lebesgue_integral_def) also have "\ = F b - F a" proof (subst integral_FTC_atLeastAtMost [OF \b \ a\]) show "continuous_on {b..a} f" using False f by linarith show "\x. \b \ x; x \ a\ \ (F has_vector_derivative f x) (at x within {b..a})" by (metis F False max_def min_def) qed auto finally show ?thesis by (metis interval_integral_endpoints_reverse) qed lemma interval_integral_FTC_nonneg: fixes f F :: "real \ real" and a b :: ereal assumes "a < b" assumes F: "\x. a < ereal x \ ereal x < b \ DERIV F x :> f x" assumes f: "\x. a < ereal x \ ereal x < b \ isCont f x" assumes f_nonneg: "AE x in lborel. a < ereal x \ ereal x < b \ 0 \ f x" assumes A: "((F \ real_of_ereal) \ A) (at_right a)" assumes B: "((F \ real_of_ereal) \ B) (at_left b)" shows "set_integrable lborel (einterval a b) f" "(LBINT x=a..b. f x) = B - A" proof - obtain u l where approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" by (blast intro: einterval_Icc_approximation[OF \a < b\]) have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) have FTCi: "\i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)" using assms approx apply (intro interval_integral_FTC_finite) apply (auto simp: less_imp_le min_def max_def has_field_derivative_iff_has_vector_derivative[symmetric]) apply (rule continuous_at_imp_continuous_on, auto intro!: f) by (rule DERIV_subset [OF F], auto) have 1: "\i. set_integrable lborel {l i..u i} f" proof - fix i show "set_integrable lborel {l i .. u i} f" using \a < l i\ \u i < b\ unfolding set_integrable_def by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI) (auto simp flip: ereal_less_eq) qed have 2: "set_borel_measurable lborel (einterval a b) f" unfolding set_borel_measurable_def by (auto simp del: real_scaleR_def intro!: borel_measurable_continuous_on_indicator simp: continuous_on_eq_continuous_at einterval_iff f) have 3: "(\i. LBINT x=l i..u i. f x) \ B - A" apply (subst FTCi) apply (intro tendsto_intros) using B approx unfolding tendsto_at_iff_sequentially comp_def using tendsto_at_iff_sequentially[where 'a=real] apply (elim allE[of _ "\i. ereal (u i)"], auto) using A approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (l i)"], auto) show "(LBINT x=a..b. f x) = B - A" by (rule interval_integral_Icc_approx_nonneg [OF \a < b\ approx 1 f_nonneg 2 3]) show "set_integrable lborel (einterval a b) f" by (rule interval_integral_Icc_approx_nonneg [OF \a < b\ approx 1 f_nonneg 2 3]) qed theorem interval_integral_FTC_integrable: fixes f F :: "real \ 'a::euclidean_space" and a b :: ereal assumes "a < b" assumes F: "\x. a < ereal x \ ereal x < b \ (F has_vector_derivative f x) (at x)" assumes f: "\x. a < ereal x \ ereal x < b \ isCont f x" assumes f_integrable: "set_integrable lborel (einterval a b) f" assumes A: "((F \ real_of_ereal) \ A) (at_right a)" assumes B: "((F \ real_of_ereal) \ B) (at_left b)" shows "(LBINT x=a..b. f x) = B - A" proof - obtain u l where approx: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" by (blast intro: einterval_Icc_approximation[OF \a < b\]) have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) have FTCi: "\i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)" using assms approx by (auto simp: less_imp_le min_def max_def intro!: f continuous_at_imp_continuous_on interval_integral_FTC_finite intro: has_vector_derivative_at_within) have "(\i. LBINT x=l i..u i. f x) \ B - A" unfolding FTCi proof (intro tendsto_intros) show "(\x. F (l x)) \ A" using A approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (l i)"], auto) show "(\x. F (u x)) \ B" using B approx unfolding tendsto_at_iff_sequentially comp_def by (elim allE[of _ "\i. ereal (u i)"], auto) qed moreover have "(\i. LBINT x=l i..u i. f x) \ (LBINT x=a..b. f x)" by (rule interval_integral_Icc_approx_integrable [OF \a < b\ approx f_integrable]) ultimately show ?thesis by (elim LIMSEQ_unique) qed (* The second Fundamental Theorem of Calculus and existence of antiderivatives on an einterval. *) theorem interval_integral_FTC2: fixes a b c :: real and f :: "real \ 'a::euclidean_space" assumes "a \ c" "c \ b" and contf: "continuous_on {a..b} f" fixes x :: real assumes "a \ x" and "x \ b" shows "((\u. LBINT y=c..u. f y) has_vector_derivative (f x)) (at x within {a..b})" proof - let ?F = "(\u. LBINT y=a..u. f y)" have intf: "set_integrable lborel {a..b} f" by (rule borel_integrable_atLeastAtMost', rule contf) have "((\u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})" using \a \ x\ \x \ b\ by (auto intro: integral_has_vector_derivative continuous_on_subset [OF contf]) then have "((\u. integral {a..u} f) has_vector_derivative (f x)) (at x within {a..b})" by simp then have "(?F has_vector_derivative (f x)) (at x within {a..b})" by (rule has_vector_derivative_weaken) (auto intro!: assms interval_integral_eq_integral[symmetric] set_integrable_subset [OF intf]) then have "((\x. (LBINT y=c..a. f y) + ?F x) has_vector_derivative (f x)) (at x within {a..b})" by (auto intro!: derivative_eq_intros) then show ?thesis proof (rule has_vector_derivative_weaken) fix u assume "u \ {a .. b}" then show "(LBINT y=c..a. f y) + (LBINT y=a..u. f y) = (LBINT y=c..u. f y)" using assms apply (intro interval_integral_sum) apply (auto simp: interval_lebesgue_integrable_def simp del: real_scaleR_def) by (rule set_integrable_subset [OF intf], auto simp: min_def max_def) qed (insert assms, auto) qed proposition einterval_antiderivative: fixes a b :: ereal and f :: "real \ 'a::euclidean_space" assumes "a < b" and contf: "\x :: real. a < x \ x < b \ isCont f x" shows "\F. \x :: real. a < x \ x < b \ (F has_vector_derivative f x) (at x)" proof - from einterval_nonempty [OF \a < b\] obtain c :: real where [simp]: "a < c" "c < b" by (auto simp: einterval_def) let ?F = "(\u. LBINT y=c..u. f y)" show ?thesis proof (rule exI, clarsimp) fix x :: real assume [simp]: "a < x" "x < b" have 1: "a < min c x" by simp from einterval_nonempty [OF 1] obtain d :: real where [simp]: "a < d" "d < c" "d < x" by (auto simp: einterval_def) have 2: "max c x < b" by simp from einterval_nonempty [OF 2] obtain e :: real where [simp]: "c < e" "x < e" "e < b" by (auto simp: einterval_def) have "(?F has_vector_derivative f x) (at x within {d<..x. \d \ x; x \ e\ \ a < ereal x" using \a < ereal d\ ereal_less_ereal_Ex by auto show "\x. \d \ x; x \ e\ \ ereal x < b" using \ereal e < b\ ereal_less_eq(3) le_less_trans by blast qed then show "(?F has_vector_derivative f x) (at x within {d..e})" by (intro interval_integral_FTC2) (use \d < c\ \c < e\ \d < x\ \x < e\ in \linarith+\) qed auto then show "(?F has_vector_derivative f x) (at x)" by (force simp: has_vector_derivative_within_open [of _ "{d<..The substitution theorem\ text\Once again, three versions: first, for finite intervals, and then two versions for arbitrary intervals.\ theorem interval_integral_substitution_finite: fixes a b :: real and f :: "real \ 'a::euclidean_space" assumes "a \ b" and derivg: "\x. a \ x \ x \ b \ (g has_real_derivative (g' x)) (at x within {a..b})" and contf : "continuous_on (g ` {a..b}) f" and contg': "continuous_on {a..b} g'" shows "LBINT x=a..b. g' x *\<^sub>R f (g x) = LBINT y=g a..g b. f y" proof- have v_derivg: "\x. a \ x \ x \ b \ (g has_vector_derivative (g' x)) (at x within {a..b})" using derivg unfolding has_field_derivative_iff_has_vector_derivative . then have contg [simp]: "continuous_on {a..b} g" by (rule continuous_on_vector_derivative) auto have 1: "\x\{a..b}. u = g x" if "min (g a) (g b) \ u" "u \ max (g a) (g b)" for u by (cases "g a \ g b") (use that assms IVT' [of g a u b] IVT2' [of g b u a] in \auto simp: min_def max_def\) obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \ d" by (metis continuous_image_closed_interval contg \a \ b\) obtain F where derivF: "\x. \a \ x; x \ b\ \ (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))" using continuous_on_subset [OF contf] g_im by (metis antiderivative_continuous atLeastAtMost_iff image_subset_iff set_eq_subset) have contfg: "continuous_on {a..b} (\x. f (g x))" by (blast intro: continuous_on_compose2 contf contg) have "LBINT x. indicat_real {a..b} x *\<^sub>R g' x *\<^sub>R f (g x) = F (g b) - F (g a)" apply (rule integral_FTC_atLeastAtMost [OF \a \ b\ vector_diff_chain_within[OF v_derivg derivF, unfolded comp_def]]) apply (auto intro!: continuous_on_scaleR contg' contfg) done then have "LBINT x=a..b. g' x *\<^sub>R f (g x) = F (g b) - F (g a)" by (simp add: assms interval_integral_Icc set_lebesgue_integral_def) moreover have "LBINT y=(g a)..(g b). f y = F (g b) - F (g a)" proof (rule interval_integral_FTC_finite) show "continuous_on {min (g a) (g b)..max (g a) (g b)} f" by (rule continuous_on_subset [OF contf]) (auto simp: image_def 1) show "(F has_vector_derivative f y) (at y within {min (g a) (g b)..max (g a) (g b)})" if y: "min (g a) (g b) \ y" "y \ max (g a) (g b)" for y proof - obtain x where "a \ x" "x \ b" "y = g x" using 1 y by force then show ?thesis by (auto simp: image_def intro!: 1 has_vector_derivative_within_subset [OF derivF]) qed qed ultimately show ?thesis by simp qed (* TODO: is it possible to lift the assumption here that g' is nonnegative? *) theorem interval_integral_substitution_integrable: fixes f :: "real \ 'a::euclidean_space" and a b u v :: ereal assumes "a < b" and deriv_g: "\x. a < ereal x \ ereal x < b \ DERIV g x :> g' x" and contf: "\x. a < ereal x \ ereal x < b \ isCont f (g x)" and contg': "\x. a < ereal x \ ereal x < b \ isCont g' x" and g'_nonneg: "\x. a \ ereal x \ ereal x \ b \ 0 \ g' x" and A: "((ereal \ g \ real_of_ereal) \ A) (at_right a)" and B: "((ereal \ g \ real_of_ereal) \ B) (at_left b)" and integrable: "set_integrable lborel (einterval a b) (\x. g' x *\<^sub>R f (g x))" and integrable2: "set_integrable lborel (einterval A B) (\x. f x)" shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))" proof - obtain u l where approx [simp]: "einterval a b = (\i. {l i .. u i})" "incseq u" "decseq l" "\i. l i < u i" "\i. a < l i" "\i. u i < b" "l \ a" "u \ b" by (blast intro: einterval_Icc_approximation[OF \a < b\]) note less_imp_le [simp] have [simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have [simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) then have lessb[simp]: "\i. l i < b" using approx(4) less_eq_real_def by blast have [simp]: "\i. a < u i" by (rule order_less_trans, rule approx, auto, rule approx) have lle[simp]: "\i j. i \ j \ l j \ l i" by (rule decseqD, rule approx) have [simp]: "\i j. i \ j \ u i \ u j" by (rule incseqD, rule approx) have g_nondec [simp]: "g x \ g y" if "a < x" "x \ y" "y < b" for x y proof (rule DERIV_nonneg_imp_nondecreasing [OF \x \ y\], intro exI conjI) show "\u. x \ u \ u \ y \ (g has_real_derivative g' u) (at u)" by (meson deriv_g ereal_less_eq(3) le_less_trans less_le_trans that) show "\u. x \ u \ u \ y \ 0 \ g' u" by (meson assms(5) dual_order.trans le_ereal_le less_imp_le order_refl that) qed have "A \ B" and un: "einterval A B = (\i. {g(l i)<..i. g (l i)) \ A" using A apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (l i)" in spec, auto) hence A3: "\i. g (l i) \ A" by (intro decseq_ge, auto simp: decseq_def) have B2: "(\i. g (u i)) \ B" using B apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (u i)" in spec, auto) hence B3: "\i. g (u i) \ B" by (intro incseq_le, auto simp: incseq_def) have "ereal (g (l 0)) \ ereal (g (u 0))" by auto then show "A \ B" by (meson A3 B3 order.trans) { fix x :: real assume "A < x" and "x < B" then have "eventually (\i. ereal (g (l i)) < x \ x < ereal (g (u i))) sequentially" by (fast intro: eventually_conj order_tendstoD A2 B2) hence "\i. g (l i) < x \ x < g (u i)" by (simp add: eventually_sequentially, auto) } note AB = this show "einterval A B = (\i. {g(l i)<.. (\i. {g(l i)<..i. {g(l i)<.. einterval A B" proof (clarsimp simp add: einterval_def, intro conjI) show "\x i. \g (l i) < x; x < g (u i)\ \ A < ereal x" using A3 le_ereal_less by blast show "\x i. \g (l i) < x; x < g (u i)\ \ ereal x < B" using B3 ereal_le_less by blast qed qed qed (* finally, the main argument *) have eq1: "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)" for i apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]]) unfolding has_field_derivative_iff_has_vector_derivative[symmetric] apply (auto intro!: continuous_at_imp_continuous_on contf contg') done have "(\i. LBINT x=l i..u i. g' x *\<^sub>R f (g x)) \ (LBINT x=a..b. g' x *\<^sub>R f (g x))" apply (rule interval_integral_Icc_approx_integrable [OF \a < b\ approx]) by (rule assms) hence 2: "(\i. (LBINT y=g (l i)..g (u i). f y)) \ (LBINT x=a..b. g' x *\<^sub>R f (g x))" by (simp add: eq1) have incseq: "incseq (\i. {g (l i)<..i. set_lebesgue_integral lborel {g (l i)<.. set_lebesgue_integral lborel (einterval A B) f" unfolding un by (rule set_integral_cont_up) (use incseq integrable2 un in auto) then have "(\i. (LBINT y=g (l i)..g (u i). f y)) \ (LBINT x = A..B. f x)" by (simp add: interval_lebesgue_integral_le_eq \A \ B\) thus ?thesis by (intro LIMSEQ_unique [OF _ 2]) qed (* TODO: the last two proofs are only slightly different. Factor out common part? An alternative: make the second one the main one, and then have another lemma that says that if f is nonnegative and all the other hypotheses hold, then it is integrable. *) theorem interval_integral_substitution_nonneg: fixes f g g':: "real \ real" and a b u v :: ereal assumes "a < b" and deriv_g: "\x. a < ereal x \ ereal x < b \ DERIV g x :> g' x" and contf: "\x. a < ereal x \ ereal x < b \ isCont f (g x)" and contg': "\x. a < ereal x \ ereal x < b \ isCont g' x" and f_nonneg: "\x. a < ereal x \ ereal x < b \ 0 \ f (g x)" (* TODO: make this AE? *) and g'_nonneg: "\x. a \ ereal x \ ereal x \ b \ 0 \ g' x" and A: "((ereal \ g \ real_of_ereal) \ A) (at_right a)" and B: "((ereal \ g \ real_of_ereal) \ B) (at_left b)" and integrable_fg: "set_integrable lborel (einterval a b) (\x. f (g x) * g' x)" shows "set_integrable lborel (einterval A B) f" "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))" proof - - from einterval_Icc_approximation[OF \a < b\] guess u l . note approx [simp] = this - note less_imp_le [simp] + from einterval_Icc_approximation[OF \a < b\] obtain u l where approx [simp]: + "einterval a b = (\i. {l i..u i})" + "incseq u" + "decseq l" + "\i. l i < u i" + "\i. a < ereal (l i)" + "\i. ereal (u i) < b" + "(\x. ereal (l x)) \ a" + "(\x. ereal (u x)) \ b" by this auto have aless[simp]: "\x i. l i \ x \ a < ereal x" by (rule order_less_le_trans, rule approx, force) have lessb[simp]: "\x i. x \ u i \ ereal x < b" by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx) have llb[simp]: "\i. l i < b" using lessb approx(4) less_eq_real_def by blast have alu[simp]: "\i. a < u i" by (rule order_less_trans, rule approx, auto, rule approx) have [simp]: "\i j. i \ j \ l j \ l i" by (rule decseqD, rule approx) have uleu[simp]: "\i j. i \ j \ u i \ u j" by (rule incseqD, rule approx) have g_nondec [simp]: "g x \ g y" if "a < x" "x \ y" "y < b" for x y proof (rule DERIV_nonneg_imp_nondecreasing [OF \x \ y\], intro exI conjI) show "\u. x \ u \ u \ y \ (g has_real_derivative g' u) (at u)" by (meson deriv_g ereal_less_eq(3) le_less_trans less_le_trans that) show "\u. x \ u \ u \ y \ 0 \ g' u" by (meson g'_nonneg less_ereal.simps(1) less_trans not_less that) qed have "A \ B" and un: "einterval A B = (\i. {g(l i)<..i. g (l i)) \ A" using A apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (l i)" in spec, auto) hence A3: "\i. g (l i) \ A" by (intro decseq_ge, auto simp: decseq_def) have B2: "(\i. g (u i)) \ B" using B apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def) by (drule_tac x = "\i. ereal (u i)" in spec, auto) hence B3: "\i. g (u i) \ B" by (intro incseq_le, auto simp: incseq_def) have "ereal (g (l 0)) \ ereal (g (u 0))" - by auto + by (auto simp: less_imp_le) then show "A \ B" by (meson A3 B3 order.trans) { fix x :: real assume "A < x" and "x < B" then have "eventually (\i. ereal (g (l i)) < x \ x < ereal (g (u i))) sequentially" by (fast intro: eventually_conj order_tendstoD A2 B2) hence "\i. g (l i) < x \ x < g (u i)" by (simp add: eventually_sequentially, auto) } note AB = this show "einterval A B = (\i. {g(l i)<.. (\i. {g (l i)<..i. {g (l i)<.. einterval A B" apply (clarsimp simp: einterval_def, intro conjI) using A3 le_ereal_less apply blast using B3 ereal_le_less by blast qed qed (* finally, the main argument *) have eq1: "(LBINT x=l i.. u i. (f (g x) * g' x)) = (LBINT y=g (l i)..g (u i). f y)" for i proof - have "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)" apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]]) unfolding has_field_derivative_iff_has_vector_derivative[symmetric] - apply (auto intro!: continuous_at_imp_continuous_on contf contg') + apply (auto simp: less_imp_le intro!: continuous_at_imp_continuous_on contf contg') done then show ?thesis by (simp add: ac_simps) qed have incseq: "incseq (\i. {g (l i)<..c \ l i. c \ u i \ x = g c" if "g (l i) \ x" "x \ g (u i)" for x i proof - have "continuous_on {l i..u i} g" by (force intro!: DERIV_isCont deriv_g continuous_at_imp_continuous_on) with that show ?thesis using IVT' [of g] approx(4) dual_order.strict_implies_order by blast qed have "continuous_on {g (l i)..g (u i)} f" for i apply (intro continuous_intros continuous_at_imp_continuous_on) using contf img by force then have int_f: "\i. set_integrable lborel {g (l i)<..i. {g (l i)<..i. LBINT x=l i..u i. f (g x) * g' x) \ ?l" by (intro assms interval_integral_Icc_approx_integrable [OF \a < b\ approx]) hence "(\i. (LBINT y=g (l i)..g (u i). f y)) \ ?l" by (simp add: eq1) then show "(\i. set_lebesgue_integral lborel {g (l i)<.. ?l" - unfolding interval_lebesgue_integral_def by auto + unfolding interval_lebesgue_integral_def by (auto simp: less_imp_le) have "\x i. g (l i) \ x \ x \ g (u i) \ 0 \ f x" using aless f_nonneg img lessb by blast then show "\x i. x \ {g (l i)<.. 0 \ f x" using less_eq_real_def by auto qed (auto simp: greaterThanLessThan_borel) thus "set_integrable lborel (einterval A B) f" by (simp add: un) have "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))" proof (rule interval_integral_substitution_integrable) show "set_integrable lborel (einterval a b) (\x. g' x *\<^sub>R f (g x))" using integrable_fg by (simp add: ac_simps) qed fact+ then show "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))" by (simp add: ac_simps) qed syntax "_complex_lebesgue_borel_integral" :: "pttrn \ real \ complex" ("(2CLBINT _. _)" [0,60] 60) translations "CLBINT x. f" == "CONST complex_lebesgue_integral CONST lborel (\x. f)" syntax "_complex_set_lebesgue_borel_integral" :: "pttrn \ real set \ real \ complex" ("(3CLBINT _:_. _)" [0,60,61] 60) translations "CLBINT x:A. f" == "CONST complex_set_lebesgue_integral CONST lborel A (\x. f)" abbreviation complex_interval_lebesgue_integral :: "real measure \ ereal \ ereal \ (real \ complex) \ complex" where "complex_interval_lebesgue_integral M a b f \ interval_lebesgue_integral M a b f" abbreviation complex_interval_lebesgue_integrable :: "real measure \ ereal \ ereal \ (real \ complex) \ bool" where "complex_interval_lebesgue_integrable M a b f \ interval_lebesgue_integrable M a b f" syntax "_ascii_complex_interval_lebesgue_borel_integral" :: "pttrn \ ereal \ ereal \ real \ complex" ("(4CLBINT _=_.._. _)" [0,60,60,61] 60) translations "CLBINT x=a..b. f" == "CONST complex_interval_lebesgue_integral CONST lborel a b (\x. f)" proposition interval_integral_norm: fixes f :: "real \ 'a :: {banach, second_countable_topology}" shows "interval_lebesgue_integrable lborel a b f \ a \ b \ norm (LBINT t=a..b. f t) \ LBINT t=a..b. norm (f t)" using integral_norm_bound[of lborel "\x. indicator (einterval a b) x *\<^sub>R f x"] by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def) proposition interval_integral_norm2: "interval_lebesgue_integrable lborel a b f \ norm (LBINT t=a..b. f t) \ \LBINT t=a..b. norm (f t)\" proof (induct a b rule: linorder_wlog) case (sym a b) then show ?case by (simp add: interval_integral_endpoints_reverse[of a b] interval_integrable_endpoints_reverse[of a b]) next case (le a b) then have "\LBINT t=a..b. norm (f t)\ = LBINT t=a..b. norm (f t)" using integrable_norm[of lborel "\x. indicator (einterval a b) x *\<^sub>R f x"] by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def intro!: integral_nonneg_AE abs_of_nonneg) then show ?case using le by (simp add: interval_integral_norm) qed (* TODO: should we have a library of facts like these? *) lemma integral_cos: "t \ 0 \ LBINT x=a..b. cos (t * x) = sin (t * b) / t - sin (t * a) / t" apply (intro interval_integral_FTC_finite continuous_intros) by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative[symmetric]) end diff --git a/src/HOL/Analysis/Measure_Space.thy b/src/HOL/Analysis/Measure_Space.thy --- a/src/HOL/Analysis/Measure_Space.thy +++ b/src/HOL/Analysis/Measure_Space.thy @@ -1,3884 +1,3892 @@ (* Title: HOL/Analysis/Measure_Space.thy Author: Lawrence C Paulson Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Measure Spaces\ theory Measure_Space imports Measurable "HOL-Library.Extended_Nonnegative_Real" begin subsection\<^marker>\tag unimportant\ "Relate extended reals and the indicator function" lemma suminf_cmult_indicator: fixes f :: "nat \ ennreal" assumes "disjoint_family A" "x \ A i" shows "(\n. f n * indicator (A n) x) = f i" proof - have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)" using \x \ A i\ assms unfolding disjoint_family_on_def indicator_def by auto then have "\n. (\jn. n \ UNIV \ (if i < n then f i else 0) \ y" from this[of "Suc i"] show "f i \ y" by auto qed (insert assms, simp) ultimately show ?thesis using assms by (subst suminf_eq_SUP) (auto simp: indicator_def) qed lemma suminf_indicator: assumes "disjoint_family A" shows "(\n. indicator (A n) x :: ennreal) = indicator (\i. A i) x" proof cases assume *: "x \ (\i. A i)" then obtain i where "x \ A i" by auto from suminf_cmult_indicator[OF assms(1), OF \x \ A i\, of "\k. 1"] show ?thesis using * by simp qed simp lemma sum_indicator_disjoint_family: fixes f :: "'d \ 'e::semiring_1" assumes d: "disjoint_family_on A P" and "x \ A j" and "finite P" and "j \ P" shows "(\i\P. f i * indicator (A i) x) = f j" proof - have "P \ {i. x \ A i} = {j}" using d \x \ A j\ \j \ P\ unfolding disjoint_family_on_def by auto with \finite P\ show ?thesis by (simp add: indicator_def) qed text \ The type for emeasure spaces is already defined in \<^theory>\HOL-Analysis.Sigma_Algebra\, as it is also used to represent sigma algebras (with an arbitrary emeasure). \ subsection\<^marker>\tag unimportant\ "Extend binary sets" lemma LIMSEQ_binaryset: assumes f: "f {} = 0" shows "(\n. \i f A + f B" proof - have "(\n. \i < Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" proof fix n show "(\i < Suc (Suc n). f (binaryset A B i)) = f A + f B" by (induct n) (auto simp add: binaryset_def f) qed moreover have "... \ f A + f B" by (rule tendsto_const) ultimately have "(\n. \i< Suc (Suc n). f (binaryset A B i)) \ f A + f B" by metis hence "(\n. \i< n+2. f (binaryset A B i)) \ f A + f B" by simp thus ?thesis by (rule LIMSEQ_offset [where k=2]) qed lemma binaryset_sums: assumes f: "f {} = 0" shows "(\n. f (binaryset A B n)) sums (f A + f B)" by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan) lemma suminf_binaryset_eq: fixes f :: "'a set \ 'b::{comm_monoid_add, t2_space}" shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) subsection\<^marker>\tag unimportant\ \Properties of a premeasure \<^term>\\\\ text \ The definitions for \<^const>\positive\ and \<^const>\countably_additive\ should be here, by they are necessary to define \<^typ>\'a measure\ in \<^theory>\HOL-Analysis.Sigma_Algebra\. \ definition subadditive where "subadditive M f \ (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" lemma subadditiveD: "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) definition countably_subadditive where "countably_subadditive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (f (\i. A i) \ (\i. f (A i))))" lemma (in ring_of_sets) countably_subadditive_subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" and cs: "countably_subadditive M f" shows "subadditive M f" proof (auto simp add: subadditive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) \ (\ n. f (binaryset x y n))" using cs by (auto simp add: countably_subadditive_def) hence "{x,y,{}} \ M \ x \ y \ M \ f (x \ y) \ (\ n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) \ f x + f y" using f x y by (auto simp add: Un o_def suminf_binaryset_eq positive_def) qed definition additive where "additive M \ \ (\x\M. \y\M. x \ y = {} \ \ (x \ y) = \ x + \ y)" definition increasing where "increasing M \ \ (\x\M. \y\M. x \ y \ \ x \ \ y)" lemma positiveD1: "positive M f \ f {} = 0" by (auto simp: positive_def) lemma positiveD_empty: "positive M f \ f {} = 0" by (auto simp add: positive_def) lemma additiveD: "additive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) = f x + f y" by (auto simp add: additive_def) lemma increasingD: "increasing M f \ x \ y \ x\M \ y\M \ f x \ f y" by (auto simp add: increasing_def) lemma countably_additiveI[case_names countably]: "(\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (\i. f (A i)) = f (\i. A i)) \ countably_additive M f" by (simp add: countably_additive_def) lemma (in ring_of_sets) disjointed_additive: assumes f: "positive M f" "additive M f" and A: "range A \ M" "incseq A" shows "(\i\n. f (disjointed A i)) = f (A n)" proof (induct n) case (Suc n) then have "(\i\Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))" by simp also have "\ = f (A n \ disjointed A (Suc n))" using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono) also have "A n \ disjointed A (Suc n) = A (Suc n)" using \incseq A\ by (auto dest: incseq_SucD simp: disjointed_mono) finally show ?case . qed simp lemma (in ring_of_sets) additive_sum: fixes A:: "'i \ 'a set" assumes f: "positive M f" and ad: "additive M f" and "finite S" and A: "A`S \ M" and disj: "disjoint_family_on A S" shows "(\i\S. f (A i)) = f (\i\S. A i)" using \finite S\ disj A proof induct case empty show ?case using f by (simp add: positive_def) next case (insert s S) then have "A s \ (\i\S. A i) = {}" by (auto simp add: disjoint_family_on_def neq_iff) moreover have "A s \ M" using insert by blast moreover have "(\i\S. A i) \ M" using insert \finite S\ by auto ultimately have "f (A s \ (\i\S. A i)) = f (A s) + f(\i\S. A i)" using ad UNION_in_sets A by (auto simp add: additive_def) with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] by (auto simp add: additive_def subset_insertI) qed lemma (in ring_of_sets) additive_increasing: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and addf: "additive M f" shows "increasing M f" proof (auto simp add: increasing_def) fix x y assume xy: "x \ M" "y \ M" "x \ y" then have "y - x \ M" by auto then have "f x + 0 \ f x + f (y-x)" by (intro add_left_mono zero_le) also have "... = f (x \ (y-x))" using addf by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2)) also have "... = f y" by (metis Un_Diff_cancel Un_absorb1 xy(3)) finally show "f x \ f y" by simp qed lemma (in ring_of_sets) subadditive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and A: "A`S \ M" and S: "finite S" shows "f (\i\S. A i) \ (\i\S. f (A i))" using S A proof (induct S) case empty thus ?case using f by (auto simp: positive_def) next case (insert x F) hence in_M: "A x \ M" "(\i\F. A i) \ M" "(\i\F. A i) - A x \ M" using A by force+ have subs: "(\i\F. A i) - A x \ (\i\F. A i)" by auto have "(\i\(insert x F). A i) = A x \ ((\i\F. A i) - A x)" by auto hence "f (\i\(insert x F). A i) = f (A x \ ((\i\F. A i) - A x))" by simp also have "\ = f (A x) + f ((\i\F. A i) - A x)" using f(2) by (rule additiveD) (insert in_M, auto) also have "\ \ f (A x) + f (\i\F. A i)" using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono) also have "\ \ f (A x) + (\i\F. f (A i))" using insert by (auto intro: add_left_mono) finally show "f (\i\(insert x F). A i) \ (\i\(insert x F). f (A i))" using insert by simp qed lemma (in ring_of_sets) countably_additive_additive: fixes f :: "'a set \ ennreal" assumes posf: "positive M f" and ca: "countably_additive M f" shows "additive M f" proof (auto simp add: additive_def) fix x y assume x: "x \ M" and y: "y \ M" and "x \ y = {}" hence "disjoint_family (binaryset x y)" by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ M \ (\i. binaryset x y i) \ M \ f (\i. binaryset x y i) = (\ n. f (binaryset x y n))" using ca by (simp add: countably_additive_def) hence "{x,y,{}} \ M \ x \ y \ M \ f (x \ y) = (\n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) = f x + f y" using posf x y by (auto simp add: Un suminf_binaryset_eq positive_def) qed lemma (in algebra) increasing_additive_bound: fixes A:: "nat \ 'a set" and f :: "'a set \ ennreal" assumes f: "positive M f" and ad: "additive M f" and inc: "increasing M f" and A: "range A \ M" and disj: "disjoint_family A" shows "(\i. f (A i)) \ f \" proof (safe intro!: suminf_le_const) fix N note disj_N = disjoint_family_on_mono[OF _ disj, of "{..ii\{.. f \" using space_closed A by (intro increasingD[OF inc] finite_UN) auto finally show "(\i f \" by simp qed (insert f A, auto simp: positive_def) lemma (in ring_of_sets) countably_additiveI_finite: fixes \ :: "'a set \ ennreal" assumes "finite \" "positive M \" "additive M \" shows "countably_additive M \" proof (rule countably_additiveI) fix F :: "nat \ 'a set" assume F: "range F \ M" "(\i. F i) \ M" and disj: "disjoint_family F" have "\i\{i. F i \ {}}. \x. x \ F i" by auto from bchoice[OF this] obtain f where f: "\i. F i \ {} \ f i \ F i" by auto have inj_f: "inj_on f {i. F i \ {}}" proof (rule inj_onI, simp) fix i j a b assume *: "f i = f j" "F i \ {}" "F j \ {}" then have "f i \ F i" "f j \ F j" using f by force+ with disj * show "i = j" by (auto simp: disjoint_family_on_def) qed have "finite (\i. F i)" by (metis F(2) assms(1) infinite_super sets_into_space) have F_subset: "{i. \ (F i) \ 0} \ {i. F i \ {}}" by (auto simp: positiveD_empty[OF \positive M \\]) moreover have fin_not_empty: "finite {i. F i \ {}}" proof (rule finite_imageD) from f have "f`{i. F i \ {}} \ (\i. F i)" by auto then show "finite (f`{i. F i \ {}})" by (rule finite_subset) fact qed fact ultimately have fin_not_0: "finite {i. \ (F i) \ 0}" by (rule finite_subset) have disj_not_empty: "disjoint_family_on F {i. F i \ {}}" using disj by (auto simp: disjoint_family_on_def) from fin_not_0 have "(\i. \ (F i)) = (\i | \ (F i) \ 0. \ (F i))" by (rule suminf_finite) auto also have "\ = (\i | F i \ {}. \ (F i))" using fin_not_empty F_subset by (rule sum.mono_neutral_left) auto also have "\ = \ (\i\{i. F i \ {}}. F i)" using \positive M \\ \additive M \\ fin_not_empty disj_not_empty F by (intro additive_sum) auto also have "\ = \ (\i. F i)" by (rule arg_cong[where f=\]) auto finally show "(\i. \ (F i)) = \ (\i. F i)" . qed lemma (in ring_of_sets) countably_additive_iff_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "countably_additive M f \ (\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i))" unfolding countably_additive_def proof safe assume count_sum: "\A. range A \ M \ disjoint_family A \ \(A ` UNIV) \ M \ (\i. f (A i)) = f (\(A ` UNIV))" fix A :: "nat \ 'a set" assume A: "range A \ M" "incseq A" "(\i. A i) \ M" then have dA: "range (disjointed A) \ M" by (auto simp: range_disjointed_sets) with count_sum[THEN spec, of "disjointed A"] A(3) have f_UN: "(\i. f (disjointed A i)) = f (\i. A i)" by (auto simp: UN_disjointed_eq disjoint_family_disjointed) moreover have "(\n. (\i (\i. f (disjointed A i))" using f(1)[unfolded positive_def] dA by (auto intro!: summable_LIMSEQ) from LIMSEQ_Suc[OF this] have "(\n. (\i\n. f (disjointed A i))) \ (\i. f (disjointed A i))" unfolding lessThan_Suc_atMost . moreover have "\n. (\i\n. f (disjointed A i)) = f (A n)" using disjointed_additive[OF f A(1,2)] . ultimately show "(\i. f (A i)) \ f (\i. A i)" by simp next assume cont: "\A. range A \ M \ incseq A \ (\i. A i) \ M \ (\i. f (A i)) \ f (\i. A i)" fix A :: "nat \ 'a set" assume A: "range A \ M" "disjoint_family A" "(\i. A i) \ M" have *: "(\n. (\ii. A i)" by auto have "(\n. f (\i f (\i. A i)" proof (unfold *[symmetric], intro cont[rule_format]) show "range (\i. \i M" "(\i. \i M" using A * by auto qed (force intro!: incseq_SucI) moreover have "\n. f (\iii. f (A i)) sums f (\i. A i)" unfolding sums_def by simp from sums_unique[OF this] show "(\i. f (A i)) = f (\i. A i)" by simp qed lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" shows "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i)) \ (\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0)" proof safe assume cont: "(\A. range A \ M \ decseq A \ (\i. A i) \ M \ (\i. f (A i) \ \) \ (\i. f (A i)) \ f (\i. A i))" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) = {}" "\i. f (A i) \ \" with cont[THEN spec, of A] show "(\i. f (A i)) \ 0" using \positive M f\[unfolded positive_def] by auto next assume cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i) \ \) \ (\i. f (A i)) \ 0" fix A :: "nat \ 'a set" assume A: "range A \ M" "decseq A" "(\i. A i) \ M" "\i. f (A i) \ \" have f_mono: "\a b. a \ M \ b \ M \ a \ b \ f a \ f b" using additive_increasing[OF f] unfolding increasing_def by simp have decseq_fA: "decseq (\i. f (A i))" using A by (auto simp: decseq_def intro!: f_mono) have decseq: "decseq (\i. A i - (\i. A i))" using A by (auto simp: decseq_def) then have decseq_f: "decseq (\i. f (A i - (\i. A i)))" using A unfolding decseq_def by (auto intro!: f_mono Diff) have "f (\x. A x) \ f (A 0)" using A by (auto intro!: f_mono) then have f_Int_fin: "f (\x. A x) \ \" using A by (auto simp: top_unique) { fix i have "f (A i - (\i. A i)) \ f (A i)" using A by (auto intro!: f_mono) then have "f (A i - (\i. A i)) \ \" using A by (auto simp: top_unique) } note f_fin = this have "(\i. f (A i - (\i. A i))) \ 0" proof (intro cont[rule_format, OF _ decseq _ f_fin]) show "range (\i. A i - (\i. A i)) \ M" "(\i. A i - (\i. A i)) = {}" using A by auto qed from INF_Lim[OF decseq_f this] have "(INF n. f (A n - (\i. A i))) = 0" . moreover have "(INF n. f (\i. A i)) = f (\i. A i)" by auto ultimately have "(INF n. f (A n - (\i. A i)) + f (\i. A i)) = 0 + f (\i. A i)" using A(4) f_fin f_Int_fin by (subst INF_ennreal_add_const) (auto simp: decseq_f) moreover { fix n have "f (A n - (\i. A i)) + f (\i. A i) = f ((A n - (\i. A i)) \ (\i. A i))" using A by (subst f(2)[THEN additiveD]) auto also have "(A n - (\i. A i)) \ (\i. A i) = A n" by auto finally have "f (A n - (\i. A i)) + f (\i. A i) = f (A n)" . } ultimately have "(INF n. f (A n)) = f (\i. A i)" by simp with LIMSEQ_INF[OF decseq_fA] show "(\i. f (A i)) \ f (\i. A i)" by simp qed lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" assumes A: "range A \ M" "incseq A" "(\i. A i) \ M" shows "(\i. f (A i)) \ f (\i. A i)" proof - from A have "(\i. f ((\i. A i) - A i)) \ 0" by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) moreover { fix i have "f ((\i. A i) - A i \ A i) = f ((\i. A i) - A i) + f (A i)" using A by (intro f(2)[THEN additiveD]) auto also have "((\i. A i) - A i) \ A i = (\i. A i)" by auto finally have "f ((\i. A i) - A i) = f (\i. A i) - f (A i)" using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) } moreover have "\\<^sub>F i in sequentially. f (A i) \ f (\i. A i)" using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\i. A i"] A by (auto intro!: always_eventually simp: subset_eq) ultimately show "(\i. f (A i)) \ f (\i. A i)" by (auto intro: ennreal_tendsto_const_minus) qed lemma (in ring_of_sets) empty_continuous_imp_countably_additive: fixes f :: "'a set \ ennreal" assumes f: "positive M f" "additive M f" and fin: "\A\M. f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) \ 0" shows "countably_additive M f" using countably_additive_iff_continuous_from_below[OF f] using empty_continuous_imp_continuous_from_below[OF f fin] cont by blast subsection\<^marker>\tag unimportant\ \Properties of \<^const>\emeasure\\ lemma emeasure_positive: "positive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma emeasure_empty[simp, intro]: "emeasure M {} = 0" using emeasure_positive[of M] by (simp add: positive_def) lemma emeasure_single_in_space: "emeasure M {x} \ 0 \ x \ space M" using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2]) lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma suminf_emeasure: "range A \ sets M \ disjoint_family A \ (\i. emeasure M (A i)) = emeasure M (\i. A i)" using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M] by (simp add: countably_additive_def) lemma sums_emeasure: "disjoint_family F \ (\i. F i \ sets M) \ (\i. emeasure M (F i)) sums emeasure M (\i. F i)" unfolding sums_iff by (intro conjI suminf_emeasure) auto lemma emeasure_additive: "additive (sets M) (emeasure M)" by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive) lemma plus_emeasure: "a \ sets M \ b \ sets M \ a \ b = {} \ emeasure M a + emeasure M b = emeasure M (a \ b)" using additiveD[OF emeasure_additive] .. lemma emeasure_Un: "A \ sets M \ B \ sets M \ emeasure M (A \ B) = emeasure M A + emeasure M (B - A)" using plus_emeasure[of A M "B - A"] by auto lemma emeasure_Un_Int: assumes "A \ sets M" "B \ sets M" shows "emeasure M A + emeasure M B = emeasure M (A \ B) + emeasure M (A \ B)" proof - have "A = (A-B) \ (A \ B)" by auto then have "emeasure M A = emeasure M (A-B) + emeasure M (A \ B)" by (metis Diff_Diff_Int Diff_disjoint assms plus_emeasure sets.Diff) moreover have "A \ B = (A-B) \ B" by auto then have "emeasure M (A \ B) = emeasure M (A-B) + emeasure M B" by (metis Diff_disjoint Int_commute assms plus_emeasure sets.Diff) ultimately show ?thesis by (metis add.assoc add.commute) qed lemma sum_emeasure: "F`I \ sets M \ disjoint_family_on F I \ finite I \ (\i\I. emeasure M (F i)) = emeasure M (\i\I. F i)" by (metis sets.additive_sum emeasure_positive emeasure_additive) lemma emeasure_mono: "a \ b \ b \ sets M \ emeasure M a \ emeasure M b" by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD) lemma emeasure_space: "emeasure M A \ emeasure M (space M)" by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le) lemma emeasure_Diff: assumes finite: "emeasure M B \ \" and [measurable]: "A \ sets M" "B \ sets M" and "B \ A" shows "emeasure M (A - B) = emeasure M A - emeasure M B" proof - have "(A - B) \ B = A" using \B \ A\ by auto then have "emeasure M A = emeasure M ((A - B) \ B)" by simp also have "\ = emeasure M (A - B) + emeasure M B" by (subst plus_emeasure[symmetric]) auto finally show "emeasure M (A - B) = emeasure M A - emeasure M B" using finite by simp qed lemma emeasure_compl: "s \ sets M \ emeasure M s \ \ \ emeasure M (space M - s) = emeasure M (space M) - emeasure M s" by (rule emeasure_Diff) (auto dest: sets.sets_into_space) lemma Lim_emeasure_incseq: "range A \ sets M \ incseq A \ (\i. (emeasure M (A i))) \ emeasure M (\i. A i)" using emeasure_countably_additive by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive emeasure_additive) lemma incseq_emeasure: assumes "range B \ sets M" "incseq B" shows "incseq (\i. emeasure M (B i))" using assms by (auto simp: incseq_def intro!: emeasure_mono) lemma SUP_emeasure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(SUP n. emeasure M (A n)) = emeasure M (\i. A i)" using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A] by (simp add: LIMSEQ_unique) lemma decseq_emeasure: assumes "range B \ sets M" "decseq B" shows "decseq (\i. emeasure M (B i))" using assms by (auto simp: decseq_def intro!: emeasure_mono) lemma INF_emeasure_decseq: assumes A: "range A \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - have le_MI: "emeasure M (\i. A i) \ emeasure M (A 0)" using A by (auto intro!: emeasure_mono) hence *: "emeasure M (\i. A i) \ \" using finite[of 0] by (auto simp: top_unique) have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))" by (simp add: ennreal_INF_const_minus) also have "\ = (SUP n. emeasure M (A 0 - A n))" using A finite \decseq A\[unfolded decseq_def] by (subst emeasure_Diff) auto also have "\ = emeasure M (\i. A 0 - A i)" proof (rule SUP_emeasure_incseq) show "range (\n. A 0 - A n) \ sets M" using A by auto show "incseq (\n. A 0 - A n)" using \decseq A\ by (auto simp add: incseq_def decseq_def) qed also have "\ = emeasure M (A 0) - emeasure M (\i. A i)" using A finite * by (simp, subst emeasure_Diff) auto finally show ?thesis by (rule ennreal_minus_cancel[rotated 3]) (insert finite A, auto intro: INF_lower emeasure_mono) qed lemma INF_emeasure_decseq': assumes A: "\i. A i \ sets M" and "decseq A" and finite: "\i. emeasure M (A i) \ \" shows "(INF n. emeasure M (A n)) = emeasure M (\i. A i)" proof - from finite obtain i where i: "emeasure M (A i) < \" by (auto simp: less_top) have fin: "i \ j \ emeasure M (A j) < \" for j by (rule le_less_trans[OF emeasure_mono i]) (auto intro!: decseqD[OF \decseq A\] A) have "(INF n. emeasure M (A n)) = (INF n. emeasure M (A (n + i)))" proof (rule INF_eq) show "\j\UNIV. emeasure M (A (j + i)) \ emeasure M (A i')" for i' by (intro bexI[of _ i'] emeasure_mono decseqD[OF \decseq A\] A) auto qed auto also have "\ = emeasure M (INF n. (A (n + i)))" using A \decseq A\ fin by (intro INF_emeasure_decseq) (auto simp: decseq_def less_top) also have "(INF n. (A (n + i))) = (INF n. A n)" by (meson INF_eq UNIV_I assms(2) decseqD le_add1) finally show ?thesis . qed lemma emeasure_INT_decseq_subset: fixes F :: "nat \ 'a set" assumes I: "I \ {}" and F: "\i j. i \ I \ j \ I \ i \ j \ F j \ F i" assumes F_sets[measurable]: "\i. i \ I \ F i \ sets M" and fin: "\i. i \ I \ emeasure M (F i) \ \" shows "emeasure M (\i\I. F i) = (INF i\I. emeasure M (F i))" proof cases assume "finite I" have "(\i\I. F i) = F (Max I)" using I \finite I\ by (intro antisym INF_lower INF_greatest F) auto moreover have "(INF i\I. emeasure M (F i)) = emeasure M (F (Max I))" using I \finite I\ by (intro antisym INF_lower INF_greatest F emeasure_mono) auto ultimately show ?thesis by simp next assume "infinite I" define L where "L n = (LEAST i. i \ I \ i \ n)" for n have L: "L n \ I \ n \ L n" for n unfolding L_def proof (rule LeastI_ex) show "\x. x \ I \ n \ x" using \infinite I\ finite_subset[of I "{..< n}"] by (rule_tac ccontr) (auto simp: not_le) qed have L_eq[simp]: "i \ I \ L i = i" for i unfolding L_def by (intro Least_equality) auto have L_mono: "i \ j \ L i \ L j" for i j using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def) have "emeasure M (\i. F (L i)) = (INF i. emeasure M (F (L i)))" proof (intro INF_emeasure_decseq[symmetric]) show "decseq (\i. F (L i))" using L by (intro antimonoI F L_mono) auto qed (insert L fin, auto) also have "\ = (INF i\I. emeasure M (F i))" proof (intro antisym INF_greatest) show "i \ I \ (INF i. emeasure M (F (L i))) \ emeasure M (F i)" for i by (intro INF_lower2[of i]) auto qed (insert L, auto intro: INF_lower) also have "(\i. F (L i)) = (\i\I. F i)" proof (intro antisym INF_greatest) show "i \ I \ (\i. F (L i)) \ F i" for i by (intro INF_lower2[of i]) auto qed (insert L, auto) finally show ?thesis . qed lemma Lim_emeasure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\i. emeasure M (A i)) \ emeasure M (\i. A i)" using LIMSEQ_INF[OF decseq_emeasure, OF A] using INF_emeasure_decseq[OF A fin] by simp lemma emeasure_lfp'[consumes 1, case_names cont measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M {x\space M. lfp F x} = (SUP i. emeasure M {x\space M. (F ^^ i) (\x. False) x})" proof - have "emeasure M {x\space M. lfp F x} = emeasure M (\i. {x\space M. (F ^^ i) (\x. False) x})" using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure]) moreover { fix i from \P M\ have "{x\space M. (F ^^ i) (\x. False) x} \ sets M" by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) } moreover have "incseq (\i. {x\space M. (F ^^ i) (\x. False) x})" proof (rule incseq_SucI) fix i have "(F ^^ i) (\x. False) \ (F ^^ (Suc i)) (\x. False)" proof (induct i) case 0 show ?case by (simp add: le_fun_def) next case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto qed then show "{x \ space M. (F ^^ i) (\x. False) x} \ {x \ space M. (F ^^ Suc i) (\x. False) x}" by auto qed ultimately show ?thesis by (subst SUP_emeasure_incseq) auto qed lemma emeasure_lfp: assumes [simp]: "\s. sets (M s) = sets N" assumes cont: "sup_continuous F" "sup_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ P \ lfp F \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" shows "emeasure (M s) {x\space N. lfp F x} = lfp f s" proof (subst lfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric]) fix C assume "incseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (SUP i. C i) x}) = (SUP i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding SUP_apply[abs_def] by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont) lemma emeasure_subadditive_finite: "finite I \ A ` I \ sets M \ emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto lemma emeasure_subadditive: "A \ sets M \ B \ sets M \ emeasure M (A \ B) \ emeasure M A + emeasure M B" using emeasure_subadditive_finite[of "{True, False}" "\True \ A | False \ B" M] by simp lemma emeasure_subadditive_countably: assumes "range f \ sets M" shows "emeasure M (\i. f i) \ (\i. emeasure M (f i))" proof - have "emeasure M (\i. f i) = emeasure M (\i. disjointed f i)" unfolding UN_disjointed_eq .. also have "\ = (\i. emeasure M (disjointed f i))" using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"] by (simp add: disjoint_family_disjointed comp_def) also have "\ \ (\i. emeasure M (f i))" using sets.range_disjointed_sets[OF assms] assms by (auto intro!: suminf_le emeasure_mono disjointed_subset) finally show ?thesis . qed lemma emeasure_insert: assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A" proof - have "{x} \ A = {}" using \x \ A\ by auto from plus_emeasure[OF sets this] show ?thesis by simp qed lemma emeasure_insert_ne: "A \ {} \ {x} \ sets M \ A \ sets M \ x \ A \ emeasure M (insert x A) = emeasure M {x} + emeasure M A" by (rule emeasure_insert) lemma emeasure_eq_sum_singleton: assumes "finite S" "\x. x \ S \ {x} \ sets M" shows "emeasure M S = (\x\S. emeasure M {x})" using sum_emeasure[of "\x. {x}" S M] assms by (auto simp: disjoint_family_on_def subset_eq) lemma sum_emeasure_cover: assumes "finite S" and "A \ sets M" and br_in_M: "B ` S \ sets M" assumes A: "A \ (\i\S. B i)" assumes disj: "disjoint_family_on B S" shows "emeasure M A = (\i\S. emeasure M (A \ (B i)))" proof - have "(\i\S. emeasure M (A \ (B i))) = emeasure M (\i\S. A \ (B i))" proof (rule sum_emeasure) show "disjoint_family_on (\i. A \ B i) S" using \disjoint_family_on B S\ unfolding disjoint_family_on_def by auto qed (insert assms, auto) also have "(\i\S. A \ (B i)) = A" using A by auto finally show ?thesis by simp qed lemma emeasure_eq_0: "N \ sets M \ emeasure M N = 0 \ K \ N \ emeasure M K = 0" by (metis emeasure_mono order_eq_iff zero_le) lemma emeasure_UN_eq_0: assumes "\i::nat. emeasure M (N i) = 0" and "range N \ sets M" shows "emeasure M (\i. N i) = 0" proof - have "emeasure M (\i. N i) \ 0" using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp then show ?thesis by (auto intro: antisym zero_le) qed lemma measure_eqI_finite: assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A" assumes eq: "\a. a \ A \ emeasure M {a} = emeasure N {a}" shows "M = N" proof (rule measure_eqI) fix X assume "X \ sets M" then have X: "X \ A" by auto then have "emeasure M X = (\a\X. emeasure M {a})" using \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) also have "\ = (\a\X. emeasure N {a})" using X eq by (auto intro!: sum.cong) also have "\ = emeasure N X" using X \finite A\ by (subst emeasure_eq_sum_singleton) (auto dest: finite_subset) finally show "emeasure M X = emeasure N X" . qed simp lemma measure_eqI_generator_eq: fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \ 'a set" assumes "Int_stable E" "E \ Pow \" and eq: "\X. X \ E \ emeasure M X = emeasure N X" and M: "sets M = sigma_sets \ E" and N: "sets N = sigma_sets \ E" and A: "range A \ E" "(\i. A i) = \" "\i. emeasure M (A i) \ \" shows "M = N" proof - let ?\ = "emeasure M" and ?\ = "emeasure N" interpret S: sigma_algebra \ "sigma_sets \ E" by (rule sigma_algebra_sigma_sets) fact have "space M = \" using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \sets M = sigma_sets \ E\ by blast { fix F D assume "F \ E" and "?\ F \ \" then have [intro]: "F \ sigma_sets \ E" by auto have "?\ F \ \" using \?\ F \ \\ \F \ E\ eq by simp assume "D \ sets M" with \Int_stable E\ \E \ Pow \\ have "emeasure M (F \ D) = emeasure N (F \ D)" unfolding M proof (induct rule: sigma_sets_induct_disjoint) case (basic A) then have "F \ A \ E" using \Int_stable E\ \F \ E\ by (auto simp: Int_stable_def) then show ?case using eq by auto next case empty then show ?case by simp next case (compl A) then have **: "F \ (\ - A) = F - (F \ A)" and [intro]: "F \ A \ sigma_sets \ E" using \F \ E\ S.sets_into_space by (auto simp: M) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) have "?\ (F \ A) \ ?\ F" by (auto intro!: emeasure_mono simp: M N) then have "?\ (F \ A) \ \" using \?\ F \ \\ by (auto simp: top_unique) then have "?\ (F \ (\ - A)) = ?\ F - ?\ (F \ A)" unfolding ** using \F \ A \ sigma_sets \ E\ by (auto intro!: emeasure_Diff simp: M N) also have "\ = ?\ F - ?\ (F \ A)" using eq \F \ E\ compl by simp also have "\ = ?\ (F \ (\ - A))" unfolding ** using \F \ A \ sigma_sets \ E\ \?\ (F \ A) \ \\ by (auto intro!: emeasure_Diff[symmetric] simp: M N) finally show ?case using \space M = \\ by auto next case (union A) then have "?\ (\x. F \ A x) = ?\ (\x. F \ A x)" by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N) with A show ?case by auto qed } note * = this show "M = N" proof (rule measure_eqI) show "sets M = sets N" using M N by simp have [simp, intro]: "\i. A i \ sets M" using A(1) by (auto simp: subset_eq M) fix F assume "F \ sets M" let ?D = "disjointed (\i. F \ A i)" from \space M = \\ have F_eq: "F = (\i. ?D i)" using \F \ sets M\[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq) have [simp, intro]: "\i. ?D i \ sets M" using sets.range_disjointed_sets[of "\i. F \ A i" M] \F \ sets M\ by (auto simp: subset_eq) have "disjoint_family ?D" by (auto simp: disjoint_family_disjointed) moreover have "(\i. emeasure M (?D i)) = (\i. emeasure N (?D i))" proof (intro arg_cong[where f=suminf] ext) fix i have "A i \ ?D i = ?D i" by (auto simp: disjointed_def) then show "emeasure M (?D i) = emeasure N (?D i)" using *[of "A i" "?D i", OF _ A(3)] A(1) by auto qed ultimately show "emeasure M F = emeasure N F" by (simp add: image_subset_iff \sets M = sets N\[symmetric] F_eq[symmetric] suminf_emeasure) qed qed lemma space_empty: "space M = {} \ M = count_space {}" by (rule measure_eqI) (simp_all add: space_empty_iff) lemma measure_eqI_generator_eq_countable: fixes M N :: "'a measure" and E :: "'a set set" and A :: "'a set set" assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" and sets: "sets M = sigma_sets \ E" "sets N = sigma_sets \ E" and A: "A \ E" "(\A) = \" "countable A" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof cases assume "\ = {}" have *: "sigma_sets \ E = sets (sigma \ E)" using E(2) by simp have "space M = \" "space N = \" using sets E(2) unfolding * by (auto dest: sets_eq_imp_space_eq simp del: sets_measure_of) then show "M = N" unfolding \\ = {}\ by (auto dest: space_empty) next assume "\ \ {}" with \\A = \\ have "A \ {}" by auto from this \countable A\ have rng: "range (from_nat_into A) = A" by (rule range_from_nat_into) show "M = N" proof (rule measure_eqI_generator_eq[OF E sets]) show "range (from_nat_into A) \ E" unfolding rng using \A \ E\ . show "(\i. from_nat_into A i) = \" unfolding rng using \\A = \\ . show "emeasure M (from_nat_into A i) \ \" for i using rng by (intro A) auto qed qed lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M" proof (intro measure_eqI emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) (emeasure M)" by (simp add: positive_def) show "countably_additive (sets M) (emeasure M)" by (simp add: emeasure_countably_additive) qed simp_all subsection \\\\-null sets\ definition\<^marker>\tag important\ null_sets :: "'a measure \ 'a set set" where "null_sets M = {N\sets M. emeasure M N = 0}" lemma null_setsD1[dest]: "A \ null_sets M \ emeasure M A = 0" by (simp add: null_sets_def) lemma null_setsD2[dest]: "A \ null_sets M \ A \ sets M" unfolding null_sets_def by simp lemma null_setsI[intro]: "emeasure M A = 0 \ A \ sets M \ A \ null_sets M" unfolding null_sets_def by simp interpretation null_sets: ring_of_sets "space M" "null_sets M" for M proof (rule ring_of_setsI) show "null_sets M \ Pow (space M)" using sets.sets_into_space by auto show "{} \ null_sets M" by auto fix A B assume null_sets: "A \ null_sets M" "B \ null_sets M" then have sets: "A \ sets M" "B \ sets M" by auto then have *: "emeasure M (A \ B) \ emeasure M A + emeasure M B" "emeasure M (A - B) \ emeasure M A" by (auto intro!: emeasure_subadditive emeasure_mono) then have "emeasure M B = 0" "emeasure M A = 0" using null_sets by auto with sets * show "A - B \ null_sets M" "A \ B \ null_sets M" by (auto intro!: antisym zero_le) qed lemma UN_from_nat_into: assumes I: "countable I" "I \ {}" shows "(\i\I. N i) = (\i. N (from_nat_into I i))" proof - have "(\i\I. N i) = \(N ` range (from_nat_into I))" using I by simp also have "\ = (\i. (N \ from_nat_into I) i)" by simp finally show ?thesis by simp qed lemma null_sets_UN': assumes "countable I" assumes "\i. i \ I \ N i \ null_sets M" shows "(\i\I. N i) \ null_sets M" proof cases assume "I = {}" then show ?thesis by simp next assume "I \ {}" show ?thesis proof (intro conjI CollectI null_setsI) show "(\i\I. N i) \ sets M" using assms by (intro sets.countable_UN') auto have "emeasure M (\i\I. N i) \ (\n. emeasure M (N (from_nat_into I n)))" unfolding UN_from_nat_into[OF \countable I\ \I \ {}\] using assms \I \ {}\ by (intro emeasure_subadditive_countably) (auto intro: from_nat_into) also have "(\n. emeasure M (N (from_nat_into I n))) = (\_. 0)" using assms \I \ {}\ by (auto intro: from_nat_into) finally show "emeasure M (\i\I. N i) = 0" by (intro antisym zero_le) simp qed qed lemma null_sets_UN[intro]: "(\i::'i::countable. N i \ null_sets M) \ (\i. N i) \ null_sets M" by (rule null_sets_UN') auto lemma null_set_Int1: assumes "B \ null_sets M" "A \ sets M" shows "A \ B \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (A \ B) = 0" using assms by (intro emeasure_eq_0[of B _ "A \ B"]) auto qed (insert assms, auto) lemma null_set_Int2: assumes "B \ null_sets M" "A \ sets M" shows "B \ A \ null_sets M" using assms by (subst Int_commute) (rule null_set_Int1) lemma emeasure_Diff_null_set: assumes "B \ null_sets M" "A \ sets M" shows "emeasure M (A - B) = emeasure M A" proof - have *: "A - B = (A - (A \ B))" by auto have "A \ B \ null_sets M" using assms by (rule null_set_Int1) then show ?thesis unfolding * using assms by (subst emeasure_Diff) auto qed lemma null_set_Diff: assumes "B \ null_sets M" "A \ sets M" shows "B - A \ null_sets M" proof (intro CollectI conjI null_setsI) show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto qed (insert assms, auto) lemma emeasure_Un_null_set: assumes "A \ sets M" "B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A" proof - have *: "A \ B = A \ (B - A)" by auto have "B - A \ null_sets M" using assms(2,1) by (rule null_set_Diff) then show ?thesis unfolding * using assms by (subst plus_emeasure[symmetric]) auto qed lemma emeasure_Un': assumes "A \ sets M" "B \ sets M" "A \ B \ null_sets M" shows "emeasure M (A \ B) = emeasure M A + emeasure M B" proof - have "A \ B = A \ (B - A \ B)" by blast also have "emeasure M \ = emeasure M A + emeasure M (B - A \ B)" using assms by (subst plus_emeasure) auto also have "emeasure M (B - A \ B) = emeasure M B" using assms by (intro emeasure_Diff_null_set) auto finally show ?thesis . qed subsection \The almost everywhere filter (i.e.\ quantifier)\ definition\<^marker>\tag important\ ae_filter :: "'a measure \ 'a filter" where "ae_filter M = (INF N\null_sets M. principal (space M - N))" abbreviation almost_everywhere :: "'a measure \ ('a \ bool) \ bool" where "almost_everywhere M P \ eventually P (ae_filter M)" syntax "_almost_everywhere" :: "pttrn \ 'a \ bool \ bool" ("AE _ in _. _" [0,0,10] 10) translations "AE x in M. P" \ "CONST almost_everywhere M (\x. P)" abbreviation "set_almost_everywhere A M P \ AE x in M. x \ A \ P x" syntax "_set_almost_everywhere" :: "pttrn \ 'a set \ 'a \ bool \ bool" ("AE _\_ in _./ _" [0,0,0,10] 10) translations "AE x\A in M. P" \ "CONST set_almost_everywhere A M (\x. P)" lemma eventually_ae_filter: "eventually P (ae_filter M) \ (\N\null_sets M. {x \ space M. \ P x} \ N)" unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq) lemma AE_I': "N \ null_sets M \ {x\space M. \ P x} \ N \ (AE x in M. P x)" unfolding eventually_ae_filter by auto lemma AE_iff_null: assumes "{x\space M. \ P x} \ sets M" (is "?P \ sets M") shows "(AE x in M. P x) \ {x\space M. \ P x} \ null_sets M" proof assume "AE x in M. P x" then obtain N where N: "N \ sets M" "?P \ N" "emeasure M N = 0" unfolding eventually_ae_filter by auto have "emeasure M ?P \ emeasure M N" using assms N(1,2) by (auto intro: emeasure_mono) then have "emeasure M ?P = 0" unfolding \emeasure M N = 0\ by auto then show "?P \ null_sets M" using assms by auto next assume "?P \ null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I') qed lemma AE_iff_null_sets: "N \ sets M \ N \ null_sets M \ (AE x in M. x \ N)" using Int_absorb1[OF sets.sets_into_space, of N M] by (subst AE_iff_null) (auto simp: Int_def[symmetric]) lemma AE_not_in: "N \ null_sets M \ AE x in M. x \ N" by (metis AE_iff_null_sets null_setsD2) lemma AE_iff_measurable: "N \ sets M \ {x\space M. \ P x} = N \ (AE x in M. P x) \ emeasure M N = 0" using AE_iff_null[of _ P] by auto lemma AE_E[consumes 1]: assumes "AE x in M. P x" obtains N where "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" using assms unfolding eventually_ae_filter by auto lemma AE_E2: assumes "AE x in M. P x" "{x\space M. P x} \ sets M" shows "emeasure M {x\space M. \ P x} = 0" (is "emeasure M ?P = 0") proof - have "{x\space M. \ P x} = space M - {x\space M. P x}" by auto with AE_iff_null[of M P] assms show ?thesis by auto qed lemma AE_E3: assumes "AE x in M. P x" obtains N where "\x. x \ space M - N \ P x" "N \ null_sets M" using assms unfolding eventually_ae_filter by auto lemma AE_I: assumes "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" shows "AE x in M. P x" using assms unfolding eventually_ae_filter by auto lemma AE_mp[elim!]: assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \ Q x" shows "AE x in M. Q x" proof - from AE_P obtain A where P: "{x\space M. \ P x} \ A" and A: "A \ sets M" "emeasure M A = 0" by (auto elim!: AE_E) from AE_imp obtain B where imp: "{x\space M. P x \ \ Q x} \ B" and B: "B \ sets M" "emeasure M B = 0" by (auto elim!: AE_E) show ?thesis proof (intro AE_I) have "emeasure M (A \ B) \ 0" using emeasure_subadditive[of A M B] A B by auto then show "A \ B \ sets M" "emeasure M (A \ B) = 0" using A B by auto show "{x\space M. \ Q x} \ A \ B" using P imp by auto qed qed text \The next lemma is convenient to combine with a lemma whose conclusion is of the form \AE x in M. P x = Q x\: for such a lemma, there is no \[symmetric]\ variant, but using \AE_symmetric[OF...]\ will replace it.\ (* depricated replace by laws about eventually *) lemma shows AE_iffI: "AE x in M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" and AE_disjI1: "AE x in M. P x \ AE x in M. P x \ Q x" and AE_disjI2: "AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conjI: "AE x in M. P x \ AE x in M. Q x \ AE x in M. P x \ Q x" and AE_conj_iff[simp]: "(AE x in M. P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_symmetric: assumes "AE x in M. P x = Q x" shows "AE x in M. Q x = P x" using assms by auto lemma AE_impI: "(P \ AE x in M. Q x) \ AE x in M. P \ Q x" by fastforce lemma AE_measure: assumes AE: "AE x in M. P x" and sets: "{x\space M. P x} \ sets M" (is "?P \ sets M") shows "emeasure M {x\space M. P x} = emeasure M (space M)" proof - - from AE_E[OF AE] guess N . note N = this + from AE_E[OF AE] obtain N + where N: "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ sets M" + by auto with sets have "emeasure M (space M) \ emeasure M (?P \ N)" by (intro emeasure_mono) auto also have "\ \ emeasure M ?P + emeasure M N" using sets N by (intro emeasure_subadditive) auto also have "\ = emeasure M ?P" using N by simp finally show "emeasure M ?P = emeasure M (space M)" using emeasure_space[of M "?P"] by auto qed lemma AE_space: "AE x in M. x \ space M" by (rule AE_I[where N="{}"]) auto lemma AE_I2[simp, intro]: "(\x. x \ space M \ P x) \ AE x in M. P x" using AE_space by force lemma AE_Ball_mp: "\x\space M. P x \ AE x in M. P x \ Q x \ AE x in M. Q x" by auto lemma AE_cong[cong]: "(\x. x \ space M \ P x \ Q x) \ (AE x in M. P x) \ (AE x in M. Q x)" by auto lemma AE_cong_simp: "M = N \ (\x. x \ space N =simp=> P x = Q x) \ (AE x in M. P x) \ (AE x in N. Q x)" by (auto simp: simp_implies_def) lemma AE_all_countable: "(AE x in M. \i. P i x) \ (\i::'i::countable. AE x in M. P i x)" proof assume "\i. AE x in M. P i x" from this[unfolded eventually_ae_filter Bex_def, THEN choice] obtain N where N: "\i. N i \ null_sets M" "\i. {x\space M. \ P i x} \ N i" by auto have "{x\space M. \ (\i. P i x)} \ (\i. {x\space M. \ P i x})" by auto also have "\ \ (\i. N i)" using N by auto finally have "{x\space M. \ (\i. P i x)} \ (\i. N i)" . moreover from N have "(\i. N i) \ null_sets M" by (intro null_sets_UN) auto ultimately show "AE x in M. \i. P i x" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable: assumes [intro]: "countable X" shows "(AE x in M. \y\X. P x y) \ (\y\X. AE x in M. P x y)" proof assume "\y\X. AE x in M. P x y" from this[unfolded eventually_ae_filter Bex_def, THEN bchoice] obtain N where N: "\y. y \ X \ N y \ null_sets M" "\y. y \ X \ {x\space M. \ P x y} \ N y" by auto have "{x\space M. \ (\y\X. P x y)} \ (\y\X. {x\space M. \ P x y})" by auto also have "\ \ (\y\X. N y)" using N by auto finally have "{x\space M. \ (\y\X. P x y)} \ (\y\X. N y)" . moreover from N have "(\y\X. N y) \ null_sets M" by (intro null_sets_UN') auto ultimately show "AE x in M. \y\X. P x y" unfolding eventually_ae_filter by auto qed auto lemma AE_ball_countable': "(\N. N \ I \ AE x in M. P N x) \ countable I \ AE x in M. \N \ I. P N x" unfolding AE_ball_countable by simp lemma AE_pairwise: "countable F \ pairwise (\A B. AE x in M. R x A B) F \ (AE x in M. pairwise (R x) F)" unfolding pairwise_alt by (simp add: AE_ball_countable) lemma AE_discrete_difference: assumes X: "countable X" assumes null: "\x. x \ X \ emeasure M {x} = 0" assumes sets: "\x. x \ X \ {x} \ sets M" shows "AE x in M. x \ X" proof - have "(\x\X. {x}) \ null_sets M" using assms by (intro null_sets_UN') auto from AE_not_in[OF this] show "AE x in M. x \ X" by auto qed lemma AE_finite_all: assumes f: "finite S" shows "(AE x in M. \i\S. P i x) \ (\i\S. AE x in M. P i x)" using f by induct auto lemma AE_finite_allI: assumes "finite S" shows "(\s. s \ S \ AE x in M. Q s x) \ AE x in M. \s\S. Q s x" using AE_finite_all[OF \finite S\] by auto lemma emeasure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "emeasure M A \ emeasure M B" proof cases assume A: "A \ sets M" from imp obtain N where N: "{x\space M. \ (x \ A \ x \ B)} \ N" "N \ null_sets M" by (auto simp: eventually_ae_filter) have "emeasure M A = emeasure M (A - N)" using N A by (subst emeasure_Diff_null_set) auto also have "emeasure M (A - N) \ emeasure M (B - N)" using N A B sets.sets_into_space by (auto intro!: emeasure_mono) also have "emeasure M (B - N) = emeasure M B" using N B by (subst emeasure_Diff_null_set) auto finally show ?thesis . qed (simp add: emeasure_notin_sets) lemma emeasure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "emeasure M A = emeasure M B" using assms by (safe intro!: antisym emeasure_mono_AE) auto lemma emeasure_Collect_eq_AE: "AE x in M. P x \ Q x \ Measurable.pred M Q \ Measurable.pred M P \ emeasure M {x\space M. P x} = emeasure M {x\space M. Q x}" by (intro emeasure_eq_AE) auto lemma emeasure_eq_0_AE: "AE x in M. \ P x \ emeasure M {x\space M. P x} = 0" using AE_iff_measurable[OF _ refl, of M "\x. \ P x"] by (cases "{x\space M. P x} \ sets M") (simp_all add: emeasure_notin_sets) lemma emeasure_0_AE: assumes "emeasure M (space M) = 0" shows "AE x in M. P x" using eventually_ae_filter assms by blast lemma emeasure_add_AE: assumes [measurable]: "A \ sets M" "B \ sets M" "C \ sets M" assumes 1: "AE x in M. x \ C \ x \ A \ x \ B" assumes 2: "AE x in M. \ (x \ A \ x \ B)" shows "emeasure M C = emeasure M A + emeasure M B" proof - have "emeasure M C = emeasure M (A \ B)" by (rule emeasure_eq_AE) (insert 1, auto) also have "\ = emeasure M A + emeasure M (B - A)" by (subst plus_emeasure) auto also have "emeasure M (B - A) = emeasure M B" by (rule emeasure_eq_AE) (insert 2, auto) finally show ?thesis . qed subsection \\\\-finite Measures\ locale\<^marker>\tag important\ sigma_finite_measure = fixes M :: "'a measure" assumes sigma_finite_countable: "\A::'a set set. countable A \ A \ sets M \ (\A) = space M \ (\a\A. emeasure M a \ \)" lemma (in sigma_finite_measure) sigma_finite: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" proof - obtain A :: "'a set set" where [simp]: "countable A" and A: "A \ sets M" "(\A) = space M" "\a. a \ A \ emeasure M a \ \" using sigma_finite_countable by metis show thesis proof cases assume "A = {}" with \(\A) = space M\ show thesis by (intro that[of "\_. {}"]) auto next assume "A \ {}" show thesis proof show "range (from_nat_into A) \ sets M" using \A \ {}\ A by auto have "(\i. from_nat_into A i) = \A" using range_from_nat_into[OF \A \ {}\ \countable A\] by auto with A show "(\i. from_nat_into A i) = space M" by auto qed (intro A from_nat_into \A \ {}\) qed qed lemma (in sigma_finite_measure) sigma_finite_disjoint: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "disjoint_family A" proof - obtain A :: "nat \ 'a set" where range: "range A \ sets M" and space: "(\i. A i) = space M" and measure: "\i. emeasure M (A i) \ \" using sigma_finite by blast show thesis proof (rule that[of "disjointed A"]) show "range (disjointed A) \ sets M" by (rule sets.range_disjointed_sets[OF range]) show "(\i. disjointed A i) = space M" and "disjoint_family (disjointed A)" using disjoint_family_disjointed UN_disjointed_eq[of A] space range by auto show "emeasure M (disjointed A i) \ \" for i proof - have "emeasure M (disjointed A i) \ emeasure M (A i)" using range disjointed_subset[of A i] by (auto intro!: emeasure_mono) then show ?thesis using measure[of i] by (auto simp: top_unique) qed qed qed lemma (in sigma_finite_measure) sigma_finite_incseq: obtains A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" proof - obtain F :: "nat \ 'a set" where F: "range F \ sets M" "(\i. F i) = space M" "\i. emeasure M (F i) \ \" using sigma_finite by blast show thesis proof (rule that[of "\n. \i\n. F i"]) show "range (\n. \i\n. F i) \ sets M" using F by (force simp: incseq_def) show "(\n. \i\n. F i) = space M" proof - from F have "\x. x \ space M \ \i. x \ F i" by auto with F show ?thesis by fastforce qed show "emeasure M (\i\n. F i) \ \" for n proof - have "emeasure M (\i\n. F i) \ (\i\n. emeasure M (F i))" using F by (auto intro!: emeasure_subadditive_finite) also have "\ < \" using F by (auto simp: sum_Pinfty less_top) finally show ?thesis by simp qed show "incseq (\n. \i\n. F i)" by (force simp: incseq_def) qed qed lemma (in sigma_finite_measure) approx_PInf_emeasure_with_finite: fixes C::real assumes W_meas: "W \ sets M" and W_inf: "emeasure M W = \" obtains Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" proof - obtain A :: "nat \ 'a set" where A: "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" "incseq A" using sigma_finite_incseq by blast define B where "B = (\i. W \ A i)" have B_meas: "\i. B i \ sets M" using W_meas \range A \ sets M\ B_def by blast have b: "\i. B i \ W" using B_def by blast { fix i have "emeasure M (B i) \ emeasure M (A i)" using A by (intro emeasure_mono) (auto simp: B_def) also have "emeasure M (A i) < \" using \\i. emeasure M (A i) \ \\ by (simp add: less_top) finally have "emeasure M (B i) < \" . } note c = this have "W = (\i. B i)" using B_def \(\i. A i) = space M\ W_meas by auto moreover have "incseq B" using B_def \incseq A\ by (simp add: incseq_def subset_eq) ultimately have "(\i. emeasure M (B i)) \ emeasure M W" using W_meas B_meas by (simp add: B_meas Lim_emeasure_incseq image_subset_iff) then have "(\i. emeasure M (B i)) \ \" using W_inf by simp from order_tendstoD(1)[OF this, of C] obtain i where d: "emeasure M (B i) > C" by (auto simp: eventually_sequentially) have "B i \ sets M" "B i \ W" "emeasure M (B i) < \" "emeasure M (B i) > C" using B_meas b c d by auto then show ?thesis using that by blast qed subsection \Measure space induced by distribution of \<^const>\measurable\-functions\ definition\<^marker>\tag important\ distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ space M))" lemma shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N" and space_distr[simp]: "space (distr M N f) = space N" by (auto simp: distr_def) lemma shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'" and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng" by (auto simp: measurable_def) lemma distr_cong: "M = K \ sets N = sets L \ (\x. x \ space M \ f x = g x) \ distr M N f = distr K L g" using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong) lemma emeasure_distr: fixes f :: "'a \ 'b" assumes f: "f \ measurable M N" and A: "A \ sets N" shows "emeasure (distr M N f) A = emeasure M (f -` A \ space M)" (is "_ = ?\ A") unfolding distr_def proof (rule emeasure_measure_of_sigma) show "positive (sets N) ?\" by (auto simp: positive_def) show "countably_additive (sets N) ?\" proof (intro countably_additiveI) fix A :: "nat \ 'b set" assume "range A \ sets N" "disjoint_family A" then have A: "\i. A i \ sets N" "(\i. A i) \ sets N" by auto then have *: "range (\i. f -` (A i) \ space M) \ sets M" using f by (auto simp: measurable_def) moreover have "(\i. f -` A i \ space M) \ sets M" using * by blast moreover have **: "disjoint_family (\i. f -` A i \ space M)" using \disjoint_family A\ by (auto simp: disjoint_family_on_def) ultimately show "(\i. ?\ (A i)) = ?\ (\i. A i)" using suminf_emeasure[OF _ **] A f by (auto simp: comp_def vimage_UN) qed show "sigma_algebra (space N) (sets N)" .. qed fact lemma emeasure_Collect_distr: assumes X[measurable]: "X \ measurable M N" "Measurable.pred N P" shows "emeasure (distr M N X) {x\space N. P x} = emeasure M {x\space M. P (X x)}" by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space]) lemma emeasure_lfp2[consumes 1, case_names cont f measurable]: assumes "P M" assumes cont: "sup_continuous F" assumes f: "\M. P M \ f \ measurable M' M" assumes *: "\M A. P M \ (\N. P N \ Measurable.pred N A) \ Measurable.pred M (F A)" shows "emeasure M' {x\space M'. lfp F (f x)} = (SUP i. emeasure M' {x\space M'. (F ^^ i) (\x. False) (f x)})" proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f]) show "f \ measurable M' M" "f \ measurable M' M" using f[OF \P M\] by auto { fix i show "Measurable.pred M ((F ^^ i) (\x. False))" using \P M\ by (induction i arbitrary: M) (auto intro!: *) } show "Measurable.pred M (lfp F)" using \P M\ cont * by (rule measurable_lfp_coinduct[of P]) have "emeasure (distr M' M f) {x \ space (distr M' M f). lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space (distr M' M f). (F ^^ i) (\x. False) x})" using \P M\ proof (coinduction arbitrary: M rule: emeasure_lfp') case (measurable A N) then have "\N. P N \ Measurable.pred (distr M' N f) A" by metis then have "\N. P N \ Measurable.pred N A" by simp with \P N\[THEN *] show ?case by auto qed fact then show "emeasure (distr M' M f) {x \ space M. lfp F x} = (SUP i. emeasure (distr M' M f) {x \ space M. (F ^^ i) (\x. False) x})" by simp qed lemma distr_id[simp]: "distr N N (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma distr_id2: "sets M = sets N \ distr N M (\x. x) = N" by (rule measure_eqI) (auto simp: emeasure_distr) lemma measure_distr: "f \ measurable M N \ S \ sets N \ measure (distr M N f) S = measure M (f -` S \ space M)" by (simp add: emeasure_distr measure_def) lemma distr_cong_AE: assumes 1: "M = K" "sets N = sets L" and 2: "(AE x in M. f x = g x)" and "f \ measurable M N" and "g \ measurable K L" shows "distr M N f = distr K L g" proof (rule measure_eqI) fix A assume "A \ sets (distr M N f)" with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A" by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets) qed (insert 1, simp) lemma AE_distrD: assumes f: "f \ measurable M M'" and AE: "AE x in distr M M' f. P x" shows "AE x in M. P (f x)" proof - - from AE[THEN AE_E] guess N . + from AE[THEN AE_E] obtain N + where "{x \ space (distr M M' f). \ P x} \ N" + "emeasure (distr M M' f) N = 0" + "N \ sets (distr M M' f)" + by auto with f show ?thesis - unfolding eventually_ae_filter - by (intro bexI[of _ "f -` N \ space M"]) + by (simp add: eventually_ae_filter, intro bexI[of _ "f -` N \ space M"]) (auto simp: emeasure_distr measurable_def) qed lemma AE_distr_iff: assumes f[measurable]: "f \ measurable M N" and P[measurable]: "{x \ space N. P x} \ sets N" shows "(AE x in distr M N f. P x) \ (AE x in M. P (f x))" proof (subst (1 2) AE_iff_measurable[OF _ refl]) have "f -` {x\space N. \ P x} \ space M = {x \ space M. \ P (f x)}" using f[THEN measurable_space] by auto then show "(emeasure (distr M N f) {x \ space (distr M N f). \ P x} = 0) = (emeasure M {x \ space M. \ P (f x)} = 0)" by (simp add: emeasure_distr) qed auto lemma null_sets_distr_iff: "f \ measurable M N \ A \ null_sets (distr M N f) \ f -` A \ space M \ null_sets M \ A \ sets N" by (auto simp add: null_sets_def emeasure_distr) proposition distr_distr: "g \ measurable N L \ f \ measurable M N \ distr (distr M N f) L g = distr M L (g \ f)" by (auto simp add: emeasure_distr measurable_space intro!: arg_cong[where f="emeasure M"] measure_eqI) subsection\<^marker>\tag unimportant\ \Real measure values\ lemma ring_of_finite_sets: "ring_of_sets (space M) {A\sets M. emeasure M A \ top}" proof (rule ring_of_setsI) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a \ b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_subadditive[of a M b] by (auto simp: top_unique) show "a \ {A \ sets M. emeasure M A \ top} \ b \ {A \ sets M. emeasure M A \ top} \ a - b \ {A \ sets M. emeasure M A \ top}" for a b using emeasure_mono[of "a - b" a M] by (auto simp: top_unique) qed (auto dest: sets.sets_into_space) lemma measure_nonneg[simp]: "0 \ measure M A" unfolding measure_def by auto lemma measure_nonneg' [simp]: "\ measure M A < 0" using measure_nonneg not_le by blast lemma zero_less_measure_iff: "0 < measure M A \ measure M A \ 0" using measure_nonneg[of M A] by (auto simp add: le_less) lemma measure_le_0_iff: "measure M X \ 0 \ measure M X = 0" using measure_nonneg[of M X] by linarith lemma measure_empty[simp]: "measure M {} = 0" unfolding measure_def by (simp add: zero_ennreal.rep_eq) lemma emeasure_eq_ennreal_measure: "emeasure M A \ top \ emeasure M A = ennreal (measure M A)" by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def) lemma measure_zero_top: "emeasure M A = top \ measure M A = 0" by (simp add: measure_def) lemma measure_eq_emeasure_eq_ennreal: "0 \ x \ emeasure M A = ennreal x \ measure M A = x" using emeasure_eq_ennreal_measure[of M A] by (cases "A \ M") (auto simp: measure_notin_sets emeasure_notin_sets) lemma enn2real_plus:"a < top \ b < top \ enn2real (a + b) = enn2real a + enn2real b" by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top del: real_of_ereal_enn2ereal) lemma enn2real_sum:"(\i. i \ I \ f i < top) \ enn2real (sum f I) = sum (enn2real \ f) I" by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus) lemma measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: measure_def) lemma measure_Union: "emeasure M A \ \ \ emeasure M B \ \ \ A \ sets M \ B \ sets M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top) lemma disjoint_family_on_insert: "i \ I \ disjoint_family_on A (insert i I) \ A i \ (\i\I. A i) = {} \ disjoint_family_on A I" by (fastforce simp: disjoint_family_on_def) lemma measure_finite_Union: "finite S \ A`S \ sets M \ disjoint_family_on A S \ (\i. i \ S \ emeasure M (A i) \ \) \ measure M (\i\S. A i) = (\i\S. measure M (A i))" by (induction S rule: finite_induct) (auto simp: disjoint_family_on_insert measure_Union sum_emeasure[symmetric] sets.countable_UN'[OF countable_finite]) lemma measure_Diff: assumes finite: "emeasure M A \ \" and measurable: "A \ sets M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" proof - have "emeasure M (A - B) \ emeasure M A" "emeasure M B \ emeasure M A" using measurable by (auto intro!: emeasure_mono) hence "measure M ((A - B) \ B) = measure M (A - B) + measure M B" using measurable finite by (rule_tac measure_Union) (auto simp: top_unique) thus ?thesis using \B \ A\ by (auto simp: Un_absorb2) qed lemma measure_UNION: assumes measurable: "range A \ sets M" "disjoint_family A" assumes finite: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" proof - have "(\i. emeasure M (A i)) sums (emeasure M (\i. A i))" unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums) moreover { fix i have "emeasure M (A i) \ emeasure M (\i. A i)" using measurable by (auto intro!: emeasure_mono) then have "emeasure M (A i) = ennreal ((measure M (A i)))" using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) } ultimately show ?thesis using finite by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all qed lemma measure_subadditive: assumes measurable: "A \ sets M" "B \ sets M" and fin: "emeasure M A \ \" "emeasure M B \ \" shows "measure M (A \ B) \ measure M A + measure M B" proof - have "emeasure M (A \ B) \ \" using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique) then show "(measure M (A \ B)) \ (measure M A) + (measure M B)" using emeasure_subadditive[OF measurable] fin apply simp apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure) apply (auto simp flip: ennreal_plus) done qed lemma measure_subadditive_finite: assumes A: "finite I" "A`I \ sets M" and fin: "\i. i \ I \ emeasure M (A i) \ \" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" proof - { have "emeasure M (\i\I. A i) \ (\i\I. emeasure M (A i))" using emeasure_subadditive_finite[OF A] . also have "\ < \" using fin by (simp add: less_top A) finally have "emeasure M (\i\I. A i) \ top" by simp } note * = this show ?thesis using emeasure_subadditive_finite[OF A] fin unfolding emeasure_eq_ennreal_measure[OF *] by (simp_all add: sum_nonneg emeasure_eq_ennreal_measure) qed lemma measure_subadditive_countably: assumes A: "range A \ sets M" and fin: "(\i. emeasure M (A i)) \ \" shows "measure M (\i. A i) \ (\i. measure M (A i))" proof - from fin have **: "\i. emeasure M (A i) \ top" using ennreal_suminf_lessD[of "\i. emeasure M (A i)"] by (simp add: less_top) { have "emeasure M (\i. A i) \ (\i. emeasure M (A i))" using emeasure_subadditive_countably[OF A] . also have "\ < \" using fin by (simp add: less_top) finally have "emeasure M (\i. A i) \ top" by simp } then have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" by (rule emeasure_eq_ennreal_measure[symmetric]) also have "\ \ (\i. emeasure M (A i))" using emeasure_subadditive_countably[OF A] . also have "\ = ennreal (\i. measure M (A i))" using fin unfolding emeasure_eq_ennreal_measure[OF **] by (subst suminf_ennreal) (auto simp: **) finally show ?thesis apply (rule ennreal_le_iff[THEN iffD1, rotated]) apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top) using fin apply (simp add: emeasure_eq_ennreal_measure[OF **]) done qed lemma measure_Un_null_set: "A \ sets M \ B \ null_sets M \ measure M (A \ B) = measure M A" by (simp add: measure_def emeasure_Un_null_set) lemma measure_Diff_null_set: "A \ sets M \ B \ null_sets M \ measure M (A - B) = measure M A" by (simp add: measure_def emeasure_Diff_null_set) lemma measure_eq_sum_singleton: "finite S \ (\x. x \ S \ {x} \ sets M) \ (\x. x \ S \ emeasure M {x} \ \) \ measure M S = (\x\S. measure M {x})" using emeasure_eq_sum_singleton[of S M] by (intro measure_eq_emeasure_eq_ennreal) (auto simp: sum_nonneg emeasure_eq_ennreal_measure) lemma Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" and fin: "emeasure M (\i. A i) \ \" shows "(\i. measure M (A i)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin by (auto simp: emeasure_eq_ennreal_measure) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using assms emeasure_mono[of "A _" "\i. A i" M] by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans) ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using A by (auto intro!: Lim_emeasure_incseq) qed auto lemma Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" and fin: "\i. emeasure M (A i) \ \" shows "(\n. measure M (A n)) \ measure M (\i. A i)" proof (rule tendsto_ennrealD) have "ennreal (measure M (\i. A i)) = emeasure M (\i. A i)" using fin[of 0] A emeasure_mono[of "\i. A i" "A 0" M] by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans) moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto ultimately show "(\x. ennreal (measure M (A x))) \ ennreal (measure M (\i. A i))" using fin A by (auto intro!: Lim_emeasure_decseq) qed auto subsection \Set of measurable sets with finite measure\ definition\<^marker>\tag important\ fmeasurable :: "'a measure \ 'a set set" where "fmeasurable M = {A\sets M. emeasure M A < \}" lemma fmeasurableD[dest, measurable_dest]: "A \ fmeasurable M \ A \ sets M" by (auto simp: fmeasurable_def) lemma fmeasurableD2: "A \ fmeasurable M \ emeasure M A \ top" by (auto simp: fmeasurable_def) lemma fmeasurableI: "A \ sets M \ emeasure M A < \ \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI_null_sets: "A \ null_sets M \ A \ fmeasurable M" by (auto simp: fmeasurable_def) lemma fmeasurableI2: "A \ fmeasurable M \ B \ A \ B \ sets M \ B \ fmeasurable M" using emeasure_mono[of B A M] by (auto simp: fmeasurable_def) lemma measure_mono_fmeasurable: "A \ B \ A \ sets M \ B \ fmeasurable M \ measure M A \ measure M B" by (auto simp: measure_def fmeasurable_def intro!: emeasure_mono enn2real_mono) lemma emeasure_eq_measure2: "A \ fmeasurable M \ emeasure M A = measure M A" by (simp add: emeasure_eq_ennreal_measure fmeasurable_def less_top) interpretation fmeasurable: ring_of_sets "space M" "fmeasurable M" proof (rule ring_of_setsI) show "fmeasurable M \ Pow (space M)" "{} \ fmeasurable M" by (auto simp: fmeasurable_def dest: sets.sets_into_space) fix a b assume *: "a \ fmeasurable M" "b \ fmeasurable M" then have "emeasure M (a \ b) \ emeasure M a + emeasure M b" by (intro emeasure_subadditive) auto also have "\ < top" using * by (auto simp: fmeasurable_def) finally show "a \ b \ fmeasurable M" using * by (auto intro: fmeasurableI) show "a - b \ fmeasurable M" using emeasure_mono[of "a - b" a M] * by (auto simp: fmeasurable_def) qed subsection\<^marker>\tag unimportant\\Measurable sets formed by unions and intersections\ lemma fmeasurable_Diff: "A \ fmeasurable M \ B \ sets M \ A - B \ fmeasurable M" using fmeasurableI2[of A M "A - B"] by auto lemma fmeasurable_Int_fmeasurable: "\S \ fmeasurable M; T \ sets M\ \ (S \ T) \ fmeasurable M" by (meson fmeasurableD fmeasurableI2 inf_le1 sets.Int) lemma fmeasurable_UN: assumes "countable I" "\i. i \ I \ F i \ A" "\i. i \ I \ F i \ sets M" "A \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "A \ fmeasurable M" "(\i\I. F i) \ A" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_UN') auto qed lemma fmeasurable_INT: assumes "countable I" "i \ I" "\i. i \ I \ F i \ sets M" "F i \ fmeasurable M" shows "(\i\I. F i) \ fmeasurable M" proof (rule fmeasurableI2) show "F i \ fmeasurable M" "(\i\I. F i) \ F i" using assms by auto show "(\i\I. F i) \ sets M" using assms by (intro sets.countable_INT') auto qed lemma measurable_measure_Diff: assumes "A \ fmeasurable M" "B \ sets M" "B \ A" shows "measure M (A - B) = measure M A - measure M B" by (simp add: assms fmeasurableD fmeasurableD2 measure_Diff) lemma measurable_Un_null_set: assumes "B \ null_sets M" shows "(A \ B \ fmeasurable M \ A \ sets M) \ A \ fmeasurable M" using assms by (fastforce simp add: fmeasurable.Un fmeasurableI_null_sets intro: fmeasurableI2) lemma measurable_Diff_null_set: assumes "B \ null_sets M" shows "(A - B) \ fmeasurable M \ A \ sets M \ A \ fmeasurable M" using assms by (metis Un_Diff_cancel2 fmeasurable.Diff fmeasurableD fmeasurableI_null_sets measurable_Un_null_set) lemma fmeasurable_Diff_D: assumes m: "T - S \ fmeasurable M" "S \ fmeasurable M" and sub: "S \ T" shows "T \ fmeasurable M" proof - have "T = S \ (T - S)" using assms by blast then show ?thesis by (metis m fmeasurable.Un) qed lemma measure_Un2: "A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M (B - A)" using measure_Union[of M A "B - A"] by (auto simp: fmeasurableD2 fmeasurable.Diff) lemma measure_Un3: assumes "A \ fmeasurable M" "B \ fmeasurable M" shows "measure M (A \ B) = measure M A + measure M B - measure M (A \ B)" proof - have "measure M (A \ B) = measure M A + measure M (B - A)" using assms by (rule measure_Un2) also have "B - A = B - (A \ B)" by auto also have "measure M (B - (A \ B)) = measure M B - measure M (A \ B)" using assms by (intro measure_Diff) (auto simp: fmeasurable_def) finally show ?thesis by simp qed lemma measure_Un_AE: "AE x in M. x \ A \ x \ B \ A \ fmeasurable M \ B \ fmeasurable M \ measure M (A \ B) = measure M A + measure M B" by (subst measure_Un2) (auto intro!: measure_eq_AE) lemma measure_UNION_AE: assumes I: "finite I" shows "(\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. AE x in M. x \ F i \ x \ F j) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" unfolding AE_pairwise[OF countable_finite, OF I] using I proof (induction I rule: finite_induct) case (insert x I) have "measure M (F x \ \(F ` I)) = measure M (F x) + measure M (\(F ` I))" by (rule measure_Un_AE) (use insert in \auto simp: pairwise_insert\) with insert show ?case by (simp add: pairwise_insert ) qed simp lemma measure_UNION': "finite I \ (\i. i \ I \ F i \ fmeasurable M) \ pairwise (\i j. disjnt (F i) (F j)) I \ measure M (\i\I. F i) = (\i\I. measure M (F i))" by (intro measure_UNION_AE) (auto simp: disjnt_def elim!: pairwise_mono intro!: always_eventually) lemma measure_Union_AE: "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise (\S T. AE x in M. x \ S \ x \ T) F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION_AE[of F "\x. x" M] by simp lemma measure_Union': "finite F \ (\S. S \ F \ S \ fmeasurable M) \ pairwise disjnt F \ measure M (\F) = (\S\F. measure M S)" using measure_UNION'[of F "\x. x" M] by simp lemma measure_Un_le: assumes "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" proof cases assume "A \ fmeasurable M \ B \ fmeasurable M" with measure_subadditive[of A M B] assms show ?thesis by (auto simp: fmeasurableD2) next assume "\ (A \ fmeasurable M \ B \ fmeasurable M)" then have "A \ B \ fmeasurable M" using fmeasurableI2[of "A \ B" M A] fmeasurableI2[of "A \ B" M B] assms by auto with assms show ?thesis by (auto simp: fmeasurable_def measure_def less_top[symmetric]) qed lemma measure_UNION_le: "finite I \ (\i. i \ I \ F i \ sets M) \ measure M (\i\I. F i) \ (\i\I. measure M (F i))" proof (induction I rule: finite_induct) case (insert i I) then have "measure M (\i\insert i I. F i) = measure M (F i \ \ (F ` I))" by simp also from insert have "measure M (F i \ \ (F ` I)) \ measure M (F i) + measure M (\ (F ` I))" by (intro measure_Un_le sets.finite_Union) auto also have "measure M (\i\I. F i) \ (\i\I. measure M (F i))" using insert by auto finally show ?case using insert by simp qed simp lemma measure_Union_le: "finite F \ (\S. S \ F \ S \ sets M) \ measure M (\F) \ (\S\F. measure M S)" using measure_UNION_le[of F "\x. x" M] by simp text\Version for indexed union over a countable set\ lemma assumes "countable I" and I: "\i. i \ I \ A i \ fmeasurable M" and bound: "\I'. I' \ I \ finite I' \ measure M (\i\I'. A i) \ B" shows fmeasurable_UN_bound: "(\i\I. A i) \ fmeasurable M" (is ?fm) and measure_UN_bound: "measure M (\i\I. A i) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof cases assume "I = {}" with \B \ 0\ show ?thesis by simp next assume "I \ {}" have "(\i\I. A i) = (\i. (\n\i. A (from_nat_into I n)))" by (subst range_from_nat_into[symmetric, OF \I \ {}\ \countable I\]) auto then have "emeasure M (\i\I. A i) = emeasure M (\i. (\n\i. A (from_nat_into I n)))" by simp also have "\ = (SUP i. emeasure M (\n\i. A (from_nat_into I n)))" using I \I \ {}\[THEN from_nat_into] by (intro SUP_emeasure_incseq[symmetric]) (fastforce simp: incseq_Suc_iff)+ also have "\ \ B" proof (intro SUP_least) fix i :: nat have "emeasure M (\n\i. A (from_nat_into I n)) = measure M (\n\i. A (from_nat_into I n))" using I \I \ {}\[THEN from_nat_into] by (intro emeasure_eq_measure2 fmeasurable.finite_UN) auto also have "\ = measure M (\n\from_nat_into I ` {..i}. A n)" by simp also have "\ \ B" by (intro ennreal_leI bound) (auto intro: from_nat_into[OF \I \ {}\]) finally show "emeasure M (\n\i. A (from_nat_into I n)) \ ennreal B" . qed finally have *: "emeasure M (\i\I. A i) \ B" . then have ?fm using I \countable I\ by (intro fmeasurableI conjI) (auto simp: less_top[symmetric] top_unique) with * \0\B\ show ?thesis by (simp add: emeasure_eq_measure2) qed then show ?fm ?m by auto qed text\Version for big union of a countable set\ lemma assumes "countable \" and meas: "\D. D \ \ \ D \ fmeasurable M" and bound: "\\. \\ \ \; finite \\ \ measure M (\\) \ B" shows fmeasurable_Union_bound: "\\ \ fmeasurable M" (is ?fm) and measure_Union_bound: "measure M (\\) \ B" (is ?m) proof - have "B \ 0" using bound by force have "?fm \ ?m" proof (cases "\ = {}") case True with \B \ 0\ show ?thesis by auto next case False then obtain D :: "nat \ 'a set" where D: "\ = range D" using \countable \\ uncountable_def by force have 1: "\i. D i \ fmeasurable M" by (simp add: D meas) have 2: "\I'. finite I' \ measure M (\x\I'. D x) \ B" by (simp add: D bound image_subset_iff) show ?thesis unfolding D by (intro conjI fmeasurable_UN_bound [OF _ 1 2] measure_UN_bound [OF _ 1 2]) auto qed then show ?fm ?m by auto qed text\Version for indexed union over the type of naturals\ lemma fixes S :: "nat \ 'a set" assumes S: "\i. S i \ fmeasurable M" and B: "\n. measure M (\i\n. S i) \ B" shows fmeasurable_countable_Union: "(\i. S i) \ fmeasurable M" and measure_countable_Union_le: "measure M (\i. S i) \ B" proof - have mB: "measure M (\i\I. S i) \ B" if "finite I" for I proof - have "(\i\I. S i) \ (\i\Max I. S i)" using Max_ge that by force then have "measure M (\i\I. S i) \ measure M (\i \ Max I. S i)" by (rule measure_mono_fmeasurable) (use S in \blast+\) then show ?thesis using B order_trans by blast qed show "(\i. S i) \ fmeasurable M" by (auto intro: fmeasurable_UN_bound [OF _ S mB]) show "measure M (\n. S n) \ B" by (auto intro: measure_UN_bound [OF _ S mB]) qed lemma measure_diff_le_measure_setdiff: assumes "S \ fmeasurable M" "T \ fmeasurable M" shows "measure M S - measure M T \ measure M (S - T)" proof - have "measure M S \ measure M ((S - T) \ T)" by (simp add: assms fmeasurable.Un fmeasurableD measure_mono_fmeasurable) also have "\ \ measure M (S - T) + measure M T" using assms by (blast intro: measure_Un_le) finally show ?thesis by (simp add: algebra_simps) qed lemma suminf_exist_split2: fixes f :: "nat \ 'a::real_normed_vector" assumes "summable f" shows "(\n. (\k. f(k+n))) \ 0" by (subst lim_sequentially, auto simp add: dist_norm suminf_exist_split[OF _ assms]) lemma emeasure_union_summable: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "emeasure M (\n. A n) < \" "emeasure M (\n. A n) \ (\n. measure M (A n))" proof - define B where "B = (\N. (\n\{.. sets M" for N unfolding B_def by auto have "(\N. emeasure M (B N)) \ emeasure M (\N. B N)" apply (rule Lim_emeasure_incseq) unfolding B_def by (auto simp add: SUP_subset_mono incseq_def) moreover have "emeasure M (B N) \ ennreal (\n. measure M (A n))" for N proof - have *: "(\n\{.. (\n. measure M (A n))" using assms(3) measure_nonneg sum_le_suminf by blast have "emeasure M (B N) \ (\n\{..n\{..n\{.. ennreal (\n. measure M (A n))" using * by (auto simp: ennreal_leI) finally show ?thesis by simp qed ultimately have "emeasure M (\N. B N) \ ennreal (\n. measure M (A n))" by (simp add: Lim_bounded) then show "emeasure M (\n. A n) \ (\n. measure M (A n))" unfolding B_def by (metis UN_UN_flatten UN_lessThan_UNIV) then show "emeasure M (\n. A n) < \" by (auto simp: less_top[symmetric] top_unique) qed lemma borel_cantelli_limsup1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "limsup A \ null_sets M" proof - have "emeasure M (limsup A) \ 0" proof (rule LIMSEQ_le_const) have "(\n. (\k. measure M (A (k+n)))) \ 0" by (rule suminf_exist_split2[OF assms(3)]) then show "(\n. ennreal (\k. measure M (A (k+n)))) \ 0" unfolding ennreal_0[symmetric] by (intro tendsto_ennrealI) have "emeasure M (limsup A) \ (\k. measure M (A (k+n)))" for n proof - have I: "(\k\{n..}. A k) = (\k. A (k+n))" by (auto, metis le_add_diff_inverse2, fastforce) have "emeasure M (limsup A) \ emeasure M (\k\{n..}. A k)" by (rule emeasure_mono, auto simp add: limsup_INF_SUP) also have "... = emeasure M (\k. A (k+n))" using I by auto also have "... \ (\k. measure M (A (k+n)))" apply (rule emeasure_union_summable) using assms summable_ignore_initial_segment[OF assms(3), of n] by auto finally show ?thesis by simp qed then show "\N. \n\N. emeasure M (limsup A) \ (\k. measure M (A (k+n)))" by auto qed then show ?thesis using assms(1) measurable_limsup by auto qed lemma borel_cantelli_AE1: assumes [measurable]: "\n. A n \ sets M" and "\n. emeasure M (A n) < \" "summable (\n. measure M (A n))" shows "AE x in M. eventually (\n. x \ space M - A n) sequentially" proof - have "AE x in M. x \ limsup A" using borel_cantelli_limsup1[OF assms] unfolding eventually_ae_filter by auto moreover { fix x assume "x \ limsup A" then obtain N where "x \ (\n\{N..}. A n)" unfolding limsup_INF_SUP by blast then have "eventually (\n. x \ A n) sequentially" using eventually_sequentially by auto } ultimately show ?thesis by auto qed subsection \Measure spaces with \<^term>\emeasure M (space M) < \\\ locale\<^marker>\tag important\ finite_measure = sigma_finite_measure M for M + assumes finite_emeasure_space: "emeasure M (space M) \ top" lemma finite_measureI[Pure.intro!]: "emeasure M (space M) \ \ \ finite_measure M" proof qed (auto intro!: exI[of _ "{space M}"]) lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \ top" using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique) lemma (in finite_measure) fmeasurable_eq_sets: "fmeasurable M = sets M" by (auto simp: fmeasurable_def less_top[symmetric]) lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)" by (intro emeasure_eq_ennreal_measure) simp lemma (in finite_measure) emeasure_real: "\r. 0 \ r \ emeasure M A = ennreal r" using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto lemma (in finite_measure) bounded_measure: "measure M A \ measure M (space M)" using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_Diff: assumes sets: "A \ sets M" "B \ sets M" and "B \ A" shows "measure M (A - B) = measure M A - measure M B" using measure_Diff[OF _ assms] by simp lemma (in finite_measure) finite_measure_Union: assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" shows "measure M (A \ B) = measure M A + measure M B" using measure_Union[OF _ _ assms] by simp lemma (in finite_measure) finite_measure_finite_Union: assumes measurable: "finite S" "A`S \ sets M" "disjoint_family_on A S" shows "measure M (\i\S. A i) = (\i\S. measure M (A i))" using measure_finite_Union[OF assms] by simp lemma (in finite_measure) finite_measure_UNION: assumes A: "range A \ sets M" "disjoint_family A" shows "(\i. measure M (A i)) sums (measure M (\i. A i))" using measure_UNION[OF A] by simp lemma (in finite_measure) finite_measure_mono: assumes "A \ B" "B \ sets M" shows "measure M A \ measure M B" using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def) lemma (in finite_measure) finite_measure_subadditive: assumes m: "A \ sets M" "B \ sets M" shows "measure M (A \ B) \ measure M A + measure M B" using measure_subadditive[OF m] by simp lemma (in finite_measure) finite_measure_subadditive_finite: assumes "finite I" "A`I \ sets M" shows "measure M (\i\I. A i) \ (\i\I. measure M (A i))" using measure_subadditive_finite[OF assms] by simp lemma (in finite_measure) finite_measure_subadditive_countably: "range A \ sets M \ summable (\i. measure M (A i)) \ measure M (\i. A i) \ (\i. measure M (A i))" by (rule measure_subadditive_countably) (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_sum_singleton: assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" shows "measure M S = (\x\S. measure M {x})" using measure_eq_sum_singleton[OF assms] by simp lemma (in finite_measure) finite_Lim_measure_incseq: assumes A: "range A \ sets M" "incseq A" shows "(\i. measure M (A i)) \ measure M (\i. A i)" using Lim_measure_incseq[OF A] by simp lemma (in finite_measure) finite_Lim_measure_decseq: assumes A: "range A \ sets M" "decseq A" shows "(\n. measure M (A n)) \ measure M (\i. A i)" using Lim_measure_decseq[OF A] by simp lemma (in finite_measure) finite_measure_compl: assumes S: "S \ sets M" shows "measure M (space M - S) = measure M (space M) - measure M S" using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp lemma (in finite_measure) finite_measure_mono_AE: assumes imp: "AE x in M. x \ A \ x \ B" and B: "B \ sets M" shows "measure M A \ measure M B" using assms emeasure_mono_AE[OF imp B] by (simp add: emeasure_eq_measure) lemma (in finite_measure) finite_measure_eq_AE: assumes iff: "AE x in M. x \ A \ x \ B" assumes A: "A \ sets M" and B: "B \ sets M" shows "measure M A = measure M B" using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure) lemma (in finite_measure) measure_increasing: "increasing M (measure M)" by (auto intro!: finite_measure_mono simp: increasing_def) lemma (in finite_measure) measure_zero_union: assumes "s \ sets M" "t \ sets M" "measure M t = 0" shows "measure M (s \ t) = measure M s" using assms proof - have "measure M (s \ t) \ measure M s" using finite_measure_subadditive[of s t] assms by auto moreover have "measure M (s \ t) \ measure M s" using assms by (blast intro: finite_measure_mono) ultimately show ?thesis by simp qed lemma (in finite_measure) measure_eq_compl: assumes "s \ sets M" "t \ sets M" assumes "measure M (space M - s) = measure M (space M - t)" shows "measure M s = measure M t" using assms finite_measure_compl by auto lemma (in finite_measure) measure_eq_bigunion_image: assumes "range f \ sets M" "range g \ sets M" assumes "disjoint_family f" "disjoint_family g" assumes "\ n :: nat. measure M (f n) = measure M (g n)" shows "measure M (\i. f i) = measure M (\i. g i)" using assms proof - have a: "(\ i. measure M (f i)) sums (measure M (\i. f i))" by (rule finite_measure_UNION[OF assms(1,3)]) have b: "(\ i. measure M (g i)) sums (measure M (\i. g i))" by (rule finite_measure_UNION[OF assms(2,4)]) show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp qed lemma (in finite_measure) measure_countably_zero: assumes "range c \ sets M" assumes "\ i. measure M (c i) = 0" shows "measure M (\i :: nat. c i) = 0" proof (rule antisym) show "measure M (\i :: nat. c i) \ 0" using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2)) qed simp lemma (in finite_measure) measure_space_inter: assumes events:"s \ sets M" "t \ sets M" assumes "measure M t = measure M (space M)" shows "measure M (s \ t) = measure M s" proof - have "measure M ((space M - s) \ (space M - t)) = measure M (space M - s)" using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union) also have "(space M - s) \ (space M - t) = space M - (s \ t)" by blast finally show "measure M (s \ t) = measure M s" using events by (auto intro!: measure_eq_compl[of "s \ t" s]) qed lemma (in finite_measure) measure_equiprobable_finite_unions: assumes s: "finite s" "\x. x \ s \ {x} \ sets M" assumes "\ x y. \x \ s; y \ s\ \ measure M {x} = measure M {y}" shows "measure M s = real (card s) * measure M {SOME x. x \ s}" proof cases assume "s \ {}" then have "\ x. x \ s" by blast from someI_ex[OF this] assms have prob_some: "\ x. x \ s \ measure M {x} = measure M {SOME y. y \ s}" by blast have "measure M s = (\ x \ s. measure M {x})" using finite_measure_eq_sum_singleton[OF s] by simp also have "\ = (\ x \ s. measure M {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * measure M {(SOME x. x \ s)}" using sum_constant assms by simp finally show ?thesis by simp qed simp lemma (in finite_measure) measure_real_sum_image_fn: assumes "e \ sets M" assumes "\ x. x \ s \ e \ f x \ sets M" assumes "finite s" assumes disjoint: "\ x y. \x \ s ; y \ s ; x \ y\ \ f x \ f y = {}" assumes upper: "space M \ (\i \ s. f i)" shows "measure M e = (\ x \ s. measure M (e \ f x))" proof - have "e \ (\i\s. f i)" using \e \ sets M\ sets.sets_into_space upper by blast then have e: "e = (\i \ s. e \ f i)" by auto hence "measure M e = measure M (\i \ s. e \ f i)" by simp also have "\ = (\ x \ s. measure M (e \ f x))" proof (rule finite_measure_finite_Union) show "finite s" by fact show "(\i. e \ f i)`s \ sets M" using assms(2) by auto show "disjoint_family_on (\i. e \ f i) s" using disjoint by (auto simp: disjoint_family_on_def) qed finally show ?thesis . qed lemma (in finite_measure) measure_exclude: assumes "A \ sets M" "B \ sets M" assumes "measure M A = measure M (space M)" "A \ B = {}" shows "measure M B = 0" using measure_space_inter[of B A] assms by (auto simp: ac_simps) lemma (in finite_measure) finite_measure_distr: assumes f: "f \ measurable M M'" shows "finite_measure (distr M M' f)" proof (rule finite_measureI) have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) with f show "emeasure (distr M M' f) (space (distr M M' f)) \ \" by (auto simp: emeasure_distr) qed lemma emeasure_gfp[consumes 1, case_names cont measurable]: assumes sets[simp]: "\s. sets (M s) = sets N" assumes "\s. finite_measure (M s)" assumes cont: "inf_continuous F" "inf_continuous f" assumes meas: "\P. Measurable.pred N P \ Measurable.pred N (F P)" assumes iter: "\P s. Measurable.pred N P \ emeasure (M s) {x\space N. F P x} = f (\s. emeasure (M s) {x\space N. P x}) s" assumes bound: "\P. f P \ f (\s. emeasure (M s) (space (M s)))" shows "emeasure (M s) {x\space N. gfp F x} = gfp f s" proof (subst gfp_transfer_bounded[where \="\F s. emeasure (M s) {x\space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric]) interpret finite_measure "M s" for s by fact fix C assume "decseq C" "\i. Measurable.pred N (C i)" then show "(\s. emeasure (M s) {x \ space N. (INF i. C i) x}) = (INF i. (\s. emeasure (M s) {x \ space N. C i x}))" unfolding INF_apply[abs_def] by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure]) next show "f x \ (\s. emeasure (M s) {x \ space N. F top x})" for x using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter) qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont) subsection\<^marker>\tag unimportant\ \Counting space\ lemma strict_monoI_Suc: assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" unfolding strict_mono_def proof safe fix n m :: nat assume "n < m" then show "f n < f m" by (induct m) (auto simp: less_Suc_eq intro: less_trans ord) qed lemma emeasure_count_space: assumes "X \ A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \)" (is "_ = ?M X") unfolding count_space_def proof (rule emeasure_measure_of_sigma) show "X \ Pow A" using \X \ A\ by auto show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow) show positive: "positive (Pow A) ?M" by (auto simp: positive_def) have additive: "additive (Pow A) ?M" by (auto simp: card_Un_disjoint additive_def) interpret ring_of_sets A "Pow A" by (rule ring_of_setsI) auto show "countably_additive (Pow A) ?M" unfolding countably_additive_iff_continuous_from_below[OF positive additive] proof safe fix F :: "nat \ 'a set" assume "incseq F" show "(\i. ?M (F i)) \ ?M (\i. F i)" proof cases assume "\i. \j\i. F i = F j" - then guess i .. note i = this - { fix j from i \incseq F\ have "F j \ F i" - by (cases "i \ j") (auto simp: incseq_def) } + then obtain i where i: "\j\i. F i = F j" .. + with \incseq F\ have "F j \ F i" for j + by (cases "i \ j") (auto simp: incseq_def) then have eq: "(\i. F i) = F i" by auto with i show ?thesis by (auto intro!: Lim_transform_eventually[OF tendsto_const] eventually_sequentiallyI[where c=i]) next assume "\ (\i. \j\i. F i = F j)" then obtain f where f: "\i. i \ f i" "\i. F i \ F (f i)" by metis then have "\i. F i \ F (f i)" using \incseq F\ by (auto simp: incseq_def) with f have *: "\i. F i \ F (f i)" by auto have "incseq (\i. ?M (F i))" using \incseq F\ unfolding incseq_def by (auto simp: card_mono dest: finite_subset) then have "(\i. ?M (F i)) \ (SUP n. ?M (F n))" by (rule LIMSEQ_SUP) moreover have "(SUP n. ?M (F n)) = top" proof (rule ennreal_SUP_eq_top) fix n :: nat show "\k::nat\UNIV. of_nat n \ ?M (F k)" proof (induct n) case (Suc n) - then guess k .. note k = this + then obtain k where "of_nat n \ ?M (F k)" .. moreover have "finite (F k) \ finite (F (f k)) \ card (F k) < card (F (f k))" using \F k \ F (f k)\ by (simp add: psubset_card_mono) moreover have "finite (F (f k)) \ finite (F k)" using \k \ f k\ \incseq F\ by (auto simp: incseq_def dest: finite_subset) ultimately show ?case by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc) qed auto qed moreover have "inj (\n. F ((f ^^ n) 0))" by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *) then have 1: "infinite (range (\i. F ((f ^^ i) 0)))" by (rule range_inj_infinite) have "infinite (Pow (\i. F i))" by (rule infinite_super[OF _ 1]) auto then have "infinite (\i. F i)" by auto ultimately show ?thesis by (simp only:) simp qed qed qed lemma distr_bij_count_space: assumes f: "bij_betw f A B" shows "distr (count_space A) (count_space B) f = count_space B" proof (rule measure_eqI) have f': "f \ measurable (count_space A) (count_space B)" using f unfolding Pi_def bij_betw_def by auto fix X assume "X \ sets (distr (count_space A) (count_space B) f)" then have X: "X \ sets (count_space B)" by auto moreover from X have "f -` X \ A = the_inv_into A f ` X" using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric]) moreover have "inj_on (the_inv_into A f) B" using X f by (auto simp: bij_betw_def inj_on_the_inv_into) with X have "inj_on (the_inv_into A f) X" by (auto intro: subset_inj_on) ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X" using f unfolding emeasure_distr[OF f' X] by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD) qed simp lemma emeasure_count_space_finite[simp]: "X \ A \ finite X \ emeasure (count_space A) X = of_nat (card X)" using emeasure_count_space[of X A] by simp lemma emeasure_count_space_infinite[simp]: "X \ A \ infinite X \ emeasure (count_space A) X = \" using emeasure_count_space[of X A] by simp lemma measure_count_space: "measure (count_space A) X = (if X \ A then of_nat (card X) else 0)" by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat measure_zero_top measure_eq_emeasure_eq_ennreal) lemma emeasure_count_space_eq_0: "emeasure (count_space A) X = 0 \ (X \ A \ X = {})" proof cases assume X: "X \ A" then show ?thesis proof (intro iffI impI) assume "emeasure (count_space A) X = 0" with X show "X = {}" by (subst (asm) emeasure_count_space) (auto split: if_split_asm) qed simp qed (simp add: emeasure_notin_sets) lemma null_sets_count_space: "null_sets (count_space A) = { {} }" unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0) lemma AE_count_space: "(AE x in count_space A. P x) \ (\x\A. P x)" unfolding eventually_ae_filter by (auto simp add: null_sets_count_space) lemma sigma_finite_measure_count_space_countable: assumes A: "countable A" shows "sigma_finite_measure (count_space A)" proof qed (insert A, auto intro!: exI[of _ "(\a. {a}) ` A"]) lemma sigma_finite_measure_count_space: fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)" by (rule sigma_finite_measure_count_space_countable) auto lemma finite_measure_count_space: assumes [simp]: "finite A" shows "finite_measure (count_space A)" by rule simp lemma sigma_finite_measure_count_space_finite: assumes A: "finite A" shows "sigma_finite_measure (count_space A)" proof - interpret finite_measure "count_space A" using A by (rule finite_measure_count_space) show "sigma_finite_measure (count_space A)" .. qed subsection\<^marker>\tag unimportant\ \Measure restricted to space\ lemma emeasure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "emeasure (restrict_space M \) A = emeasure M A" proof (cases "A \ sets M") case True show ?thesis proof (rule emeasure_measure_of[OF restrict_space_def]) show "(\) \ ` sets M \ Pow (\ \ space M)" "A \ sets (restrict_space M \)" using \A \ \\ \A \ sets M\ sets.space_closed by (auto simp: sets_restrict_space) show "positive (sets (restrict_space M \)) (emeasure M)" by (auto simp: positive_def) show "countably_additive (sets (restrict_space M \)) (emeasure M)" proof (rule countably_additiveI) fix A :: "nat \ _" assume "range A \ sets (restrict_space M \)" "disjoint_family A" with assms have "\i. A i \ sets M" "\i. A i \ space M" "disjoint_family A" by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff dest: sets.sets_into_space)+ then show "(\i. emeasure M (A i)) = emeasure M (\i. A i)" by (subst suminf_emeasure) (auto simp: disjoint_family_subset) qed qed next case False with assms have "A \ sets (restrict_space M \)" by (simp add: sets_restrict_space_iff) with False show ?thesis by (simp add: emeasure_notin_sets) qed lemma measure_restrict_space: assumes "\ \ space M \ sets M" "A \ \" shows "measure (restrict_space M \) A = measure M A" using emeasure_restrict_space[OF assms] by (simp add: measure_def) lemma AE_restrict_space_iff: assumes "\ \ space M \ sets M" shows "(AE x in restrict_space M \. P x) \ (AE x in M. x \ \ \ P x)" proof - have ex_cong: "\P Q f. (\x. P x \ Q x) \ (\x. Q x \ P (f x)) \ (\x. P x) \ (\x. Q x)" by auto { fix X assume X: "X \ sets M" "emeasure M X = 0" then have "emeasure M (\ \ space M \ X) \ emeasure M X" by (intro emeasure_mono) auto then have "emeasure M (\ \ space M \ X) = 0" using X by (auto intro!: antisym) } with assms show ?thesis unfolding eventually_ae_filter by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff emeasure_restrict_space cong: conj_cong intro!: ex_cong[where f="\X. (\ \ space M) \ X"]) qed lemma restrict_restrict_space: assumes "A \ space M \ sets M" "B \ space M \ sets M" shows "restrict_space (restrict_space M A) B = restrict_space M (A \ B)" (is "?l = ?r") proof (rule measure_eqI[symmetric]) show "sets ?r = sets ?l" unfolding sets_restrict_space image_comp by (intro image_cong) auto next fix X assume "X \ sets (restrict_space M (A \ B))" then obtain Y where "Y \ sets M" "X = Y \ A \ B" by (auto simp: sets_restrict_space) with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X" by (subst (1 2) emeasure_restrict_space) (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps) qed lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \ B)" proof (rule measure_eqI) show "sets (restrict_space (count_space B) A) = sets (count_space (A \ B))" by (subst sets_restrict_space) auto moreover fix X assume "X \ sets (restrict_space (count_space B) A)" ultimately have "X \ A \ B" by auto then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \ B)) X" by (cases "finite X") (auto simp add: emeasure_restrict_space) qed lemma sigma_finite_measure_restrict_space: assumes "sigma_finite_measure M" and A: "A \ sets M" shows "sigma_finite_measure (restrict_space M A)" proof - interpret sigma_finite_measure M by fact from sigma_finite_countable obtain C where C: "countable C" "C \ sets M" "(\C) = space M" "\a\C. emeasure M a \ \" by blast let ?C = "(\) A ` C" from C have "countable ?C" "?C \ sets (restrict_space M A)" "(\?C) = space (restrict_space M A)" by(auto simp add: sets_restrict_space space_restrict_space) moreover { fix a assume "a \ ?C" then obtain a' where "a = A \ a'" "a' \ C" .. then have "emeasure (restrict_space M A) a \ emeasure M a'" using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono) also have "\ < \" using C(4)[rule_format, of a'] \a' \ C\ by (simp add: less_top) finally have "emeasure (restrict_space M A) a \ \" by simp } ultimately show ?thesis by unfold_locales (rule exI conjI|assumption|blast)+ qed lemma finite_measure_restrict_space: assumes "finite_measure M" and A: "A \ sets M" shows "finite_measure (restrict_space M A)" proof - interpret finite_measure M by fact show ?thesis by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space) qed lemma restrict_distr: assumes [measurable]: "f \ measurable M N" assumes [simp]: "\ \ space N \ sets N" and restrict: "f \ space M \ \" shows "restrict_space (distr M N f) \ = distr M (restrict_space N \) f" (is "?l = ?r") proof (rule measure_eqI) fix A assume "A \ sets (restrict_space (distr M N f) \)" with restrict show "emeasure ?l A = emeasure ?r A" by (subst emeasure_distr) (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr intro!: measurable_restrict_space2) qed (simp add: sets_restrict_space) lemma measure_eqI_restrict_generator: assumes E: "Int_stable E" "E \ Pow \" "\X. X \ E \ emeasure M X = emeasure N X" assumes sets_eq: "sets M = sets N" and \: "\ \ sets M" assumes "sets (restrict_space M \) = sigma_sets \ E" assumes "sets (restrict_space N \) = sigma_sets \ E" assumes ae: "AE x in M. x \ \" "AE x in N. x \ \" assumes A: "countable A" "A \ {}" "A \ E" "\A = \" "\a. a \ A \ emeasure M a \ \" shows "M = N" proof (rule measure_eqI) fix X assume X: "X \ sets M" then have "emeasure M X = emeasure (restrict_space M \) (X \ \)" using ae \ by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE) also have "restrict_space M \ = restrict_space N \" proof (rule measure_eqI_generator_eq) fix X assume "X \ E" then show "emeasure (restrict_space M \) X = emeasure (restrict_space N \) X" using E \ by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq]) next show "range (from_nat_into A) \ E" "(\i. from_nat_into A i) = \" using A by (auto cong del: SUP_cong_simp) next fix i have "emeasure (restrict_space M \) (from_nat_into A i) = emeasure M (from_nat_into A i)" using A \ by (subst emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into) with A show "emeasure (restrict_space M \) (from_nat_into A i) \ \" by (auto intro: from_nat_into) qed fact+ also have "emeasure (restrict_space N \) (X \ \) = emeasure N X" using X ae \ by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE) finally show "emeasure M X = emeasure N X" . qed fact subsection\<^marker>\tag unimportant\ \Null measure\ definition null_measure :: "'a measure \ 'a measure" where "null_measure M = sigma (space M) (sets M)" lemma space_null_measure[simp]: "space (null_measure M) = space M" by (simp add: null_measure_def) lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M" by (simp add: null_measure_def) lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0" by (cases "X \ sets M", rule emeasure_measure_of) (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def dest: sets.sets_into_space) lemma measure_null_measure[simp]: "measure (null_measure M) X = 0" by (intro measure_eq_emeasure_eq_ennreal) auto lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M" by(rule measure_eqI) simp_all subsection \Scaling a measure\ definition\<^marker>\tag important\ scale_measure :: "ennreal \ 'a measure \ 'a measure" where "scale_measure r M = measure_of (space M) (sets M) (\A. r * emeasure M A)" lemma space_scale_measure: "space (scale_measure r M) = space M" by (simp add: scale_measure_def) lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M" by (simp add: scale_measure_def) lemma emeasure_scale_measure [simp]: "emeasure (scale_measure r M) A = r * emeasure M A" (is "_ = ?\ A") proof(cases "A \ sets M") case True show ?thesis unfolding scale_measure_def proof(rule emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) ?\" by (simp add: positive_def) show "countably_additive (sets M) ?\" proof (rule countably_additiveI) fix A :: "nat \ _" assume *: "range A \ sets M" "disjoint_family A" have "(\i. ?\ (A i)) = r * (\i. emeasure M (A i))" by simp also have "\ = ?\ (\i. A i)" using * by(simp add: suminf_emeasure) finally show "(\i. ?\ (A i)) = ?\ (\i. A i)" . qed qed(fact True) qed(simp add: emeasure_notin_sets) lemma scale_measure_1 [simp]: "scale_measure 1 M = M" by(rule measure_eqI) simp_all lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M" by(rule measure_eqI) simp_all lemma measure_scale_measure [simp]: "0 \ r \ measure (scale_measure r M) A = r * measure M A" using emeasure_scale_measure[of r M A] emeasure_eq_ennreal_measure[of M A] measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A] by (cases "emeasure (scale_measure r M) A = top") (auto simp del: emeasure_scale_measure simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric]) lemma scale_scale_measure [simp]: "scale_measure r (scale_measure r' M) = scale_measure (r * r') M" by (rule measure_eqI) (simp_all add: max_def mult.assoc) lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M" by (rule measure_eqI) simp_all subsection \Complete lattice structure on measures\ lemma (in finite_measure) finite_measure_Diff': "A \ sets M \ B \ sets M \ measure M (A - B) = measure M A - measure M (A \ B)" using finite_measure_Diff[of A "A \ B"] by (auto simp: Diff_Int) lemma (in finite_measure) finite_measure_Union': "A \ sets M \ B \ sets M \ measure M (A \ B) = measure M A + measure M (B - A)" using finite_measure_Union[of A "B - A"] by auto lemma finite_unsigned_Hahn_decomposition: assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M" shows "\Y\sets M. (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ Y = {} \ M X \ N X)" proof - interpret M: finite_measure M by fact interpret N: finite_measure N by fact define d where "d X = measure M X - measure N X" for X have [intro]: "bdd_above (d`sets M)" using sets.sets_into_space[of _ M] by (intro bdd_aboveI[where M="measure M (space M)"]) (auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono) define \ where "\ = (SUP X\sets M. d X)" have le_\[intro]: "X \ sets M \ d X \ \" for X by (auto simp: \_def intro!: cSUP_upper) have "\f. \n. f n \ sets M \ d (f n) > \ - 1 / 2^n" proof (intro choice_iff[THEN iffD1] allI) fix n have "\X\sets M. \ - 1 / 2^n < d X" unfolding \_def by (intro less_cSUP_iff[THEN iffD1]) auto then show "\y. y \ sets M \ \ - 1 / 2 ^ n < d y" by auto qed then obtain E where [measurable]: "E n \ sets M" and E: "d (E n) > \ - 1 / 2^n" for n by auto define F where "F m n = (if m \ n then \i\{m..n}. E i else space M)" for m n have [measurable]: "m \ n \ F m n \ sets M" for m n by (auto simp: F_def) have 1: "\ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" if "m \ n" for m n using that proof (induct rule: dec_induct) case base with E[of m] show ?case by (simp add: F_def field_simps) next case (step i) have F_Suc: "F m (Suc i) = F m i \ E (Suc i)" using \m \ i\ by (auto simp: F_def le_Suc_eq) have "\ + (\ - 2 / 2^m + 1 / 2 ^ Suc i) \ (\ - 1 / 2^Suc i) + (\ - 2 / 2^m + 1 / 2^i)" by (simp add: field_simps) also have "\ \ d (E (Suc i)) + d (F m i)" using E[of "Suc i"] by (intro add_mono step) auto also have "\ = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff') also have "\ = d (E (Suc i) \ F m i) + d (F m (Suc i))" using \m \ i\ by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union') also have "\ \ \ + d (F m (Suc i))" using \m \ i\ by auto finally show ?case by auto qed define F' where "F' m = (\i\{m..}. E i)" for m have F'_eq: "F' m = (\i. F m (i + m))" for m by (fastforce simp: le_iff_add[of m] F'_def F_def) have [measurable]: "F' m \ sets M" for m by (auto simp: F'_def) have \_le: "\ - 0 \ d (\m. F' m)" proof (rule LIMSEQ_le) show "(\n. \ - 2 / 2 ^ n) \ \ - 0" by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto have "incseq F'" by (auto simp: incseq_def F'_def) then show "(\m. d (F' m)) \ d (\m. F' m)" unfolding d_def by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto have "\ - 2 / 2 ^ m + 0 \ d (F' m)" for m proof (rule LIMSEQ_le) have *: "decseq (\n. F m (n + m))" by (auto simp: decseq_def F_def) show "(\n. d (F m n)) \ d (F' m)" unfolding d_def F'_eq by (rule LIMSEQ_offset[where k=m]) (auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *) show "(\n. \ - 2 / 2 ^ m + 1 / 2 ^ n) \ \ - 2 / 2 ^ m + 0" by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto show "\N. \n\N. \ - 2 / 2 ^ m + 1 / 2 ^ n \ d (F m n)" using 1[of m] by (intro exI[of _ m]) auto qed then show "\N. \n\N. \ - 2 / 2 ^ n \ d (F' n)" by auto qed show ?thesis proof (safe intro!: bexI[of _ "\m. F' m"]) fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m)" have "d (\m. F' m) - d X = d ((\m. F' m) - X)" using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff) also have "\ \ \" by auto finally have "0 \ d X" using \_le by auto then show "emeasure N X \ emeasure M X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) next fix X assume [measurable]: "X \ sets M" and X: "X \ (\m. F' m) = {}" then have "d (\m. F' m) + d X = d (X \ (\m. F' m))" by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union) also have "\ \ \" by auto finally have "d X \ 0" using \_le by auto then show "emeasure M X \ emeasure N X" by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure) qed auto qed proposition unsigned_Hahn_decomposition: assumes [simp]: "sets N = sets M" and [measurable]: "A \ sets M" and [simp]: "emeasure M A \ top" "emeasure N A \ top" shows "\Y\sets M. Y \ A \ (\X\sets M. X \ Y \ N X \ M X) \ (\X\sets M. X \ A \ X \ Y = {} \ M X \ N X)" proof - have "\Y\sets (restrict_space M A). (\X\sets (restrict_space M A). X \ Y \ (restrict_space N A) X \ (restrict_space M A) X) \ (\X\sets (restrict_space M A). X \ Y = {} \ (restrict_space M A) X \ (restrict_space N A) X)" proof (rule finite_unsigned_Hahn_decomposition) show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)" by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI) qed (simp add: sets_restrict_space) - then guess Y .. then show ?thesis - apply (intro bexI[of _ Y] conjI ballI conjI) - apply (simp_all add: sets_restrict_space emeasure_restrict_space) - apply safe - subgoal for X Z - by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1) - subgoal for X Z - by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1 ac_simps) - apply auto + apply - + apply (erule bexE) + subgoal for Y + apply (intro bexI[of _ Y] conjI ballI conjI) + apply (simp_all add: sets_restrict_space emeasure_restrict_space) + apply safe + subgoal for X Z + by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1) + subgoal for X Z + by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1 ac_simps) + apply auto + done done qed text\<^marker>\tag important\ \ Define a lexicographical order on \<^type>\measure\, in the order space, sets and measure. The parts of the lexicographical order are point-wise ordered. \ instantiation measure :: (type) order_bot begin inductive less_eq_measure :: "'a measure \ 'a measure \ bool" where "space M \ space N \ less_eq_measure M N" | "space M = space N \ sets M \ sets N \ less_eq_measure M N" | "space M = space N \ sets M = sets N \ emeasure M \ emeasure N \ less_eq_measure M N" lemma le_measure_iff: "M \ N \ (if space M = space N then if sets M = sets N then emeasure M \ emeasure N else sets M \ sets N else space M \ space N)" by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros) definition\<^marker>\tag important\ less_measure :: "'a measure \ 'a measure \ bool" where "less_measure M N \ (M \ N \ \ N \ M)" definition\<^marker>\tag important\ bot_measure :: "'a measure" where "bot_measure = sigma {} {}" lemma shows space_bot[simp]: "space bot = {}" and sets_bot[simp]: "sets bot = {{}}" and emeasure_bot[simp]: "emeasure bot X = 0" by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma) instance proof standard show "bot \ a" for a :: "'a measure" by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def) qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI) end proposition le_measure: "sets M = sets N \ M \ N \ (\A\sets M. emeasure M A \ emeasure N A)" apply - apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq) subgoal for X by (cases "X \ sets M") (auto simp: emeasure_notin_sets) done definition\<^marker>\tag important\ sup_measure' :: "'a measure \ 'a measure \ 'a measure" where "sup_measure' A B = measure_of (space A) (sets A) (\X. SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" lemma assumes [simp]: "sets B = sets A" shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A" and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A" using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def) lemma emeasure_sup_measure': assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \ sets A" shows "emeasure (sup_measure' A B) X = (SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" (is "_ = ?S X") proof - note sets_eq_imp_space_eq[OF sets_eq, simp] show ?thesis using sup_measure'_def proof (rule emeasure_measure_of) let ?d = "\X Y. emeasure A (X \ Y) + emeasure B (X \ - Y)" show "countably_additive (sets (sup_measure' A B)) (\X. SUP Y \ sets A. emeasure A (X \ Y) + emeasure B (X \ - Y))" proof (rule countably_additiveI, goal_cases) case (1 X) then have [measurable]: "\i. X i \ sets A" and "disjoint_family X" by auto have disjoint: "disjoint_family (\i. X i \ Y)" "disjoint_family (\i. X i - Y)" for Y by (auto intro: disjoint_family_on_bisimulation [OF \disjoint_family X\, simplified]) have "(\i. ?S (X i)) = (SUP Y\sets A. \i. ?d (X i) Y)" proof (rule ennreal_suminf_SUP_eq_directed) fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \ sets A" "b \ sets A" have "\c\sets A. c \ X i \ (\a\sets A. ?d (X i) a \ ?d (X i) c)" for i proof cases assume "emeasure A (X i) = top \ emeasure B (X i) = top" then show ?thesis proof assume "emeasure A (X i) = top" then show ?thesis by (intro bexI[of _ "X i"]) auto next assume "emeasure B (X i) = top" then show ?thesis by (intro bexI[of _ "{}"]) auto qed next assume finite: "\ (emeasure A (X i) = top \ emeasure B (X i) = top)" then have "\Y\sets A. Y \ X i \ (\C\sets A. C \ Y \ B C \ A C) \ (\C\sets A. C \ X i \ C \ Y = {} \ A C \ B C)" using unsigned_Hahn_decomposition[of B A "X i"] by simp then obtain Y where [measurable]: "Y \ sets A" and [simp]: "Y \ X i" and B_le_A: "\C. C \ sets A \ C \ Y \ B C \ A C" and A_le_B: "\C. C \ sets A \ C \ X i \ C \ Y = {} \ A C \ B C" by auto show ?thesis proof (intro bexI[of _ Y] ballI conjI) fix a assume [measurable]: "a \ sets A" have *: "(X i \ a \ Y \ (X i \ a - Y)) = X i \ a" "(X i - a) \ Y \ (X i - a - Y) = X i \ - a" for a Y by auto then have "?d (X i) a = (A (X i \ a \ Y) + A (X i \ a \ - Y)) + (B (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ a \ Y) + B (X i \ a \ - Y)) + (A (X i \ - a \ Y) + B (X i \ - a \ - Y))" by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric]) also have "\ \ (A (X i \ Y \ a) + A (X i \ Y \ - a)) + (B (X i \ - Y \ a) + B (X i \ - Y \ - a))" by (simp add: ac_simps) also have "\ \ A (X i \ Y) + B (X i \ - Y)" by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *) finally show "?d (X i) a \ ?d (X i) Y" . qed auto qed then obtain C where [measurable]: "C i \ sets A" and "C i \ X i" and C: "\a. a \ sets A \ ?d (X i) a \ ?d (X i) (C i)" for i by metis have *: "X i \ (\i. C i) = X i \ C i" for i proof safe fix x j assume "x \ X i" "x \ C j" moreover have "i = j \ X i \ X j = {}" using \disjoint_family X\ by (auto simp: disjoint_family_on_def) ultimately show "x \ C i" using \C i \ X i\ \C j \ X j\ by auto qed auto have **: "X i \ - (\i. C i) = X i \ - C i" for i proof safe fix x j assume "x \ X i" "x \ C i" "x \ C j" moreover have "i = j \ X i \ X j = {}" using \disjoint_family X\ by (auto simp: disjoint_family_on_def) ultimately show False using \C i \ X i\ \C j \ X j\ by auto qed auto show "\c\sets A. \i\J. ?d (X i) a \ ?d (X i) c \ ?d (X i) b \ ?d (X i) c" apply (intro bexI[of _ "\i. C i"]) unfolding * ** apply (intro C ballI conjI) apply auto done qed also have "\ = ?S (\i. X i)" apply (simp only: UN_extend_simps(4)) apply (rule arg_cong [of _ _ Sup]) apply (rule image_cong) apply (fact refl) using disjoint apply (auto simp add: suminf_add [symmetric] Diff_eq [symmetric] image_subset_iff suminf_emeasure simp del: UN_simps) done finally show "(\i. ?S (X i)) = ?S (\i. X i)" . qed qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const) qed lemma le_emeasure_sup_measure'1: assumes "sets B = sets A" "X \ sets A" shows "emeasure A X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms) lemma le_emeasure_sup_measure'2: assumes "sets B = sets A" "X \ sets A" shows "emeasure B X \ emeasure (sup_measure' A B) X" by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms) lemma emeasure_sup_measure'_le2: assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \ sets C" assumes A: "\Y. Y \ X \ Y \ sets A \ emeasure A Y \ emeasure C Y" assumes B: "\Y. Y \ X \ Y \ sets A \ emeasure B Y \ emeasure C Y" shows "emeasure (sup_measure' A B) X \ emeasure C X" proof (subst emeasure_sup_measure') show "(SUP Y\sets A. emeasure A (X \ Y) + emeasure B (X \ - Y)) \ emeasure C X" unfolding \sets A = sets C\ proof (intro SUP_least) fix Y assume [measurable]: "Y \ sets C" have [simp]: "X \ Y \ (X - Y) = X" by auto have "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C (X \ Y) + emeasure C (X \ - Y)" by (intro add_mono A B) (auto simp: Diff_eq[symmetric]) also have "\ = emeasure C X" by (subst plus_emeasure) (auto simp: Diff_eq[symmetric]) finally show "emeasure A (X \ Y) + emeasure B (X \ - Y) \ emeasure C X" . qed qed simp_all definition\<^marker>\tag important\ sup_lexord :: "'a \ 'a \ ('a \ 'b::order) \ 'a \ 'a \ 'a" where "sup_lexord A B k s c = (if k A = k B then c else if \ k A \ k B \ \ k B \ k A then s else if k B \ k A then A else B)" lemma sup_lexord: "(k A < k B \ P B) \ (k B < k A \ P A) \ (k A = k B \ P c) \ (\ k B \ k A \ \ k A \ k B \ P s) \ P (sup_lexord A B k s c)" by (auto simp: sup_lexord_def) lemmas le_sup_lexord = sup_lexord[where P="\a. c \ a" for c] lemma sup_lexord1: "k A = k B \ sup_lexord A B k s c = c" by (simp add: sup_lexord_def) lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c" by (auto simp: sup_lexord_def) lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \ \ sets x) = (\ \ sets x)" using sets.sigma_sets_subset[of \ x] by auto lemma sigma_le_iff: "\ \ Pow \ \ sigma \ \ \ x \ (\ \ space x \ (space x = \ \ \ \ sets x))" by (cases "\ = space x") (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def sigma_sets_superset_generator sigma_sets_le_sets_iff) instantiation measure :: (type) semilattice_sup begin definition\<^marker>\tag important\ sup_measure :: "'a measure \ 'a measure \ 'a measure" where "sup_measure A B = sup_lexord A B space (sigma (space A \ space B) {}) (sup_lexord A B sets (sigma (space A) (sets A \ sets B)) (sup_measure' A B))" instance proof fix x y z :: "'a measure" show "x \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume "space x = space y" then have *: "sets x \ sets y \ Pow (space x)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets x \ sets x \ sets y" by auto also have "\ \ sigma (space x) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "x \ sigma (space x) (sets x \ sets y)" by (simp add: space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "x \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "x \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'1) qed (auto intro: less_eq_measure.intros) show "y \ sup x y" unfolding sup_measure_def proof (intro le_sup_lexord) assume **: "space x = space y" then have *: "sets x \ sets y \ Pow (space y)" using sets.space_closed by auto assume "\ sets y \ sets x" "\ sets x \ sets y" then have "sets y \ sets x \ sets y" by auto also have "\ \ sigma (space y) (sets x \ sets y)" by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator) finally show "y \ sigma (space x) (sets x \ sets y)" by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2)) next assume "\ space y \ space x" "\ space x \ space y" then show "y \ sigma (space x \ space y) {}" by (intro less_eq_measure.intros) auto next assume "sets x = sets y" then show "y \ sup_measure' x y" by (simp add: le_measure le_emeasure_sup_measure'2) qed (auto intro: less_eq_measure.intros) show "x \ y \ z \ y \ sup x z \ y" unfolding sup_measure_def proof (intro sup_lexord[where P="\x. x \ y"]) assume "x \ y" "z \ y" and [simp]: "space x = space z" "sets x = sets z" from \x \ y\ show "sup_measure' x z \ y" proof cases case 1 then show ?thesis by (intro less_eq_measure.intros(1)) simp next case 2 then show ?thesis by (intro less_eq_measure.intros(2)) simp_all next case 3 with \z \ y\ \x \ y\ show ?thesis by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2) qed next assume **: "x \ y" "z \ y" "space x = space z" "\ sets z \ sets x" "\ sets x \ sets z" then have *: "sets x \ sets z \ Pow (space x)" using sets.space_closed by auto show "sigma (space x) (sets x \ sets z) \ y" unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm) next assume "x \ y" "z \ y" "\ space z \ space x" "\ space x \ space z" then have "space x \ space y" "space z \ space y" by (auto simp: le_measure_iff split: if_split_asm) then show "sigma (space x \ space z) {} \ y" by (simp add: sigma_le_iff) qed qed end lemma space_empty_eq_bot: "space a = {} \ a = bot" using space_empty[of a] by (auto intro!: measure_eqI) lemma sets_eq_iff_bounded: "A \ B \ B \ C \ sets A = sets C \ sets B = sets A" by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm) lemma sets_sup: "sets A = sets M \ sets B = sets M \ sets (sup A B) = sets M" by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq) lemma le_measureD1: "A \ B \ space A \ space B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD2: "A \ B \ space A = space B \ sets A \ sets B" by (auto simp: le_measure_iff split: if_split_asm) lemma le_measureD3: "A \ B \ sets A = sets B \ emeasure A X \ emeasure B X" by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm) lemma UN_space_closed: "\(sets ` S) \ Pow (\(space ` S))" using sets.space_closed by auto definition\<^marker>\tag important\ Sup_lexord :: "('a \ 'b::complete_lattice) \ ('a set \ 'a) \ ('a set \ 'a) \ 'a set \ 'a" where "Sup_lexord k c s A = (let U = (SUP a\A. k a) in if \a\A. k a = U then c {a\A. k a = U} else s A)" lemma Sup_lexord: "(\a S. a \ A \ k a = (SUP a\A. k a) \ S = {a'\A. k a' = k a} \ P (c S)) \ ((\a. a \ A \ k a \ (SUP a\A. k a)) \ P (s A)) \ P (Sup_lexord k c s A)" by (auto simp: Sup_lexord_def Let_def) lemma Sup_lexord1: assumes A: "A \ {}" "(\a. a \ A \ k a = (\a\A. k a))" "P (c A)" shows "P (Sup_lexord k c s A)" unfolding Sup_lexord_def Let_def proof (clarsimp, safe) show "\a\A. k a \ (\x\A. k x) \ P (s A)" by (metis assms(1,2) ex_in_conv) next fix a assume "a \ A" "k a = (\x\A. k x)" then have "{a \ A. k a = (\x\A. k x)} = {a \ A. k a = k a}" by (metis A(2)[symmetric]) then show "P (c {a \ A. k a = (\x\A. k x)})" by (simp add: A(3)) qed instantiation measure :: (type) complete_lattice begin interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure" by standard (auto intro!: antisym) lemma sup_measure_F_mono': "finite J \ finite I \ sup_measure.F id I \ sup_measure.F id (I \ J)" proof (induction J rule: finite_induct) case empty then show ?case by simp next case (insert i J) show ?case proof cases assume "i \ I" with insert show ?thesis by (auto simp: insert_absorb) next assume "i \ I" have "sup_measure.F id I \ sup_measure.F id (I \ J)" by (intro insert) also have "\ \ sup_measure.F id (insert i (I \ J))" using insert \i \ I\ by (subst sup_measure.insert) auto finally show ?thesis by auto qed qed lemma sup_measure_F_mono: "finite I \ J \ I \ sup_measure.F id J \ sup_measure.F id I" using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1) lemma sets_sup_measure_F: "finite I \ I \ {} \ (\i. i \ I \ sets i = sets M) \ sets (sup_measure.F id I) = sets M" by (induction I rule: finite_ne_induct) (simp_all add: sets_sup) definition\<^marker>\tag important\ Sup_measure' :: "'a measure set \ 'a measure" where "Sup_measure' M = measure_of (\a\M. space a) (\a\M. sets a) (\X. (SUP P\{P. finite P \ P \ M }. sup_measure.F id P X))" lemma space_Sup_measure'2: "space (Sup_measure' M) = (\m\M. space m)" unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed]) lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\m\M. space m) (\m\M. sets m)" unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed]) lemma sets_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "sets (Sup_measure' M) = sets A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def) lemma space_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "M \ {}" shows "space (Sup_measure' M) = space A" using sets_eq[THEN sets_eq_imp_space_eq, simp] \M \ {}\ by (simp add: Sup_measure'_def ) lemma emeasure_Sup_measure': assumes sets_eq[simp]: "\m. m \ M \ sets m = sets A" and "X \ sets A" "M \ {}" shows "emeasure (Sup_measure' M) X = (SUP P\{P. finite P \ P \ M}. sup_measure.F id P X)" (is "_ = ?S X") using Sup_measure'_def proof (rule emeasure_measure_of) note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A" using \M \ {}\ by (simp_all add: Sup_measure'_def) let ?\ = "sup_measure.F id" show "countably_additive (sets (Sup_measure' M)) ?S" proof (rule countably_additiveI, goal_cases) case (1 F) then have **: "range F \ sets A" by (auto simp: *) show "(\i. ?S (F i)) = ?S (\i. F i)" proof (subst ennreal_suminf_SUP_eq_directed) fix i j and N :: "nat set" assume ij: "i \ {P. finite P \ P \ M}" "j \ {P. finite P \ P \ M}" have "(i \ {} \ sets (?\ i) = sets A) \ (j \ {} \ sets (?\ j) = sets A) \ (i \ {} \ j \ {} \ sets (?\ (i \ j)) = sets A)" using ij by (intro impI sets_sup_measure_F conjI) auto then have "?\ j (F n) \ ?\ (i \ j) (F n) \ ?\ i (F n) \ ?\ (i \ j) (F n)" for n using ij by (cases "i = {}"; cases "j = {}") (auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F simp del: id_apply) with ij show "\k\{P. finite P \ P \ M}. \n\N. ?\ i (F n) \ ?\ k (F n) \ ?\ j (F n) \ ?\ k (F n)" by (safe intro!: bexI[of _ "i \ j"]) auto next show "(SUP P \ {P. finite P \ P \ M}. \n. ?\ P (F n)) = (SUP P \ {P. finite P \ P \ M}. ?\ P (\(F ` UNIV)))" proof (intro arg_cong [of _ _ Sup] image_cong refl) fix i assume i: "i \ {P. finite P \ P \ M}" show "(\n. ?\ i (F n)) = ?\ i (\(F ` UNIV))" proof cases assume "i \ {}" with i ** show ?thesis apply (intro suminf_emeasure \disjoint_family F\) apply (subst sets_sup_measure_F[OF _ _ sets_eq]) apply auto done qed simp qed qed qed show "positive (sets (Sup_measure' M)) ?S" by (auto simp: positive_def bot_ennreal[symmetric]) show "X \ sets (Sup_measure' M)" using assms * by auto qed (rule UN_space_closed) definition\<^marker>\tag important\ Sup_measure :: "'a measure set \ 'a measure" where "Sup_measure = Sup_lexord space (Sup_lexord sets Sup_measure' (\U. sigma (\u\U. space u) (\u\U. sets u))) (\U. sigma (\u\U. space u) {})" definition\<^marker>\tag important\ Inf_measure :: "'a measure set \ 'a measure" where "Inf_measure A = Sup {x. \a\A. x \ a}" definition\<^marker>\tag important\ inf_measure :: "'a measure \ 'a measure \ 'a measure" where "inf_measure a b = Inf {a, b}" definition\<^marker>\tag important\ top_measure :: "'a measure" where "top_measure = Inf {}" instance proof note UN_space_closed [simp] show upper: "x \ Sup A" if x: "x \ A" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. x \ y"]) assume "\a. a \ A \ space a \ (\a\A. space a)" from this[OF \x \ A\] \x \ A\ show "x \ sigma (\a\A. space a) {}" by (intro less_eq_measure.intros) auto next fix a S assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and neq: "\aa. aa \ S \ sets aa \ (\a\S. sets a)" have sp_a: "space a = (\(space ` S))" using \a\A\ by (auto simp: S) show "x \ sigma (\(space ` S)) (\(sets ` S))" proof cases assume [simp]: "space x = space a" have "sets x \ (\a\S. sets a)" using \x\A\ neq[of x] by (auto simp: S) also have "\ \ sigma_sets (\x\S. space x) (\x\S. sets x)" by (rule sigma_sets_superset_generator) finally show ?thesis by (intro less_eq_measure.intros(2)) (simp_all add: sp_a) next assume "space x \ space a" moreover have "space x \ space a" unfolding a using \x\A\ by auto ultimately show ?thesis by (intro less_eq_measure.intros) (simp add: less_le sp_a) qed next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "x \ Sup_measure' S'" proof cases assume "x \ S" with \b \ S\ have "space x = space b" by (simp add: S) show ?thesis proof cases assume "x \ S'" show "x \ Sup_measure' S'" proof (intro le_measure[THEN iffD2] ballI) show "sets x = sets (Sup_measure' S')" using \x\S'\ * by (simp add: S') fix X assume "X \ sets x" show "emeasure x X \ emeasure (Sup_measure' S') X" proof (subst emeasure_Sup_measure'[OF _ \X \ sets x\]) show "emeasure x X \ (SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X)" using \x\S'\ by (intro SUP_upper2[where i="{x}"]) auto qed (insert \x\S'\ S', auto) qed next assume "x \ S'" then have "sets x \ sets b" using \x\S\ by (auto simp: S') moreover have "sets x \ sets b" using \x\S\ unfolding b by auto ultimately show ?thesis using * \x \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * \space x = space b\ less_le) qed next assume "x \ S" with \x\A\ \x \ S\ \space b = space a\ show ?thesis by (intro less_eq_measure.intros) (simp_all add: * less_le a SUP_upper S) qed qed show least: "Sup A \ x" if x: "\z. z \ A \ z \ x" for x :: "'a measure" and A unfolding Sup_measure_def proof (intro Sup_lexord[where P="\y. y \ x"]) assume "\a. a \ A \ space a \ (\a\A. space a)" show "sigma (\(space ` A)) {} \ x" using x[THEN le_measureD1] by (subst sigma_le_iff) auto next fix a S assume "a \ A" "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" "\a. a \ S \ sets a \ (\a\S. sets a)" have "\(space ` S) \ space x" using S le_measureD1[OF x] by auto moreover have "\(space ` S) = space a" using \a\A\ S by auto then have "space x = \(space ` S) \ \(sets ` S) \ sets x" using \a \ A\ le_measureD2[OF x] by (auto simp: S) ultimately show "sigma (\(space ` S)) (\(sets ` S)) \ x" by (subst sigma_le_iff) simp_all next fix a b S S' assume "a \ A" and a: "space a = (\a\A. space a)" and S: "S = {a' \ A. space a' = space a}" and "b \ S" and b: "sets b = (\a\S. sets a)" and S': "S' = {a' \ S. sets a' = sets b}" then have "S' \ {}" "space b = space a" by auto have sets_eq: "\x. x \ S' \ sets x = sets b" by (auto simp: S') note sets_eq[THEN sets_eq_imp_space_eq, simp] have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b" using \S' \ {}\ by (simp_all add: Sup_measure'_def sets_eq) show "Sup_measure' S' \ x" proof cases assume "space x = space a" show ?thesis proof cases assume **: "sets x = sets b" show ?thesis proof (intro le_measure[THEN iffD2] ballI) show ***: "sets (Sup_measure' S') = sets x" by (simp add: * **) fix X assume "X \ sets (Sup_measure' S')" show "emeasure (Sup_measure' S') X \ emeasure x X" unfolding *** proof (subst emeasure_Sup_measure'[OF _ \X \ sets (Sup_measure' S')\]) show "(SUP P \ {P. finite P \ P \ S'}. emeasure (sup_measure.F id P) X) \ emeasure x X" proof (safe intro!: SUP_least) fix P assume P: "finite P" "P \ S'" show "emeasure (sup_measure.F id P) X \ emeasure x X" proof cases assume "P = {}" then show ?thesis by auto next assume "P \ {}" from P have "finite P" "P \ A" unfolding S' S by (simp_all add: subset_eq) then have "sup_measure.F id P \ x" by (induction P) (auto simp: x) moreover have "sets (sup_measure.F id P) = sets x" using \finite P\ \P \ {}\ \P \ S'\ \sets x = sets b\ by (intro sets_sup_measure_F) (auto simp: S') ultimately show "emeasure (sup_measure.F id P) X \ emeasure x X" by (rule le_measureD3) qed qed show "m \ S' \ sets m = sets (Sup_measure' S')" for m unfolding * by (simp add: S') qed fact qed next assume "sets x \ sets b" moreover have "sets b \ sets x" unfolding b S using x[THEN le_measureD2] \space x = space a\ by auto ultimately show "Sup_measure' S' \ x" using \space x = space a\ \b \ S\ by (intro less_eq_measure.intros(2)) (simp_all add: * S) qed next assume "space x \ space a" then have "space a < space x" using le_measureD1[OF x[OF \a\A\]] by auto then show "Sup_measure' S' \ x" by (intro less_eq_measure.intros) (simp add: * \space b = space a\) qed qed show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)" by (auto intro!: antisym least simp: top_measure_def) show lower: "x \ A \ Inf A \ x" for x :: "'a measure" and A unfolding Inf_measure_def by (intro least) auto show greatest: "(\z. z \ A \ x \ z) \ x \ Inf A" for x :: "'a measure" and A unfolding Inf_measure_def by (intro upper) auto show "inf x y \ x" "inf x y \ y" "x \ y \ x \ z \ x \ inf y z" for x y z :: "'a measure" by (auto simp: inf_measure_def intro!: lower greatest) qed end lemma sets_SUP: assumes "\x. x \ I \ sets (M x) = sets N" shows "I \ {} \ sets (SUP i\I. M i) = sets N" unfolding Sup_measure_def using assms assms[THEN sets_eq_imp_space_eq] sets_Sup_measure'[where A=N and M="M`I"] by (intro Sup_lexord1[where P="\x. sets x = sets N"]) auto lemma emeasure_SUP: assumes sets: "\i. i \ I \ sets (M i) = sets N" "X \ sets N" "I \ {}" shows "emeasure (SUP i\I. M i) X = (SUP J\{J. J \ {} \ finite J \ J \ I}. emeasure (SUP i\J. M i) X)" proof - interpret sup_measure: comm_monoid_set sup "bot :: 'b measure" by standard (auto intro!: antisym) have eq: "finite J \ sup_measure.F id J = (SUP i\J. i)" for J :: "'b measure set" by (induction J rule: finite_induct) auto have 1: "J \ {} \ J \ I \ sets (SUP x\J. M x) = sets N" for J by (intro sets_SUP sets) (auto ) from \I \ {}\ obtain i where "i\I" by auto have "Sup_measure' (M`I) X = (SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X)" using sets by (intro emeasure_Sup_measure') auto also have "Sup_measure' (M`I) = (SUP i\I. M i)" unfolding Sup_measure_def using \I \ {}\ sets sets(1)[THEN sets_eq_imp_space_eq] by (intro Sup_lexord1[where P="\x. _ = x"]) auto also have "(SUP P\{P. finite P \ P \ M`I}. sup_measure.F id P X) = (SUP J\{J. J \ {} \ finite J \ J \ I}. (SUP i\J. M i) X)" proof (intro SUP_eq) fix J assume "J \ {P. finite P \ P \ M`I}" then obtain J' where J': "J' \ I" "finite J'" and J: "J = M`J'" and "finite J" using finite_subset_image[of J M I] by auto show "\j\{J. J \ {} \ finite J \ J \ I}. sup_measure.F id J X \ (SUP i\j. M i) X" proof cases assume "J' = {}" with \i \ I\ show ?thesis by (auto simp add: J) next assume "J' \ {}" with J J' show ?thesis by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply) qed next fix J assume J: "J \ {P. P \ {} \ finite P \ P \ I}" show "\J'\{J. finite J \ J \ M`I}. (SUP i\J. M i) X \ sup_measure.F id J' X" using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply) qed finally show ?thesis . qed lemma emeasure_SUP_chain: assumes sets: "\i. i \ A \ sets (M i) = sets N" "X \ sets N" assumes ch: "Complete_Partial_Order.chain (\) (M ` A)" and "A \ {}" shows "emeasure (SUP i\A. M i) X = (SUP i\A. emeasure (M i) X)" proof (subst emeasure_SUP[OF sets \A \ {}\]) show "(SUP J\{J. J \ {} \ finite J \ J \ A}. emeasure (Sup (M ` J)) X) = (SUP i\A. emeasure (M i) X)" proof (rule SUP_eq) fix J assume "J \ {J. J \ {} \ finite J \ J \ A}" then have J: "Complete_Partial_Order.chain (\) (M ` J)" "finite J" "J \ {}" and "J \ A" using ch[THEN chain_subset, of "M`J"] by auto with in_chain_finite[OF J(1)] obtain j where "j \ J" "(SUP j\J. M j) = M j" by auto with \J \ A\ show "\j\A. emeasure (Sup (M ` J)) X \ emeasure (M j) X" by auto next fix j assume "j\A" then show "\i\{J. J \ {} \ finite J \ J \ A}. emeasure (M j) X \ emeasure (Sup (M ` i)) X" by (intro bexI[of _ "{j}"]) auto qed qed subsubsection\<^marker>\tag unimportant\ \Supremum of a set of \\\-algebras\ lemma space_Sup_eq_UN: "space (Sup M) = (\x\M. space x)" unfolding Sup_measure_def apply (intro Sup_lexord[where P="\x. space x = _"]) apply (subst space_Sup_measure'2) apply auto [] apply (subst space_measure_of[OF UN_space_closed]) apply auto done lemma sets_Sup_eq: assumes *: "\m. m \ M \ space m = X" and "M \ {}" shows "sets (Sup M) = sigma_sets X (\x\M. sets x)" unfolding Sup_measure_def apply (rule Sup_lexord1) apply fact apply (simp add: assms) apply (rule Sup_lexord) subgoal premises that for a S unfolding that(3) that(2)[symmetric] using that(1) apply (subst sets_Sup_measure'2) apply (intro arg_cong2[where f=sigma_sets]) apply (auto simp: *) done apply (subst sets_measure_of[OF UN_space_closed]) apply (simp add: assms) done lemma in_sets_Sup: "(\m. m \ M \ space m = X) \ m \ M \ A \ sets m \ A \ sets (Sup M)" by (subst sets_Sup_eq[where X=X]) auto lemma Sup_lexord_rel: assumes "\i. i \ I \ k (A i) = k (B i)" "R (c (A ` {a \ I. k (B a) = (SUP x\I. k (B x))})) (c (B ` {a \ I. k (B a) = (SUP x\I. k (B x))}))" "R (s (A`I)) (s (B`I))" shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))" proof - have "A ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ A ` I. k a = (SUP x\I. k (B x))}" using assms(1) by auto moreover have "B ` {a \ I. k (B a) = (SUP x\I. k (B x))} = {a \ B ` I. k a = (SUP x\I. k (B x))}" by auto ultimately show ?thesis using assms by (auto simp: Sup_lexord_def Let_def image_comp) qed lemma sets_SUP_cong: assumes eq: "\i. i \ I \ sets (M i) = sets (N i)" shows "sets (SUP i\I. M i) = sets (SUP i\I. N i)" unfolding Sup_measure_def using eq eq[THEN sets_eq_imp_space_eq] apply (intro Sup_lexord_rel[where R="\x y. sets x = sets y"]) apply simp apply simp apply (simp add: sets_Sup_measure'2) apply (intro arg_cong2[where f="\x y. sets (sigma x y)"]) apply auto done lemma sets_Sup_in_sets: assumes "M \ {}" assumes "\m. m \ M \ space m = space N" assumes "\m. m \ M \ sets m \ sets N" shows "sets (Sup M) \ sets N" proof - have *: "\(space ` M) = space N" using assms by auto show ?thesis unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset) qed lemma measurable_Sup1: assumes m: "m \ M" and f: "f \ measurable m N" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable (Sup M) N" proof - have "space (Sup M) = space m" using m by (auto simp add: space_Sup_eq_UN dest: const_space) then show ?thesis using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space]) qed lemma measurable_Sup2: assumes M: "M \ {}" assumes f: "\m. m \ M \ f \ measurable N m" and const_space: "\m n. m \ M \ n \ M \ space m = space n" shows "f \ measurable N (Sup M)" proof - from M obtain m where "m \ M" by auto have space_eq: "\n. n \ M \ space n = space m" by (intro const_space \m \ M\) have "f \ measurable N (sigma (\m\M. space m) (\m\M. sets m))" proof (rule measurable_measure_of) show "f \ space N \ \(space ` M)" using measurable_space[OF f] M by auto qed (auto intro: measurable_sets f dest: sets.sets_into_space) also have "measurable N (sigma (\m\M. space m) (\m\M. sets m)) = measurable N (Sup M)" apply (intro measurable_cong_sets refl) apply (subst sets_Sup_eq[OF space_eq M]) apply simp apply (subst sets_measure_of[OF UN_space_closed]) apply (simp add: space_eq M) done finally show ?thesis . qed lemma measurable_SUP2: "I \ {} \ (\i. i \ I \ f \ measurable N (M i)) \ (\i j. i \ I \ j \ I \ space (M i) = space (M j)) \ f \ measurable N (SUP i\I. M i)" by (auto intro!: measurable_Sup2) lemma sets_Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" proof - { fix a m assume "a \ sigma_sets \ m" "m \ M" then have "a \ sigma_sets \ (\M)" by induction (auto intro: sigma_sets.intros(2-)) } then show "sets (SUP m\M. sigma \ m) = sets (sigma \ (\M))" apply (subst sets_Sup_eq[where X="\"]) apply (auto simp add: M) [] apply auto [] apply (simp add: space_measure_of_conv M Union_least) apply (rule sigma_sets_eqI) apply auto done qed lemma Sup_sigma: assumes [simp]: "M \ {}" and M: "\m. m \ M \ m \ Pow \" shows "(SUP m\M. sigma \ m) = (sigma \ (\M))" proof (intro antisym SUP_least) have *: "\M \ Pow \" using M by auto show "sigma \ (\M) \ (SUP m\M. sigma \ m)" proof (intro less_eq_measure.intros(3)) show "space (sigma \ (\M)) = space (SUP m\M. sigma \ m)" "sets (sigma \ (\M)) = sets (SUP m\M. sigma \ m)" using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq] by auto qed (simp add: emeasure_sigma le_fun_def) fix m assume "m \ M" then show "sigma \ m \ sigma \ (\M)" by (subst sigma_le_iff) (auto simp add: M *) qed lemma SUP_sigma_sigma: "M \ {} \ (\m. m \ M \ f m \ Pow \) \ (SUP m\M. sigma \ (f m)) = sigma \ (\m\M. f m)" using Sup_sigma[of "f`M" \] by (auto simp: image_comp) lemma sets_vimage_Sup_eq: assumes *: "M \ {}" "f \ X \ Y" "\m. m \ M \ space m = Y" shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m \ M. vimage_algebra X f m)" (is "?IS = ?SI") proof show "?IS \ ?SI" apply (intro sets_image_in_sets measurable_Sup2) apply (simp add: space_Sup_eq_UN *) apply (simp add: *) apply (intro measurable_Sup1) apply (rule imageI) apply assumption apply (rule measurable_vimage_algebra1) apply (auto simp: *) done show "?SI \ ?IS" apply (intro sets_Sup_in_sets) apply (auto simp: *) [] apply (auto simp: *) [] apply (elim imageE) apply simp apply (rule sets_image_in_sets) apply simp apply (simp add: measurable_def) apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2) apply (auto intro: in_sets_Sup[OF *(3)]) done qed lemma restrict_space_eq_vimage_algebra': "sets (restrict_space M \) = sets (vimage_algebra (\ \ space M) (\x. x) M)" proof - have *: "{A \ (\ \ space M) |A. A \ sets M} = {A \ \ |A. A \ sets M}" using sets.sets_into_space[of _ M] by blast show ?thesis unfolding restrict_space_def by (subst sets_measure_of) (auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets]) qed lemma sigma_le_sets: assumes [simp]: "A \ Pow X" shows "sets (sigma X A) \ sets N \ X \ sets N \ A \ sets N" proof have "X \ sigma_sets X A" "A \ sigma_sets X A" by (auto intro: sigma_sets_top) moreover assume "sets (sigma X A) \ sets N" ultimately show "X \ sets N \ A \ sets N" by auto next assume *: "X \ sets N \ A \ sets N" { fix Y assume "Y \ sigma_sets X A" from this * have "Y \ sets N" by induction auto } then show "sets (sigma X A) \ sets N" by auto qed lemma measurable_iff_sets: "f \ measurable M N \ (f \ space M \ space N \ sets (vimage_algebra (space M) f N) \ sets M)" proof - have *: "{f -` A \ space M |A. A \ sets N} \ Pow (space M)" by auto show ?thesis unfolding measurable_def by (auto simp add: vimage_algebra_def sigma_le_sets[OF *]) qed lemma sets_vimage_algebra_space: "X \ sets (vimage_algebra X f M)" using sets.top[of "vimage_algebra X f M"] by simp lemma measurable_mono: assumes N: "sets N' \ sets N" "space N = space N'" assumes M: "sets M \ sets M'" "space M = space M'" shows "measurable M N \ measurable M' N'" unfolding measurable_def proof safe fix f A assume "f \ space M \ space N" "A \ sets N'" moreover assume "\y\sets N. f -` y \ space M \ sets M" note this[THEN bspec, of A] ultimately show "f -` A \ space M' \ sets M'" using assms by auto qed (insert N M, auto) lemma measurable_Sup_measurable: assumes f: "f \ space N \ A" shows "f \ measurable N (Sup {M. space M = A \ f \ measurable N M})" proof (rule measurable_Sup2) show "{M. space M = A \ f \ measurable N M} \ {}" using f unfolding ex_in_conv[symmetric] by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of) qed auto lemma (in sigma_algebra) sigma_sets_subset': assumes a: "a \ M" "\' \ M" shows "sigma_sets \' a \ M" proof show "x \ M" if x: "x \ sigma_sets \' a" for x using x by (induct rule: sigma_sets.induct) (insert a, auto) qed lemma in_sets_SUP: "i \ I \ (\i. i \ I \ space (M i) = Y) \ X \ sets (M i) \ X \ sets (SUP i\I. M i)" by (intro in_sets_Sup[where X=Y]) auto lemma measurable_SUP1: "i \ I \ f \ measurable (M i) N \ (\m n. m \ I \ n \ I \ space (M m) = space (M n)) \ f \ measurable (SUP i\I. M i) N" by (auto intro: measurable_Sup1) lemma sets_image_in_sets': assumes X: "X \ sets N" assumes f: "\A. A \ sets M \ f -` A \ X \ sets N" shows "sets (vimage_algebra X f M) \ sets N" unfolding sets_vimage_algebra by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f) lemma mono_vimage_algebra: "sets M \ sets N \ sets (vimage_algebra X f M) \ sets (vimage_algebra X f N)" using sets.top[of "sigma X {f -` A \ X |A. A \ sets N}"] unfolding vimage_algebra_def apply (subst (asm) space_measure_of) apply auto [] apply (subst sigma_le_sets) apply auto done lemma mono_restrict_space: "sets M \ sets N \ sets (restrict_space M X) \ sets (restrict_space N X)" unfolding sets_restrict_space by (rule image_mono) lemma sets_eq_bot: "sets M = {{}} \ M = bot" apply safe apply (intro measure_eqI) apply auto done lemma sets_eq_bot2: "{{}} = sets M \ M = bot" using sets_eq_bot[of M] by blast lemma (in finite_measure) countable_support: "countable {x. measure M {x} \ 0}" proof cases assume "measure M (space M) = 0" with bounded_measure measure_le_0_iff have "{x. measure M {x} \ 0} = {}" by auto then show ?thesis by simp next let ?M = "measure M (space M)" and ?m = "\x. measure M {x}" assume "?M \ 0" then have *: "{x. ?m x \ 0} = (\n. {x. ?M / Suc n < ?m x})" using reals_Archimedean[of "?m x / ?M" for x] by (auto simp: field_simps not_le[symmetric] divide_le_0_iff measure_le_0_iff) have **: "\n. finite {x. ?M / Suc n < ?m x}" proof (rule ccontr) fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X") then obtain X where "finite X" "card X = Suc (Suc n)" "X \ ?X" by (metis infinite_arbitrarily_large) from this(3) have *: "\x. x \ X \ ?M / Suc n \ ?m x" by auto { fix x assume "x \ X" from \?M \ 0\ *[OF this] have "?m x \ 0" by (auto simp: field_simps measure_le_0_iff) then have "{x} \ sets M" by (auto dest: measure_notin_sets) } note singleton_sets = this have "?M < (\x\X. ?M / Suc n)" using \?M \ 0\ by (simp add: \card X = Suc (Suc n)\ field_simps less_le) also have "\ \ (\x\X. ?m x)" by (rule sum_mono) fact also have "\ = measure M (\x\X. {x})" using singleton_sets \finite X\ by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def) finally have "?M < measure M (\x\X. {x})" . moreover have "measure M (\x\X. {x}) \ ?M" using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto ultimately show False by simp qed show ?thesis unfolding * by (intro countable_UN countableI_type countable_finite[OF **]) qed end diff --git a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy +++ b/src/HOL/Analysis/Nonnegative_Lebesgue_Integration.thy @@ -1,2633 +1,2647 @@ (* Title: HOL/Analysis/Nonnegative_Lebesgue_Integration.thy Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Lebesgue Integration for Nonnegative Functions\ theory Nonnegative_Lebesgue_Integration imports Measure_Space Borel_Space begin subsection\<^marker>\tag unimportant\ \Approximating functions\ lemma AE_upper_bound_inf_ennreal: fixes F G::"'a \ ennreal" assumes "\e. (e::real) > 0 \ AE x in M. F x \ G x + e" shows "AE x in M. F x \ G x" proof - have "AE x in M. \n::nat. F x \ G x + ennreal (1 / Suc n)" using assms by (auto simp: AE_all_countable) then show ?thesis proof (eventually_elim) fix x assume x: "\n::nat. F x \ G x + ennreal (1 / Suc n)" show "F x \ G x" proof (rule ennreal_le_epsilon) fix e :: real assume "0 < e" then obtain n where n: "1 / Suc n < e" by (blast elim: nat_approx_posE) have "F x \ G x + 1 / Suc n" using x by simp also have "\ \ G x + e" using n by (intro add_mono ennreal_leI) auto finally show "F x \ G x + ennreal e" . qed qed qed lemma AE_upper_bound_inf: fixes F G::"'a \ real" assumes "\e. e > 0 \ AE x in M. F x \ G x + e" shows "AE x in M. F x \ G x" proof - have "AE x in M. F x \ G x + 1/real (n+1)" for n::nat by (rule assms, auto) then have "AE x in M. \n::nat \ UNIV. F x \ G x + 1/real (n+1)" by (rule AE_ball_countable', auto) moreover { fix x assume i: "\n::nat \ UNIV. F x \ G x + 1/real (n+1)" have "(\n. G x + 1/real (n+1)) \ G x + 0" by (rule tendsto_add, simp, rule LIMSEQ_ignore_initial_segment[OF lim_1_over_n, of 1]) then have "F x \ G x" using i LIMSEQ_le_const by fastforce } ultimately show ?thesis by auto qed lemma not_AE_zero_ennreal_E: fixes f::"'a \ ennreal" assumes "\ (AE x in M. f x = 0)" and [measurable]: "f \ borel_measurable M" shows "\A\sets M. \e::real>0. emeasure M A > 0 \ (\x \ A. f x \ e)" proof - { assume "\ (\e::real>0. {x \ space M. f x \ e} \ null_sets M)" then have "0 < e \ AE x in M. f x \ e" for e :: real by (auto simp: not_le less_imp_le dest!: AE_not_in) then have "AE x in M. f x \ 0" by (intro AE_upper_bound_inf_ennreal[where G="\_. 0"]) simp then have False using assms by auto } then obtain e::real where e: "e > 0" "{x \ space M. f x \ e} \ null_sets M" by auto define A where "A = {x \ space M. f x \ e}" have 1 [measurable]: "A \ sets M" unfolding A_def by auto have 2: "emeasure M A > 0" using e(2) A_def \A \ sets M\ by auto have 3: "\x. x \ A \ f x \ e" unfolding A_def by auto show ?thesis using e(1) 1 2 3 by blast qed lemma not_AE_zero_E: fixes f::"'a \ real" assumes "AE x in M. f x \ 0" "\(AE x in M. f x = 0)" and [measurable]: "f \ borel_measurable M" shows "\A e. A \ sets M \ e>0 \ emeasure M A > 0 \ (\x \ A. f x \ e)" proof - have "\e. e > 0 \ {x \ space M. f x \ e} \ null_sets M" proof (rule ccontr) assume *: "\(\e. e > 0 \ {x \ space M. f x \ e} \ null_sets M)" { fix e::real assume "e > 0" then have "{x \ space M. f x \ e} \ null_sets M" using * by blast then have "AE x in M. x \ {x \ space M. f x \ e}" using AE_not_in by blast then have "AE x in M. f x \ e" by auto } then have "AE x in M. f x \ 0" by (rule AE_upper_bound_inf, auto) then have "AE x in M. f x = 0" using assms(1) by auto then show False using assms(2) by auto qed then obtain e where e: "e>0" "{x \ space M. f x \ e} \ null_sets M" by auto define A where "A = {x \ space M. f x \ e}" have 1 [measurable]: "A \ sets M" unfolding A_def by auto have 2: "emeasure M A > 0" using e(2) A_def \A \ sets M\ by auto have 3: "\x. x \ A \ f x \ e" unfolding A_def by auto show ?thesis using e(1) 1 2 3 by blast qed subsection "Simple function" text \ Our simple functions are not restricted to nonnegative real numbers. Instead they are just functions with a finite range and are measurable when singleton sets are measurable. \ definition\<^marker>\tag important\ "simple_function M g \ finite (g ` space M) \ (\x \ g ` space M. g -` {x} \ space M \ sets M)" lemma simple_functionD: assumes "simple_function M g" shows "finite (g ` space M)" and "g -` X \ space M \ sets M" proof - show "finite (g ` space M)" using assms unfolding simple_function_def by auto have "g -` X \ space M = g -` (X \ g`space M) \ space M" by auto also have "\ = (\x\X \ g`space M. g-`{x} \ space M)" by auto finally show "g -` X \ space M \ sets M" using assms by (auto simp del: UN_simps simp: simple_function_def) qed lemma measurable_simple_function[measurable_dest]: "simple_function M f \ f \ measurable M (count_space UNIV)" unfolding simple_function_def measurable_def proof safe fix A assume "finite (f ` space M)" "\x\f ` space M. f -` {x} \ space M \ sets M" then have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) \ sets M" by (intro sets.finite_UN) auto also have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) = f -` A \ space M" by (auto split: if_split_asm) finally show "f -` A \ space M \ sets M" . qed simp lemma borel_measurable_simple_function: "simple_function M f \ f \ borel_measurable M" by (auto dest!: measurable_simple_function simp: measurable_def) lemma simple_function_measurable2[intro]: assumes "simple_function M f" "simple_function M g" shows "f -` A \ g -` B \ space M \ sets M" proof - have "f -` A \ g -` B \ space M = (f -` A \ space M) \ (g -` B \ space M)" by auto then show ?thesis using assms[THEN simple_functionD(2)] by auto qed lemma simple_function_indicator_representation: fixes f ::"'a \ ennreal" assumes f: "simple_function M f" and x: "x \ space M" shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" (is "?l = ?r") proof - have "?r = (\y \ f ` space M. (if y = f x then y * indicator (f -` {y} \ space M) x else 0))" by (auto intro!: sum.cong) also have "... = f x * indicator (f -` {f x} \ space M) x" using assms by (auto dest: simple_functionD) also have "... = f x" using x by (auto simp: indicator_def) finally show ?thesis by auto qed lemma simple_function_notspace: "simple_function M (\x. h x * indicator (- space M) x::ennreal)" (is "simple_function M ?h") proof - have "?h ` space M \ {0}" unfolding indicator_def by auto hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) have "?h -` {0} \ space M = space M" by auto thus ?thesis unfolding simple_function_def by (auto simp add: image_constant_conv) qed lemma simple_function_cong: assumes "\t. t \ space M \ f t = g t" shows "simple_function M f \ simple_function M g" proof - have "\x. f -` {x} \ space M = g -` {x} \ space M" using assms by auto with assms show ?thesis by (simp add: simple_function_def cong: image_cong) qed lemma simple_function_cong_algebra: assumes "sets N = sets M" "space N = space M" shows "simple_function M f \ simple_function N f" unfolding simple_function_def assms .. lemma simple_function_borel_measurable: fixes f :: "'a \ 'x::{t2_space}" assumes "f \ borel_measurable M" and "finite (f ` space M)" shows "simple_function M f" using assms unfolding simple_function_def by (auto intro: borel_measurable_vimage) lemma simple_function_iff_borel_measurable: fixes f :: "'a \ 'x::{t2_space}" shows "simple_function M f \ finite (f ` space M) \ f \ borel_measurable M" by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable) lemma simple_function_eq_measurable: "simple_function M f \ finite (f`space M) \ f \ measurable M (count_space UNIV)" using measurable_simple_function[of M f] by (fastforce simp: simple_function_def) lemma simple_function_const[intro, simp]: "simple_function M (\x. c)" by (auto intro: finite_subset simp: simple_function_def) lemma simple_function_compose[intro, simp]: assumes "simple_function M f" shows "simple_function M (g \ f)" unfolding simple_function_def proof safe show "finite ((g \ f) ` space M)" using assms unfolding simple_function_def image_comp [symmetric] by auto next fix x assume "x \ space M" let ?G = "g -` {g (f x)} \ (f`space M)" have *: "(g \ f) -` {(g \ f) x} \ space M = (\x\?G. f -` {x} \ space M)" by auto show "(g \ f) -` {(g \ f) x} \ space M \ sets M" using assms unfolding simple_function_def * by (rule_tac sets.finite_UN) auto qed lemma simple_function_indicator[intro, simp]: assumes "A \ sets M" shows "simple_function M (indicator A)" proof - have "indicator A ` space M \ {0, 1}" (is "?S \ _") by (auto simp: indicator_def) hence "finite ?S" by (rule finite_subset) simp moreover have "- A \ space M = space M - A" by auto ultimately show ?thesis unfolding simple_function_def using assms by (auto simp: indicator_def [abs_def]) qed lemma simple_function_Pair[intro, simp]: assumes "simple_function M f" assumes "simple_function M g" shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") unfolding simple_function_def proof safe show "finite (?p ` space M)" using assms unfolding simple_function_def by (rule_tac finite_subset[of _ "f`space M \ g`space M"]) auto next fix x assume "x \ space M" have "(\x. (f x, g x)) -` {(f x, g x)} \ space M = (f -` {f x} \ space M) \ (g -` {g x} \ space M)" by auto with \x \ space M\ show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" using assms unfolding simple_function_def by auto qed lemma simple_function_compose1: assumes "simple_function M f" shows "simple_function M (\x. g (f x))" using simple_function_compose[OF assms, of g] by (simp add: comp_def) lemma simple_function_compose2: assumes "simple_function M f" and "simple_function M g" shows "simple_function M (\x. h (f x) (g x))" proof - have "simple_function M ((\(x, y). h x y) \ (\x. (f x, g x)))" using assms by auto thus ?thesis by (simp_all add: comp_def) qed lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="(+)"] and simple_function_diff[intro, simp] = simple_function_compose2[where h="(-)"] and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] and simple_function_mult[intro, simp] = simple_function_compose2[where h="(*)"] and simple_function_div[intro, simp] = simple_function_compose2[where h="(/)"] and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] and simple_function_max[intro, simp] = simple_function_compose2[where h=max] lemma simple_function_sum[intro, simp]: assumes "\i. i \ P \ simple_function M (f i)" shows "simple_function M (\x. \i\P. f i x)" proof cases assume "finite P" from this assms show ?thesis by induct auto qed auto lemma simple_function_ennreal[intro, simp]: fixes f g :: "'a \ real" assumes sf: "simple_function M f" shows "simple_function M (\x. ennreal (f x))" by (rule simple_function_compose1[OF sf]) lemma simple_function_real_of_nat[intro, simp]: fixes f g :: "'a \ nat" assumes sf: "simple_function M f" shows "simple_function M (\x. real (f x))" by (rule simple_function_compose1[OF sf]) lemma\<^marker>\tag important\ borel_measurable_implies_simple_function_sequence: fixes u :: "'a \ ennreal" assumes u[measurable]: "u \ borel_measurable M" shows "\f. incseq f \ (\i. (\x. f i x < top) \ simple_function M (f i)) \ u = (SUP i. f i)" proof - define f where [abs_def]: "f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x have [simp]: "0 \ f i x" for i x by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg) have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x by simp have "real_of_int \real i * 2 ^ i\ = real_of_int \i * 2 ^ i\" for i by (intro arg_cong[where f=real_of_int]) simp then have [simp]: "real_of_int \real i * 2 ^ i\ = i * 2 ^ i" for i unfolding floor_of_nat by simp have "incseq f" proof (intro monoI le_funI) fix m n :: nat and x assume "m \ n" moreover { fix d :: nat have "\2^d::real\ * \2^m * enn2real (min (of_nat m) (u x))\ \ \2^d * (2^m * enn2real (min (of_nat m) (u x)))\" by (rule le_mult_floor) (auto) also have "\ \ \2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\" by (intro floor_mono mult_mono enn2real_mono min.mono) (auto simp: min_less_iff_disj of_nat_less_top) finally have "f m x \ f (m + d) x" unfolding f_def by (auto simp: field_simps power_add * simp del: of_int_mult) } ultimately show "f m x \ f n x" by (auto simp add: le_iff_add) qed then have inc_f: "incseq (\i. ennreal (f i x))" for x by (auto simp: incseq_def le_fun_def) then have "incseq (\i x. ennreal (f i x))" by (auto simp: incseq_def le_fun_def) moreover have "simple_function M (f i)" for i proof (rule simple_function_borel_measurable) have "\enn2real (min (of_nat i) (u x)) * 2 ^ i\ \ \int i * 2 ^ i\" for x by (cases "u x" rule: ennreal_cases) (auto split: split_min intro!: floor_mono) then have "f i ` space M \ (\n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}" unfolding floor_of_int by (auto simp: f_def intro!: imageI) then show "finite (f i ` space M)" by (rule finite_subset) auto show "f i \ borel_measurable M" unfolding f_def enn2real_def by measurable qed moreover { fix x have "(SUP i. ennreal (f i x)) = u x" proof (cases "u x" rule: ennreal_cases) case top then show ?thesis by (simp add: f_def inf_min[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric] ennreal_SUP_of_nat_eq_top) next case (real r) obtain n where "r \ of_nat n" using real_arch_simple by auto then have min_eq_r: "\\<^sub>F x in sequentially. min (real x) r = r" by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min) have "(\i. real_of_int \min (real i) r * 2^i\ / 2^i) \ r" proof (rule tendsto_sandwich) show "(\n. r - (1/2)^n) \ r" by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero) show "\\<^sub>F n in sequentially. real_of_int \min (real n) r * 2 ^ n\ / 2 ^ n \ r" using min_eq_r by eventually_elim (auto simp: field_simps) have *: "r * (2 ^ n * 2 ^ n) \ 2^n + 2^n * real_of_int \r * 2 ^ n\" for n using real_of_int_floor_ge_diff_one[of "r * 2^n", THEN mult_left_mono, of "2^n"] by (auto simp: field_simps) show "\\<^sub>F n in sequentially. r - (1/2)^n \ real_of_int \min (real n) r * 2 ^ n\ / 2 ^ n" using min_eq_r by eventually_elim (insert *, auto simp: field_simps) qed auto then have "(\i. ennreal (f i x)) \ ennreal r" by (simp add: real f_def ennreal_of_nat_eq_real_of_nat min_ennreal) from LIMSEQ_unique[OF LIMSEQ_SUP[OF inc_f] this] show ?thesis by (simp add: real) qed } ultimately show ?thesis by (intro exI [of _ "\i x. ennreal (f i x)"]) (auto simp add: image_comp) qed lemma borel_measurable_implies_simple_function_sequence': fixes u :: "'a \ ennreal" assumes u: "u \ borel_measurable M" obtains f where "\i. simple_function M (f i)" "incseq f" "\i x. f i x < top" "\x. (SUP i. f i x) = u x" using borel_measurable_implies_simple_function_sequence [OF u] by (metis SUP_apply) lemma\<^marker>\tag important\ simple_function_induct [consumes 1, case_names cong set mult add, induct set: simple_function]: fixes u :: "'a \ ennreal" assumes u: "simple_function M u" assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. P u \ P (\x. c * u x)" assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" shows "P u" proof (rule cong) from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" proof eventually_elim fix x assume x: "x \ space M" from simple_function_indicator_representation[OF u x] show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. qed next from u have "finite (u ` space M)" unfolding simple_function_def by auto then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" proof induct case empty show ?case using set[of "{}"] by (simp add: indicator_def[abs_def]) qed (auto intro!: add mult set simple_functionD u) next show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" apply (subst simple_function_cong) apply (rule simple_function_indicator_representation[symmetric]) apply (auto intro: u) done qed fact lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]: fixes u :: "'a \ ennreal" assumes u: "simple_function M u" assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. simple_function M u \ P u \ P (\x. c * u x)" assumes add: "\u v. simple_function M u \ P u \ simple_function M v \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" shows "P u" proof - show ?thesis proof (rule cong) fix x assume x: "x \ space M" from simple_function_indicator_representation[OF u x] show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. next show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" apply (subst simple_function_cong) apply (rule simple_function_indicator_representation[symmetric]) apply (auto intro: u) done next from u have "finite (u ` space M)" unfolding simple_function_def by auto then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" proof induct case empty show ?case using set[of "{}"] by (simp add: indicator_def[abs_def]) next case (insert x S) { fix z have "(\y\S. y * indicator (u -` {y} \ space M) z) = 0 \ x * indicator (u -` {x} \ space M) z = 0" using insert by (subst sum_eq_0_iff) (auto simp: indicator_def) } note disj = this from insert show ?case by (auto intro!: add mult set simple_functionD u simple_function_sum disj) qed qed fact qed lemma\<^marker>\tag important\ borel_measurable_induct [consumes 1, case_names cong set mult add seq, induct set: borel_measurable]: fixes u :: "'a \ ennreal" assumes u: "u \ borel_measurable M" assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult': "\u c. c < top \ u \ borel_measurable M \ (\x. x \ space M \ u x < top) \ P u \ P (\x. c * u x)" assumes add: "\u v. u \ borel_measurable M\ (\x. x \ space M \ u x < top) \ P u \ v \ borel_measurable M \ (\x. x \ space M \ v x < top) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. x \ space M \ U i x < top) \ (\i. P (U i)) \ incseq U \ u = (SUP i. U i) \ P (SUP i. U i)" shows "P u" using u proof (induct rule: borel_measurable_implies_simple_function_sequence') fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i x. U i x < top" and sup: "\x. (SUP i. U i x) = u x" have u_eq: "u = (SUP i. U i)" using u by (auto simp add: image_comp sup) have not_inf: "\x i. x \ space M \ U i x < top" using U by (auto simp: image_iff eq_commute) from U have "\i. U i \ borel_measurable M" by (simp add: borel_measurable_simple_function) show "P u" unfolding u_eq proof (rule seq) fix i show "P (U i)" using \simple_function M (U i)\ not_inf[of _ i] proof (induct rule: simple_function_induct_nn) case (mult u c) show ?case proof cases assume "c = 0 \ space M = {} \ (\x\space M. u x = 0)" with mult(1) show ?thesis by (intro cong[of "\x. c * u x" "indicator {}"] set) (auto dest!: borel_measurable_simple_function) next assume "\ (c = 0 \ space M = {} \ (\x\space M. u x = 0))" then obtain x where "space M \ {}" and x: "x \ space M" "u x \ 0" "c \ 0" by auto with mult(3)[of x] have "c < top" by (auto simp: ennreal_mult_less_top) then have u_fin: "x' \ space M \ u x' < top" for x' using mult(3)[of x'] \c \ 0\ by (auto simp: ennreal_mult_less_top) then have "P u" by (rule mult) with u_fin \c < top\ mult(1) show ?thesis by (intro mult') (auto dest!: borel_measurable_simple_function) qed qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function) qed fact+ qed lemma simple_function_If_set: assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ sets M" shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") proof - define F where "F x = f -` {x} \ space M" for x define G where "G x = g -` {x} \ space M" for x show ?thesis unfolding simple_function_def proof safe have "?IF ` space M \ f ` space M \ g ` space M" by auto from finite_subset[OF this] assms show "finite (?IF ` space M)" unfolding simple_function_def by auto next fix x assume "x \ space M" then have *: "?IF -` {?IF x} \ space M = (if x \ A then ((F (f x) \ (A \ space M)) \ (G (f x) - (G (f x) \ (A \ space M)))) else ((F (g x) \ (A \ space M)) \ (G (g x) - (G (g x) \ (A \ space M)))))" using sets.sets_into_space[OF A] by (auto split: if_split_asm simp: G_def F_def) have [intro]: "\x. F x \ sets M" "\x. G x \ sets M" unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto show "?IF -` {?IF x} \ space M \ sets M" unfolding * using A by auto qed qed lemma simple_function_If: assumes sf: "simple_function M f" "simple_function M g" and P: "{x\space M. P x} \ sets M" shows "simple_function M (\x. if P x then f x else g x)" proof - have "{x\space M. P x} = {x. P x} \ space M" by auto with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp qed lemma simple_function_subalgebra: assumes "simple_function N f" and N_subalgebra: "sets N \ sets M" "space N = space M" shows "simple_function M f" using assms unfolding simple_function_def by auto lemma simple_function_comp: assumes T: "T \ measurable M M'" and f: "simple_function M' f" shows "simple_function M (\x. f (T x))" proof (intro simple_function_def[THEN iffD2] conjI ballI) have "(\x. f (T x)) ` space M \ f ` space M'" using T unfolding measurable_def by auto then show "finite ((\x. f (T x)) ` space M)" using f unfolding simple_function_def by (auto intro: finite_subset) fix i assume i: "i \ (\x. f (T x)) ` space M" then have "i \ f ` space M'" using T unfolding measurable_def by auto then have "f -` {i} \ space M' \ sets M'" using f unfolding simple_function_def by auto then have "T -` (f -` {i} \ space M') \ space M \ sets M" using T unfolding measurable_def by auto also have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" using T unfolding measurable_def by auto finally show "(\x. f (T x)) -` {i} \ space M \ sets M" . qed subsection "Simple integral" definition\<^marker>\tag important\ simple_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>S") where "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" syntax "_simple_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" ("\\<^sup>S _. _ \_" [60,61] 110) translations "\\<^sup>S x. f \M" == "CONST simple_integral M (%x. f)" lemma simple_integral_cong: assumes "\t. t \ space M \ f t = g t" shows "integral\<^sup>S M f = integral\<^sup>S M g" proof - have "f ` space M = g ` space M" "\x. f -` {x} \ space M = g -` {x} \ space M" using assms by (auto intro!: image_eqI) thus ?thesis unfolding simple_integral_def by simp qed lemma simple_integral_const[simp]: "(\\<^sup>Sx. c \M) = c * (emeasure M) (space M)" proof (cases "space M = {}") case True thus ?thesis unfolding simple_integral_def by simp next case False hence "(\x. c) ` space M = {c}" by auto thus ?thesis unfolding simple_integral_def by simp qed lemma simple_function_partition: assumes f: "simple_function M f" and g: "simple_function M g" assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" assumes v: "\x. x \ space M \ f x = v (g x)" shows "integral\<^sup>S M f = (\y\g ` space M. v y * emeasure M {x\space M. g x = y})" (is "_ = ?r") proof - from f g have [simp]: "finite (f`space M)" "finite (g`space M)" by (auto simp: simple_function_def) from f g have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" by (auto intro: measurable_simple_function) { fix y assume "y \ space M" then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" by (auto cong: sub simp: v[symmetric]) } note eq = this have "integral\<^sup>S M f = (\y\f`space M. y * (\z\g`space M. if \x\space M. y = f x \ z = g x then emeasure M {x\space M. g x = z} else 0))" unfolding simple_integral_def proof (safe intro!: sum.cong ennreal_mult_left_cong) fix y assume y: "y \ space M" "f y \ 0" have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = {z. \x\space M. f y = f x \ z = g x}" by auto have eq:"(\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i}) = f -` {f y} \ space M" by (auto simp: eq_commute cong: sub rev_conj_cong) have "finite (g`space M)" by simp then have "finite {z. \x\space M. f y = f x \ z = g x}" by (rule rev_finite_subset) auto then show "emeasure M (f -` {f y} \ space M) = (\z\g ` space M. if \x\space M. f y = f x \ z = g x then emeasure M {x \ space M. g x = z} else 0)" apply (simp add: sum.If_cases) apply (subst sum_emeasure) apply (auto simp: disjoint_family_on_def eq) done qed also have "\ = (\y\f`space M. (\z\g`space M. if \x\space M. y = f x \ z = g x then y * emeasure M {x\space M. g x = z} else 0))" by (auto intro!: sum.cong simp: sum_distrib_left) also have "\ = ?r" by (subst sum.swap) (auto intro!: sum.cong simp: sum.If_cases scaleR_sum_right[symmetric] eq) finally show "integral\<^sup>S M f = ?r" . qed lemma simple_integral_add[simp]: assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" shows "(\\<^sup>Sx. f x + g x \M) = integral\<^sup>S M f + integral\<^sup>S M g" proof - have "(\\<^sup>Sx. f x + g x \M) = (\y\(\x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\space M. (f x, g x) = y})" by (intro simple_function_partition) (auto intro: f g) also have "\ = (\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) + (\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y})" using assms(2,4) by (auto intro!: sum.cong distrib_right simp: sum.distrib[symmetric]) also have "(\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. f x \M)" by (intro simple_function_partition[symmetric]) (auto intro: f g) also have "(\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. g x \M)" by (intro simple_function_partition[symmetric]) (auto intro: f g) finally show ?thesis . qed lemma simple_integral_sum[simp]: assumes "\i x. i \ P \ 0 \ f i x" assumes "\i. i \ P \ simple_function M (f i)" shows "(\\<^sup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^sup>S M (f i))" proof cases assume "finite P" from this assms show ?thesis by induct (auto) qed auto lemma simple_integral_mult[simp]: assumes f: "simple_function M f" shows "(\\<^sup>Sx. c * f x \M) = c * integral\<^sup>S M f" proof - have "(\\<^sup>Sx. c * f x \M) = (\y\f ` space M. (c * y) * emeasure M {x\space M. f x = y})" using f by (intro simple_function_partition) auto also have "\ = c * integral\<^sup>S M f" using f unfolding simple_integral_def by (subst sum_distrib_left) (auto simp: mult.assoc Int_def conj_commute) finally show ?thesis . qed lemma simple_integral_mono_AE: assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" and mono: "AE x in M. f x \ g x" shows "integral\<^sup>S M f \ integral\<^sup>S M g" proof - let ?\ = "\P. emeasure M {x\space M. P x}" have "integral\<^sup>S M f = (\y\(\x. (f x, g x))`space M. fst y * ?\ (\x. (f x, g x) = y))" using f g by (intro simple_function_partition) auto also have "\ \ (\y\(\x. (f x, g x))`space M. snd y * ?\ (\x. (f x, g x) = y))" proof (clarsimp intro!: sum_mono) fix x assume "x \ space M" let ?M = "?\ (\y. f y = f x \ g y = g x)" show "f x * ?M \ g x * ?M" proof cases assume "?M \ 0" then have "0 < ?M" by (simp add: less_le) also have "\ \ ?\ (\y. f x \ g x)" using mono by (intro emeasure_mono_AE) auto finally have "\ \ f x \ g x" by (intro notI) auto then show ?thesis by (intro mult_right_mono) auto qed simp qed also have "\ = integral\<^sup>S M g" using f g by (intro simple_function_partition[symmetric]) auto finally show ?thesis . qed lemma simple_integral_mono: assumes "simple_function M f" and "simple_function M g" and mono: "\ x. x \ space M \ f x \ g x" shows "integral\<^sup>S M f \ integral\<^sup>S M g" using assms by (intro simple_integral_mono_AE) auto lemma simple_integral_cong_AE: assumes "simple_function M f" and "simple_function M g" and "AE x in M. f x = g x" shows "integral\<^sup>S M f = integral\<^sup>S M g" using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) lemma simple_integral_cong': assumes sf: "simple_function M f" "simple_function M g" and mea: "(emeasure M) {x\space M. f x \ g x} = 0" shows "integral\<^sup>S M f = integral\<^sup>S M g" proof (intro simple_integral_cong_AE sf AE_I) show "(emeasure M) {x\space M. f x \ g x} = 0" by fact show "{x \ space M. f x \ g x} \ sets M" using sf[THEN borel_measurable_simple_function] by auto qed simp lemma simple_integral_indicator: assumes A: "A \ sets M" assumes f: "simple_function M f" shows "(\\<^sup>Sx. f x * indicator A x \M) = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" proof - have eq: "(\x. (f x, indicator A x)) ` space M \ {x. snd x = 1} = (\x. (f x, 1::ennreal))`A" using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: if_split_asm) have eq2: "\x. f x \ f ` A \ f -` {f x} \ space M \ A = {}" by (auto simp: image_iff) have "(\\<^sup>Sx. f x * indicator A x \M) = (\y\(\x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\space M. (f x, indicator A x) = y})" using assms by (intro simple_function_partition) auto also have "\ = (\y\(\x. (f x, indicator A x::ennreal))`space M. if snd y = 1 then fst y * emeasure M (f -` {fst y} \ space M \ A) else 0)" by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="(*)"] arg_cong2[where f=emeasure] sum.cong) also have "\ = (\y\(\x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \ space M \ A))" using assms by (subst sum.If_cases) (auto intro!: simple_functionD(1) simp: eq) also have "\ = (\y\fst`(\x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \ space M \ A))" by (subst sum.reindex [of fst]) (auto simp: inj_on_def) also have "\ = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" using A[THEN sets.sets_into_space] by (intro sum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) finally show ?thesis . qed lemma simple_integral_indicator_only[simp]: assumes "A \ sets M" shows "integral\<^sup>S M (indicator A) = emeasure M A" using simple_integral_indicator[OF assms, of "\x. 1"] sets.sets_into_space[OF assms] by (simp_all add: image_constant_conv Int_absorb1 split: if_split_asm) lemma simple_integral_null_set: assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets M" shows "(\\<^sup>Sx. u x * indicator N x \M) = 0" proof - have "AE x in M. indicator N x = (0 :: ennreal)" using \N \ null_sets M\ by (auto simp: indicator_def intro!: AE_I[of _ _ N]) then have "(\\<^sup>Sx. u x * indicator N x \M) = (\\<^sup>Sx. 0 \M)" using assms apply (intro simple_integral_cong_AE) by auto then show ?thesis by simp qed lemma simple_integral_cong_AE_mult_indicator: assumes sf: "simple_function M f" and eq: "AE x in M. x \ S" and "S \ sets M" shows "integral\<^sup>S M f = (\\<^sup>Sx. f x * indicator S x \M)" using assms by (intro simple_integral_cong_AE) auto lemma simple_integral_cmult_indicator: assumes A: "A \ sets M" shows "(\\<^sup>Sx. c * indicator A x \M) = c * emeasure M A" using simple_integral_mult[OF simple_function_indicator[OF A]] unfolding simple_integral_indicator_only[OF A] by simp lemma simple_integral_nonneg: assumes f: "simple_function M f" and ae: "AE x in M. 0 \ f x" shows "0 \ integral\<^sup>S M f" proof - have "integral\<^sup>S M (\x. 0) \ integral\<^sup>S M f" using simple_integral_mono_AE[OF _ f ae] by auto then show ?thesis by simp qed subsection \Integral on nonnegative functions\ definition\<^marker>\tag important\ nn_integral :: "'a measure \ ('a \ ennreal) \ ennreal" ("integral\<^sup>N") where "integral\<^sup>N M f = (SUP g \ {g. simple_function M g \ g \ f}. integral\<^sup>S M g)" syntax "_nn_integral" :: "pttrn \ ennreal \ 'a measure \ ennreal" ("\\<^sup>+((2 _./ _)/ \_)" [60,61] 110) translations "\\<^sup>+x. f \M" == "CONST nn_integral M (\x. f)" lemma nn_integral_def_finite: "integral\<^sup>N M f = (SUP g \ {g. simple_function M g \ g \ f \ (\x. g x < top)}. integral\<^sup>S M g)" (is "_ = Sup (?A ` ?f)") unfolding nn_integral_def proof (safe intro!: antisym SUP_least) fix g assume g[measurable]: "simple_function M g" "g \ f" show "integral\<^sup>S M g \ Sup (?A ` ?f)" proof cases assume ae: "AE x in M. g x \ top" let ?G = "{x \ space M. g x \ top}" have "integral\<^sup>S M g = integral\<^sup>S M (\x. g x * indicator ?G x)" proof (rule simple_integral_cong_AE) show "AE x in M. g x = g x * indicator ?G x" using ae AE_space by eventually_elim auto qed (insert g, auto) also have "\ \ Sup (?A ` ?f)" using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator) finally show ?thesis . next assume nAE: "\ (AE x in M. g x \ top)" then have "emeasure M {x\space M. g x = top} \ 0" (is "emeasure M ?G \ 0") by (subst (asm) AE_iff_measurable[OF _ refl]) auto then have "top = (SUP n. (\\<^sup>Sx. of_nat n * indicator ?G x \M))" by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric]) also have "\ \ Sup (?A ` ?f)" using g by (safe intro!: SUP_least SUP_upper) (auto simp: le_fun_def of_nat_less_top top_unique[symmetric] split: split_indicator intro: order_trans[of _ "g x" "f x" for x, OF order_trans[of _ top]]) finally show ?thesis by (simp add: top_unique del: SUP_eq_top_iff Sup_eq_top_iff) qed qed (auto intro: SUP_upper) lemma nn_integral_mono_AE: assumes ae: "AE x in M. u x \ v x" shows "integral\<^sup>N M u \ integral\<^sup>N M v" unfolding nn_integral_def proof (safe intro!: SUP_mono) fix n assume n: "simple_function M n" "n \ u" - from ae[THEN AE_E] guess N . note N = this + from ae[THEN AE_E] obtain N + where N: "{x \ space M. \ u x \ v x} \ N" "emeasure M N = 0" "N \ sets M" + by auto then have ae_N: "AE x in M. x \ N" by (auto intro: AE_not_in) let ?n = "\x. n x * indicator (space M - N) x" have "AE x in M. n x \ ?n x" "simple_function M ?n" using n N ae_N by auto moreover { fix x have "?n x \ v x" proof cases assume x: "x \ space M - N" with N have "u x \ v x" by auto with n(2)[THEN le_funD, of x] x show ?thesis by (auto simp: max_def split: if_split_asm) qed simp } then have "?n \ v" by (auto simp: le_funI) moreover have "integral\<^sup>S M n \ integral\<^sup>S M ?n" using ae_N N n by (auto intro!: simple_integral_mono_AE) ultimately show "\m\{g. simple_function M g \ g \ v}. integral\<^sup>S M n \ integral\<^sup>S M m" by force qed lemma nn_integral_mono: "(\x. x \ space M \ u x \ v x) \ integral\<^sup>N M u \ integral\<^sup>N M v" by (auto intro: nn_integral_mono_AE) lemma mono_nn_integral: "mono F \ mono (\x. integral\<^sup>N M (F x))" by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono) lemma nn_integral_cong_AE: "AE x in M. u x = v x \ integral\<^sup>N M u = integral\<^sup>N M v" by (auto simp: eq_iff intro!: nn_integral_mono_AE) lemma nn_integral_cong: "(\x. x \ space M \ u x = v x) \ integral\<^sup>N M u = integral\<^sup>N M v" by (auto intro: nn_integral_cong_AE) lemma nn_integral_cong_simp: "(\x. x \ space M =simp=> u x = v x) \ integral\<^sup>N M u = integral\<^sup>N M v" by (auto intro: nn_integral_cong simp: simp_implies_def) lemma incseq_nn_integral: assumes "incseq f" shows "incseq (\i. integral\<^sup>N M (f i))" proof - have "\i x. f i x \ f (Suc i) x" using assms by (auto dest!: incseq_SucD simp: le_fun_def) then show ?thesis by (auto intro!: incseq_SucI nn_integral_mono) qed lemma nn_integral_eq_simple_integral: assumes f: "simple_function M f" shows "integral\<^sup>N M f = integral\<^sup>S M f" proof - let ?f = "\x. f x * indicator (space M) x" have f': "simple_function M ?f" using f by auto have "integral\<^sup>N M ?f \ integral\<^sup>S M ?f" using f' by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def) moreover have "integral\<^sup>S M ?f \ integral\<^sup>N M ?f" unfolding nn_integral_def using f' by (auto intro!: SUP_upper) ultimately show ?thesis by (simp cong: nn_integral_cong simple_integral_cong) qed text \Beppo-Levi monotone convergence theorem\ lemma nn_integral_monotone_convergence_SUP: assumes f: "incseq f" and [measurable]: "\i. f i \ borel_measurable M" shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" proof (rule antisym) show "(\\<^sup>+ x. (SUP i. f i x) \M) \ (SUP i. (\\<^sup>+ x. f i x \M))" unfolding nn_integral_def_finite[of _ "\x. SUP i. f i x"] proof (safe intro!: SUP_least) fix u assume sf_u[simp]: "simple_function M u" and u: "u \ (\x. SUP i. f i x)" and u_range: "\x. u x < top" note sf_u[THEN borel_measurable_simple_function, measurable] show "integral\<^sup>S M u \ (SUP j. \\<^sup>+x. f j x \M)" proof (rule ennreal_approx_unit) fix a :: ennreal assume "a < 1" let ?au = "\x. a * u x" let ?B = "\c i. {x\space M. ?au x = c \ c \ f i x}" have "integral\<^sup>S M ?au = (\c\?au`space M. c * (SUP i. emeasure M (?B c i)))" unfolding simple_integral_def proof (intro sum.cong ennreal_mult_left_cong refl) fix c assume "c \ ?au ` space M" "c \ 0" { fix x' assume x': "x' \ space M" "?au x' = c" with \c \ 0\ u_range have "?au x' < 1 * u x'" by (intro ennreal_mult_strict_right_mono \a < 1\) (auto simp: less_le) also have "\ \ (SUP i. f i x')" using u by (auto simp: le_fun_def) finally have "\i. ?au x' \ f i x'" by (auto simp: less_SUP_iff intro: less_imp_le) } then have *: "?au -` {c} \ space M = (\i. ?B c i)" by auto show "emeasure M (?au -` {c} \ space M) = (SUP i. emeasure M (?B c i))" unfolding * using f by (intro SUP_emeasure_incseq[symmetric]) (auto simp: incseq_def le_fun_def intro: order_trans) qed also have "\ = (SUP i. \c\?au`space M. c * emeasure M (?B c i))" unfolding SUP_mult_left_ennreal using f by (intro ennreal_SUP_sum[symmetric]) (auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans) also have "\ \ (SUP i. integral\<^sup>N M (f i))" proof (intro SUP_subset_mono order_refl) fix i have "(\c\?au`space M. c * emeasure M (?B c i)) = (\\<^sup>Sx. (a * u x) * indicator {x\space M. a * u x \ f i x} x \M)" by (subst simple_integral_indicator) (auto intro!: sum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure]) also have "\ = (\\<^sup>+x. (a * u x) * indicator {x\space M. a * u x \ f i x} x \M)" by (rule nn_integral_eq_simple_integral[symmetric]) simp also have "\ \ (\\<^sup>+x. f i x \M)" by (intro nn_integral_mono) (auto split: split_indicator) finally show "(\c\?au`space M. c * emeasure M (?B c i)) \ (\\<^sup>+x. f i x \M)" . qed finally show "a * integral\<^sup>S M u \ (SUP i. integral\<^sup>N M (f i))" by simp qed qed qed (auto intro!: SUP_least SUP_upper nn_integral_mono) lemma sup_continuous_nn_integral[order_continuous_intros]: assumes f: "\y. sup_continuous (f y)" assumes [measurable]: "\x. (\y. f y x) \ borel_measurable M" shows "sup_continuous (\x. (\\<^sup>+y. f y x \M))" unfolding sup_continuous_def proof safe fix C :: "nat \ 'b" assume C: "incseq C" with sup_continuous_mono[OF f] show "(\\<^sup>+ y. f y (Sup (C ` UNIV)) \M) = (SUP i. \\<^sup>+ y. f y (C i) \M)" unfolding sup_continuousD[OF f C] by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def) qed theorem nn_integral_monotone_convergence_SUP_AE: assumes f: "\i. AE x in M. f i x \ f (Suc i) x" "\i. f i \ borel_measurable M" shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" proof - from f have "AE x in M. \i. f i x \ f (Suc i) x" by (simp add: AE_all_countable) - from this[THEN AE_E] guess N . note N = this + from this[THEN AE_E] obtain N + where N: "{x \ space M. \ (\i. f i x \ f (Suc i) x)} \ N" "emeasure M N = 0" "N \ sets M" + by auto let ?f = "\i x. if x \ space M - N then f i x else 0" have f_eq: "AE x in M. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) then have "(\\<^sup>+ x. (SUP i. f i x) \M) = (\\<^sup>+ x. (SUP i. ?f i x) \M)" by (auto intro!: nn_integral_cong_AE) also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \M))" proof (rule nn_integral_monotone_convergence_SUP) show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) { fix i show "(\x. if x \ space M - N then f i x else 0) \ borel_measurable M" using f N(3) by (intro measurable_If_set) auto } qed also have "\ = (SUP i. (\\<^sup>+ x. f i x \M))" using f_eq by (force intro!: arg_cong[where f = "\f. Sup (range f)"] nn_integral_cong_AE ext) finally show ?thesis . qed lemma nn_integral_monotone_convergence_simple: "incseq f \ (\i. simple_function M (f i)) \ (SUP i. \\<^sup>S x. f i x \M) = (\\<^sup>+x. (SUP i. f i x) \M)" using nn_integral_monotone_convergence_SUP[of f M] by (simp add: nn_integral_eq_simple_integral[symmetric] borel_measurable_simple_function) lemma SUP_simple_integral_sequences: assumes f: "incseq f" "\i. simple_function M (f i)" and g: "incseq g" "\i. simple_function M (g i)" and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" (is "Sup (?F ` _) = Sup (?G ` _)") proof - have "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" using f by (rule nn_integral_monotone_convergence_simple) also have "\ = (\\<^sup>+x. (SUP i. g i x) \M)" unfolding eq[THEN nn_integral_cong_AE] .. also have "\ = (SUP i. ?G i)" using g by (rule nn_integral_monotone_convergence_simple[symmetric]) finally show ?thesis by simp qed lemma nn_integral_const[simp]: "(\\<^sup>+ x. c \M) = c * emeasure M (space M)" by (subst nn_integral_eq_simple_integral) auto lemma nn_integral_linear: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" shows "(\\<^sup>+ x. a * f x + g x \M) = a * integral\<^sup>N M f + integral\<^sup>N M g" (is "integral\<^sup>N M ?L = _") proof - - from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . + obtain u + where "\i. simple_function M (u i)" "incseq u" "\i x. u i x < top" "\x. (SUP i. u i x) = f x" + using borel_measurable_implies_simple_function_sequence' f(1) + by auto note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this - from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . + + obtain v where + "\i. simple_function M (v i)" "incseq v" "\i x. v i x < top" "\x. (SUP i. v i x) = g x" + using borel_measurable_implies_simple_function_sequence' g(1) + by auto note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this + let ?L' = "\i x. a * u i x + v i x" have "?L \ borel_measurable M" using assms by auto - from borel_measurable_implies_simple_function_sequence'[OF this] guess l . + from borel_measurable_implies_simple_function_sequence'[OF this] + obtain l where "\i. simple_function M (l i)" "incseq l" "\i x. l i x < top" "\x. (SUP i. l i x) = a * f x + g x" + by auto note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this have inc: "incseq (\i. a * integral\<^sup>S M (u i))" "incseq (\i. integral\<^sup>S M (v i))" using u v by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono mult_left_mono simple_integral_mono) have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" proof (rule SUP_simple_integral_sequences[OF l(3,2)]) show "incseq ?L'" "\i. simple_function M (?L' i)" using u v unfolding incseq_Suc_iff le_fun_def by (auto intro!: add_mono mult_left_mono) { fix x have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" using u(3) v(3) u(4)[of _ x] v(4)[of _ x] unfolding SUP_mult_left_ennreal by (auto intro!: ennreal_SUP_add simp: incseq_Suc_iff le_fun_def add_mono mult_left_mono) } then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" unfolding l(5) using u(5) v(5) by (intro AE_I2) auto qed also have "\ = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" using u(2) v(2) by auto finally show ?thesis unfolding l(5)[symmetric] l(1)[symmetric] by (simp add: ennreal_SUP_add[OF inc] v u SUP_mult_left_ennreal[symmetric]) qed lemma nn_integral_cmult: "f \ borel_measurable M \ (\\<^sup>+ x. c * f x \M) = c * integral\<^sup>N M f" using nn_integral_linear[of f M "\x. 0" c] by simp lemma nn_integral_multc: "f \ borel_measurable M \ (\\<^sup>+ x. f x * c \M) = integral\<^sup>N M f * c" unfolding mult.commute[of _ c] nn_integral_cmult by simp lemma nn_integral_divide: "f \ borel_measurable M \ (\\<^sup>+ x. f x / c \M) = (\\<^sup>+ x. f x \M) / c" unfolding divide_ennreal_def by (rule nn_integral_multc) lemma nn_integral_indicator[simp]: "A \ sets M \ (\\<^sup>+ x. indicator A x\M) = (emeasure M) A" by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator) lemma nn_integral_cmult_indicator: "A \ sets M \ (\\<^sup>+ x. c * indicator A x \M) = c * emeasure M A" by (subst nn_integral_eq_simple_integral) (auto) lemma nn_integral_indicator': assumes [measurable]: "A \ space M \ sets M" shows "(\\<^sup>+ x. indicator A x \M) = emeasure M (A \ space M)" proof - have "(\\<^sup>+ x. indicator A x \M) = (\\<^sup>+ x. indicator (A \ space M) x \M)" by (intro nn_integral_cong) (simp split: split_indicator) also have "\ = emeasure M (A \ space M)" by simp finally show ?thesis . qed lemma nn_integral_indicator_singleton[simp]: assumes [measurable]: "{y} \ sets M" shows "(\\<^sup>+x. f x * indicator {y} x \M) = f y * emeasure M {y}" proof - have "(\\<^sup>+x. f x * indicator {y} x \M) = (\\<^sup>+x. f y * indicator {y} x \M)" by (auto intro!: nn_integral_cong split: split_indicator) then show ?thesis by (simp add: nn_integral_cmult) qed lemma nn_integral_set_ennreal: "(\\<^sup>+x. ennreal (f x) * indicator A x \M) = (\\<^sup>+x. ennreal (f x * indicator A x) \M)" by (rule nn_integral_cong) (simp split: split_indicator) lemma nn_integral_indicator_singleton'[simp]: assumes [measurable]: "{y} \ sets M" shows "(\\<^sup>+x. ennreal (f x * indicator {y} x) \M) = f y * emeasure M {y}" by (subst nn_integral_set_ennreal[symmetric]) (simp) lemma nn_integral_add: "f \ borel_measurable M \ g \ borel_measurable M \ (\\<^sup>+ x. f x + g x \M) = integral\<^sup>N M f + integral\<^sup>N M g" using nn_integral_linear[of f M g 1] by simp lemma nn_integral_sum: "(\i. i \ P \ f i \ borel_measurable M) \ (\\<^sup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^sup>N M (f i))" by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add) theorem nn_integral_suminf: assumes f: "\i. f i \ borel_measurable M" shows "(\\<^sup>+ x. (\i. f i x) \M) = (\i. integral\<^sup>N M (f i))" proof - have all_pos: "AE x in M. \i. 0 \ f i x" using assms by (auto simp: AE_all_countable) have "(\i. integral\<^sup>N M (f i)) = (SUP n. \iN M (f i))" by (rule suminf_eq_SUP) also have "\ = (SUP n. \\<^sup>+x. (\iM)" unfolding nn_integral_sum[OF f] .. also have "\ = \\<^sup>+x. (SUP n. \iM" using f all_pos by (intro nn_integral_monotone_convergence_SUP_AE[symmetric]) (elim AE_mp, auto simp: sum_nonneg simp del: sum.lessThan_Suc intro!: AE_I2 sum_mono2) also have "\ = \\<^sup>+x. (\i. f i x) \M" using all_pos by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP) finally show ?thesis by simp qed lemma nn_integral_bound_simple_function: assumes bnd: "\x. x \ space M \ f x < \" assumes f[measurable]: "simple_function M f" assumes supp: "emeasure M {x\space M. f x \ 0} < \" shows "nn_integral M f < \" proof cases assume "space M = {}" then have "nn_integral M f = (\\<^sup>+x. 0 \M)" by (intro nn_integral_cong) auto then show ?thesis by simp next assume "space M \ {}" with simple_functionD(1)[OF f] bnd have bnd: "0 \ Max (f`space M) \ Max (f`space M) < \" by (subst Max_less_iff) (auto simp: Max_ge_iff) have "nn_integral M f \ (\\<^sup>+x. Max (f`space M) * indicator {x\space M. f x \ 0} x \M)" proof (rule nn_integral_mono) fix x assume "x \ space M" with f show "f x \ Max (f ` space M) * indicator {x \ space M. f x \ 0} x" by (auto split: split_indicator intro!: Max_ge simple_functionD) qed also have "\ < \" using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top) finally show ?thesis . qed theorem nn_integral_Markov_inequality: assumes u: "(\x. u x * indicator A x) \ borel_measurable M" and "A \ sets M" shows "(emeasure M) ({x\A. 1 \ c * u x}) \ c * (\\<^sup>+ x. u x * indicator A x \M)" (is "(emeasure M) ?A \ _ * ?PI") proof - define u' where "u' = (\x. u x * indicator A x)" have [measurable]: "u' \ borel_measurable M" using u unfolding u'_def . have "{x\space M. c * u' x \ 1} \ sets M" by measurable also have "{x\space M. c * u' x \ 1} = ?A" using sets.sets_into_space[OF \A \ sets M\] by (auto simp: u'_def indicator_def) finally have "(emeasure M) ?A = (\\<^sup>+ x. indicator ?A x \M)" using nn_integral_indicator by simp also have "\ \ (\\<^sup>+ x. c * (u x * indicator A x) \M)" using u by (auto intro!: nn_integral_mono_AE simp: indicator_def) also have "\ = c * (\\<^sup>+ x. u x * indicator A x \M)" using assms by (auto intro!: nn_integral_cmult) finally show ?thesis . qed lemma Chernoff_ineq_nn_integral_ge: assumes s: "s > 0" and [measurable]: "A \ sets M" assumes [measurable]: "(\x. f x * indicator A x) \ borel_measurable M" shows "emeasure M {x\A. f x \ a} \ ennreal (exp (-s * a)) * nn_integral M (\x. ennreal (exp (s * f x)) * indicator A x)" proof - define f' where "f' = (\x. f x * indicator A x)" have [measurable]: "f' \ borel_measurable M" using assms(3) unfolding f'_def by assumption have "(\x. ennreal (exp (s * f' x)) * indicator A x) \ borel_measurable M" by simp also have "(\x. ennreal (exp (s * f' x)) * indicator A x) = (\x. ennreal (exp (s * f x)) * indicator A x)" by (auto simp: f'_def indicator_def fun_eq_iff) finally have meas: "\ \ borel_measurable M" . have "{x\A. f x \ a} = {x\A. ennreal (exp (-s * a)) * ennreal (exp (s * f x)) \ 1}" using s by (auto simp: exp_minus field_simps simp flip: ennreal_mult) also have "emeasure M \ \ ennreal (exp (-s * a)) * (\\<^sup>+x. ennreal (exp (s * f x)) * indicator A x \M)" by (intro order.trans[OF nn_integral_Markov_inequality] meas) auto finally show ?thesis . qed lemma Chernoff_ineq_nn_integral_le: assumes s: "s > 0" and [measurable]: "A \ sets M" assumes [measurable]: "f \ borel_measurable M" shows "emeasure M {x\A. f x \ a} \ ennreal (exp (s * a)) * nn_integral M (\x. ennreal (exp (-s * f x)) * indicator A x)" using Chernoff_ineq_nn_integral_ge[of s A M "\x. -f x" "-a"] assms by simp lemma nn_integral_noteq_infinite: assumes g: "g \ borel_measurable M" and "integral\<^sup>N M g \ \" shows "AE x in M. g x \ \" proof (rule ccontr) assume c: "\ (AE x in M. g x \ \)" have "(emeasure M) {x\space M. g x = \} \ 0" using c g by (auto simp add: AE_iff_null) then have "0 < (emeasure M) {x\space M. g x = \}" by (auto simp: zero_less_iff_neq_zero) then have "\ = \ * (emeasure M) {x\space M. g x = \}" by (auto simp: ennreal_top_eq_mult_iff) also have "\ \ (\\<^sup>+x. \ * indicator {x\space M. g x = \} x \M)" using g by (subst nn_integral_cmult_indicator) auto also have "\ \ integral\<^sup>N M g" using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def) finally show False using \integral\<^sup>N M g \ \\ by (auto simp: top_unique) qed lemma nn_integral_PInf: assumes f: "f \ borel_measurable M" and not_Inf: "integral\<^sup>N M f \ \" shows "emeasure M (f -` {\} \ space M) = 0" proof - have "\ * emeasure M (f -` {\} \ space M) = (\\<^sup>+ x. \ * indicator (f -` {\} \ space M) x \M)" using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets) also have "\ \ integral\<^sup>N M f" by (auto intro!: nn_integral_mono simp: indicator_def) finally have "\ * (emeasure M) (f -` {\} \ space M) \ integral\<^sup>N M f" by simp then show ?thesis using assms by (auto simp: ennreal_top_mult top_unique split: if_split_asm) qed lemma simple_integral_PInf: "simple_function M f \ integral\<^sup>S M f \ \ \ emeasure M (f -` {\} \ space M) = 0" by (rule nn_integral_PInf) (auto simp: nn_integral_eq_simple_integral borel_measurable_simple_function) lemma nn_integral_PInf_AE: assumes "f \ borel_measurable M" "integral\<^sup>N M f \ \" shows "AE x in M. f x \ \" proof (rule AE_I) show "(emeasure M) (f -` {\} \ space M) = 0" by (rule nn_integral_PInf[OF assms]) show "f -` {\} \ space M \ sets M" using assms by (auto intro: borel_measurable_vimage) qed auto lemma nn_integral_diff: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" and fin: "integral\<^sup>N M g \ \" and mono: "AE x in M. g x \ f x" shows "(\\<^sup>+ x. f x - g x \M) = integral\<^sup>N M f - integral\<^sup>N M g" proof - have diff: "(\x. f x - g x) \ borel_measurable M" using assms by auto have "AE x in M. f x = f x - g x + g x" using diff_add_cancel_ennreal mono nn_integral_noteq_infinite[OF g fin] assms by auto then have **: "integral\<^sup>N M f = (\\<^sup>+x. f x - g x \M) + integral\<^sup>N M g" unfolding nn_integral_add[OF diff g, symmetric] by (rule nn_integral_cong_AE) show ?thesis unfolding ** using fin by (cases rule: ennreal2_cases[of "\\<^sup>+ x. f x - g x \M" "integral\<^sup>N M g"]) auto qed lemma nn_integral_mult_bounded_inf: assumes f: "f \ borel_measurable M" "(\\<^sup>+x. f x \M) < \" and c: "c \ \" and ae: "AE x in M. g x \ c * f x" shows "(\\<^sup>+x. g x \M) < \" proof - have "(\\<^sup>+x. g x \M) \ (\\<^sup>+x. c * f x \M)" by (intro nn_integral_mono_AE ae) also have "(\\<^sup>+x. c * f x \M) < \" using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less) finally show ?thesis . qed text \Fatou's lemma: convergence theorem on limes inferior\ lemma nn_integral_monotone_convergence_INF_AE': assumes f: "\i. AE x in M. f (Suc i) x \ f i x" and [measurable]: "\i. f i \ borel_measurable M" and *: "(\\<^sup>+ x. f 0 x \M) < \" shows "(\\<^sup>+ x. (INF i. f i x) \M) = (INF i. integral\<^sup>N M (f i))" proof (rule ennreal_minus_cancel) have "integral\<^sup>N M (f 0) - (\\<^sup>+ x. (INF i. f i x) \M) = (\\<^sup>+x. f 0 x - (INF i. f i x) \M)" proof (rule nn_integral_diff[symmetric]) have "(\\<^sup>+ x. (INF i. f i x) \M) \ (\\<^sup>+ x. f 0 x \M)" by (intro nn_integral_mono INF_lower) simp with * show "(\\<^sup>+ x. (INF i. f i x) \M) \ \" by simp qed (auto intro: INF_lower) also have "\ = (\\<^sup>+x. (SUP i. f 0 x - f i x) \M)" by (simp add: ennreal_INF_const_minus) also have "\ = (SUP i. (\\<^sup>+x. f 0 x - f i x \M))" proof (intro nn_integral_monotone_convergence_SUP_AE) show "AE x in M. f 0 x - f i x \ f 0 x - f (Suc i) x" for i using f[of i] by eventually_elim (auto simp: ennreal_mono_minus) qed simp also have "\ = (SUP i. nn_integral M (f 0) - (\\<^sup>+x. f i x \M))" proof (subst nn_integral_diff[symmetric]) fix i have dec: "AE x in M. \i. f (Suc i) x \ f i x" unfolding AE_all_countable using f by auto then show "AE x in M. f i x \ f 0 x" using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "\i. f i x" 0 i for x]) then have "(\\<^sup>+ x. f i x \M) \ (\\<^sup>+ x. f 0 x \M)" by (rule nn_integral_mono_AE) with * show "(\\<^sup>+ x. f i x \M) \ \" by simp qed (insert f, auto simp: decseq_def le_fun_def) finally show "integral\<^sup>N M (f 0) - (\\<^sup>+ x. (INF i. f i x) \M) = integral\<^sup>N M (f 0) - (INF i. \\<^sup>+ x. f i x \M)" by (simp add: ennreal_INF_const_minus) qed (insert *, auto intro!: nn_integral_mono intro: INF_lower) theorem nn_integral_monotone_convergence_INF_AE: fixes f :: "nat \ 'a \ ennreal" assumes f: "\i. AE x in M. f (Suc i) x \ f i x" and [measurable]: "\i. f i \ borel_measurable M" and fin: "(\\<^sup>+ x. f i x \M) < \" shows "(\\<^sup>+ x. (INF i. f i x) \M) = (INF i. integral\<^sup>N M (f i))" proof - { fix f :: "nat \ ennreal" and j assume "decseq f" then have "(INF i. f i) = (INF i. f (i + j))" apply (intro INF_eq) apply (rule_tac x="i" in bexI) apply (auto simp: decseq_def le_fun_def) done } note INF_shift = this have mono: "AE x in M. \i. f (Suc i) x \ f i x" using f by (auto simp: AE_all_countable) then have "AE x in M. (INF i. f i x) = (INF n. f (n + i) x)" by eventually_elim (auto intro!: decseq_SucI INF_shift) then have "(\\<^sup>+ x. (INF i. f i x) \M) = (\\<^sup>+ x. (INF n. f (n + i) x) \M)" by (rule nn_integral_cong_AE) also have "\ = (INF n. (\\<^sup>+ x. f (n + i) x \M))" by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto) also have "\ = (INF n. (\\<^sup>+ x. f n x \M))" by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f) finally show ?thesis . qed lemma nn_integral_monotone_convergence_INF_decseq: assumes f: "decseq f" and *: "\i. f i \ borel_measurable M" "(\\<^sup>+ x. f i x \M) < \" shows "(\\<^sup>+ x. (INF i. f i x) \M) = (INF i. integral\<^sup>N M (f i))" using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (auto simp: decseq_Suc_iff le_fun_def) theorem nn_integral_liminf: fixes u :: "nat \ 'a \ ennreal" assumes u: "\i. u i \ borel_measurable M" shows "(\\<^sup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>N M (u n))" proof - have "(\\<^sup>+ x. liminf (\n. u n x) \M) = (SUP n. \\<^sup>+ x. (INF i\{n..}. u i x) \M)" unfolding liminf_SUP_INF using u by (intro nn_integral_monotone_convergence_SUP_AE) (auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) also have "\ \ liminf (\n. integral\<^sup>N M (u n))" by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower) finally show ?thesis . qed theorem nn_integral_limsup: fixes u :: "nat \ 'a \ ennreal" assumes [measurable]: "\i. u i \ borel_measurable M" "w \ borel_measurable M" assumes bounds: "\i. AE x in M. u i x \ w x" and w: "(\\<^sup>+x. w x \M) < \" shows "limsup (\n. integral\<^sup>N M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" proof - have bnd: "AE x in M. \i. u i x \ w x" using bounds by (auto simp: AE_all_countable) then have "(\\<^sup>+ x. (SUP n. u n x) \M) \ (\\<^sup>+ x. w x \M)" by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least) then have "(\\<^sup>+ x. limsup (\n. u n x) \M) = (INF n. \\<^sup>+ x. (SUP i\{n..}. u i x) \M)" unfolding limsup_INF_SUP using bnd w by (intro nn_integral_monotone_convergence_INF_AE') (auto intro!: AE_I2 intro: SUP_least SUP_subset_mono) also have "\ \ limsup (\n. integral\<^sup>N M (u n))" by (auto simp: limsup_INF_SUP intro!: INF_mono SUP_least exI nn_integral_mono SUP_upper) finally (xtrans) show ?thesis . qed lemma nn_integral_LIMSEQ: assumes f: "incseq f" "\i. f i \ borel_measurable M" and u: "\x. (\i. f i x) \ u x" shows "(\n. integral\<^sup>N M (f n)) \ integral\<^sup>N M u" proof - have "(\n. integral\<^sup>N M (f n)) \ (SUP n. integral\<^sup>N M (f n))" using f by (intro LIMSEQ_SUP[of "\n. integral\<^sup>N M (f n)"] incseq_nn_integral) also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\x. SUP n. f n x)" using f by (intro nn_integral_monotone_convergence_SUP[symmetric]) also have "integral\<^sup>N M (\x. SUP n. f n x) = integral\<^sup>N M (\x. u x)" using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def) finally show ?thesis . qed theorem nn_integral_dominated_convergence: assumes [measurable]: "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" and bound: "\j. AE x in M. u j x \ w x" and w: "(\\<^sup>+x. w x \M) < \" and u': "AE x in M. (\i. u i x) \ u' x" shows "(\i. (\\<^sup>+x. u i x \M)) \ (\\<^sup>+x. u' x \M)" proof - have "limsup (\n. integral\<^sup>N M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" by (intro nn_integral_limsup[OF _ _ bound w]) auto moreover have "(\\<^sup>+ x. limsup (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) moreover have "(\\<^sup>+ x. liminf (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) moreover have "(\\<^sup>+x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>N M (u n))" by (intro nn_integral_liminf) auto moreover have "liminf (\n. integral\<^sup>N M (u n)) \ limsup (\n. integral\<^sup>N M (u n))" by (intro Liminf_le_Limsup sequentially_bot) ultimately show ?thesis by (intro Liminf_eq_Limsup) auto qed lemma inf_continuous_nn_integral[order_continuous_intros]: assumes f: "\y. inf_continuous (f y)" assumes [measurable]: "\x. (\y. f y x) \ borel_measurable M" assumes bnd: "\x. (\\<^sup>+ y. f y x \M) \ \" shows "inf_continuous (\x. (\\<^sup>+y. f y x \M))" unfolding inf_continuous_def proof safe fix C :: "nat \ 'b" assume C: "decseq C" then show "(\\<^sup>+ y. f y (Inf (C ` UNIV)) \M) = (INF i. \\<^sup>+ y. f y (C i) \M)" using inf_continuous_mono[OF f] bnd by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def less_top intro!: nn_integral_monotone_convergence_INF_decseq) qed lemma nn_integral_null_set: assumes "N \ null_sets M" shows "(\\<^sup>+ x. u x * indicator N x \M) = 0" proof - have "(\\<^sup>+ x. u x * indicator N x \M) = (\\<^sup>+ x. 0 \M)" proof (intro nn_integral_cong_AE AE_I) show "{x \ space M. u x * indicator N x \ 0} \ N" by (auto simp: indicator_def) show "(emeasure M) N = 0" "N \ sets M" using assms by auto qed then show ?thesis by simp qed lemma nn_integral_0_iff: assumes u [measurable]: "u \ borel_measurable M" shows "integral\<^sup>N M u = 0 \ emeasure M {x\space M. u x \ 0} = 0" (is "_ \ (emeasure M) ?A = 0") proof - have u_eq: "(\\<^sup>+ x. u x * indicator ?A x \M) = integral\<^sup>N M u" by (auto intro!: nn_integral_cong simp: indicator_def) show ?thesis proof assume "(emeasure M) ?A = 0" with nn_integral_null_set[of ?A M u] u show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def) next assume *: "integral\<^sup>N M u = 0" let ?M = "\n. {x \ space M. 1 \ real (n::nat) * u x}" have "0 = (SUP n. (emeasure M) (?M n \ ?A))" proof - { fix n :: nat have "emeasure M {x \ ?A. 1 \ of_nat n * u x} \ of_nat n * \\<^sup>+ x. u x * indicator ?A x \M" by (intro nn_integral_Markov_inequality) auto also have "{x \ ?A. 1 \ of_nat n * u x} = (?M n \ ?A)" by (auto simp: ennreal_of_nat_eq_real_of_nat u_eq * ) finally have "emeasure M (?M n \ ?A) \ 0" by (simp add: ennreal_of_nat_eq_real_of_nat u_eq * ) moreover have "0 \ (emeasure M) (?M n \ ?A)" using u by auto ultimately have "(emeasure M) (?M n \ ?A) = 0" by auto } thus ?thesis by simp qed also have "\ = (emeasure M) (\n. ?M n \ ?A)" proof (safe intro!: SUP_emeasure_incseq) fix n show "?M n \ ?A \ sets M" using u by (auto intro!: sets.Int) next show "incseq (\n. {x \ space M. 1 \ real n * u x} \ {x \ space M. u x \ 0})" proof (safe intro!: incseq_SucI) fix n :: nat and x assume *: "1 \ real n * u x" also have "real n * u x \ real (Suc n) * u x" by (auto intro!: mult_right_mono) finally show "1 \ real (Suc n) * u x" by auto qed qed also have "\ = (emeasure M) {x\space M. 0 < u x}" proof (safe intro!: arg_cong[where f="(emeasure M)"]) fix x assume "0 < u x" and [simp, intro]: "x \ space M" show "x \ (\n. ?M n \ ?A)" proof (cases "u x" rule: ennreal_cases) case (real r) with \0 < u x\ have "0 < r" by auto obtain j :: nat where "1 / r \ real j" using real_arch_simple .. hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using \0 < r\ by auto hence "1 \ real j * r" using real \0 < r\ by auto thus ?thesis using \0 < r\ real by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric] simp del: ennreal_1) qed (insert \0 < u x\, auto simp: ennreal_mult_top) qed (auto simp: zero_less_iff_neq_zero) finally show "emeasure M ?A = 0" by (simp add: zero_less_iff_neq_zero) qed qed lemma nn_integral_0_iff_AE: assumes u: "u \ borel_measurable M" shows "integral\<^sup>N M u = 0 \ (AE x in M. u x = 0)" proof - have sets: "{x\space M. u x \ 0} \ sets M" using u by auto show "integral\<^sup>N M u = 0 \ (AE x in M. u x = 0)" using nn_integral_0_iff[of u] AE_iff_null[OF sets] u by auto qed lemma AE_iff_nn_integral: "{x\space M. P x} \ sets M \ (AE x in M. P x) \ integral\<^sup>N M (indicator {x. \ P x}) = 0" by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def]) lemma nn_integral_less: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" assumes f: "(\\<^sup>+x. f x \M) \ \" assumes ord: "AE x in M. f x \ g x" "\ (AE x in M. g x \ f x)" shows "(\\<^sup>+x. f x \M) < (\\<^sup>+x. g x \M)" proof - have "0 < (\\<^sup>+x. g x - f x \M)" proof (intro order_le_neq_trans notI) assume "0 = (\\<^sup>+x. g x - f x \M)" then have "AE x in M. g x - f x = 0" using nn_integral_0_iff_AE[of "\x. g x - f x" M] by simp with ord(1) have "AE x in M. g x \ f x" by eventually_elim (auto simp: ennreal_minus_eq_0) with ord show False by simp qed simp also have "\ = (\\<^sup>+x. g x \M) - (\\<^sup>+x. f x \M)" using f by (subst nn_integral_diff) (auto simp: ord) finally show ?thesis using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top) qed lemma nn_integral_subalgebra: assumes f: "f \ borel_measurable N" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integral\<^sup>N N f = integral\<^sup>N M f" proof - have [simp]: "\f :: 'a \ ennreal. f \ borel_measurable N \ f \ borel_measurable M" using N by (auto simp: measurable_def) have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq) have [simp]: "\A. A \ sets N \ A \ sets M" using N by auto from f show ?thesis apply induct apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N image_comp) apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric]) done qed lemma nn_integral_nat_function: fixes f :: "'a \ nat" assumes "f \ measurable M (count_space UNIV)" shows "(\\<^sup>+x. of_nat (f x) \M) = (\t. emeasure M {x\space M. t < f x})" proof - define F where "F i = {x\space M. i < f x}" for i with assms have [measurable]: "\i. F i \ sets M" by auto { fix x assume "x \ space M" have "(\i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" using sums_If_finite[of "\i. i < f x" "\_. 1::real"] by simp then have "(\i. ennreal (if i < f x then 1 else 0)) sums of_nat(f x)" unfolding ennreal_of_nat_eq_real_of_nat by (subst sums_ennreal) auto moreover have "\i. ennreal (if i < f x then 1 else 0) = indicator (F i) x" using \x \ space M\ by (simp add: one_ennreal_def F_def) ultimately have "of_nat (f x) = (\i. indicator (F i) x :: ennreal)" by (simp add: sums_iff) } then have "(\\<^sup>+x. of_nat (f x) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" by (simp cong: nn_integral_cong) also have "\ = (\i. emeasure M (F i))" by (simp add: nn_integral_suminf) finally show ?thesis by (simp add: F_def) qed theorem nn_integral_lfp: assumes sets[simp]: "\s. sets (M s) = sets N" assumes f: "sup_continuous f" assumes g: "sup_continuous g" assumes meas: "\F. F \ borel_measurable N \ f F \ borel_measurable N" assumes step: "\F s. F \ borel_measurable N \ integral\<^sup>N (M s) (f F) = g (\s. integral\<^sup>N (M s) F) s" shows "(\\<^sup>+\. lfp f \ \M s) = lfp g s" proof (subst lfp_transfer_bounded[where \="\F s. \\<^sup>+x. F x \M s" and g=g and f=f and P="\f. f \ borel_measurable N", symmetric]) fix C :: "nat \ 'b \ ennreal" assume "incseq C" "\i. C i \ borel_measurable N" then show "(\s. \\<^sup>+x. (SUP i. C i) x \M s) = (SUP i. (\s. \\<^sup>+x. C i x \M s))" unfolding SUP_apply[abs_def] by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g) theorem nn_integral_gfp: assumes sets[simp]: "\s. sets (M s) = sets N" assumes f: "inf_continuous f" and g: "inf_continuous g" assumes meas: "\F. F \ borel_measurable N \ f F \ borel_measurable N" assumes bound: "\F s. F \ borel_measurable N \ (\\<^sup>+x. f F x \M s) < \" assumes non_zero: "\s. emeasure (M s) (space (M s)) \ 0" assumes step: "\F s. F \ borel_measurable N \ integral\<^sup>N (M s) (f F) = g (\s. integral\<^sup>N (M s) F) s" shows "(\\<^sup>+\. gfp f \ \M s) = gfp g s" proof (subst gfp_transfer_bounded[where \="\F s. \\<^sup>+x. F x \M s" and g=g and f=f and P="\F. F \ borel_measurable N \ (\s. (\\<^sup>+x. F x \M s) < \)", symmetric]) fix C :: "nat \ 'b \ ennreal" assume "decseq C" "\i. C i \ borel_measurable N \ (\s. integral\<^sup>N (M s) (C i) < \)" then show "(\s. \\<^sup>+x. (INF i. C i) x \M s) = (INF i. (\s. \\<^sup>+x. C i x \M s))" unfolding INF_apply[abs_def] by (subst nn_integral_monotone_convergence_INF_decseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) next show "\x. g x \ (\s. integral\<^sup>N (M s) (f top))" by (subst step) (auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult cong del: if_weak_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD]) next fix C assume "\i::nat. C i \ borel_measurable N \ (\s. integral\<^sup>N (M s) (C i) < \)" "decseq C" with bound show "Inf (C ` UNIV) \ borel_measurable N \ (\s. integral\<^sup>N (M s) (Inf (C ` UNIV)) < \)" unfolding INF_apply[abs_def] by (subst nn_integral_monotone_convergence_INF_decseq) (auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF) next show "\x. x \ borel_measurable N \ (\s. integral\<^sup>N (M s) x < \) \ (\s. integral\<^sup>N (M s) (f x)) = g (\s. integral\<^sup>N (M s) x)" by (subst step) auto qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g) text \Cauchy--Schwarz inequality for \<^const>\nn_integral\\ lemma sum_of_squares_ge_ennreal: fixes a b :: ennreal shows "2 * a * b \ a\<^sup>2 + b\<^sup>2" proof (cases a; cases b) fix x y assume xy: "x \ 0" "y \ 0" and [simp]: "a = ennreal x" "b = ennreal y" have "0 \ (x - y)\<^sup>2" by simp also have "\ = x\<^sup>2 + y\<^sup>2 - 2 * x * y" by (simp add: algebra_simps power2_eq_square) finally have "2 * x * y \ x\<^sup>2 + y\<^sup>2" by simp hence "ennreal (2 * x * y) \ ennreal (x\<^sup>2 + y\<^sup>2)" by (intro ennreal_leI) thus ?thesis using xy by (simp add: ennreal_mult ennreal_power) qed auto lemma Cauchy_Schwarz_nn_integral: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" shows "(\\<^sup>+x. f x * g x \M)\<^sup>2 \ (\\<^sup>+x. f x ^ 2 \M) * (\\<^sup>+x. g x ^ 2 \M)" proof (cases "(\\<^sup>+x. f x * g x \M) = 0") case False define F where "F = nn_integral M (\x. f x ^ 2)" define G where "G = nn_integral M (\x. g x ^ 2)" from False have "\(AE x in M. f x = 0 \ g x = 0)" by (auto simp: nn_integral_0_iff_AE) hence "\(AE x in M. f x = 0)" and "\(AE x in M. g x = 0)" by (auto intro: AE_disjI1 AE_disjI2) hence nz: "F \ 0" "G \ 0" by (auto simp: nn_integral_0_iff_AE F_def G_def) show ?thesis proof (cases "F = \ \ G = \") case True thus ?thesis using nz by (auto simp: F_def G_def) next case False define F' where "F' = ennreal (sqrt (enn2real F))" define G' where "G' = ennreal (sqrt (enn2real G))" from False have fin: "F < top" "G < top" by (simp_all add: top.not_eq_extremum) have F'_sqr: "F'\<^sup>2 = F" using False by (cases F) (auto simp: F'_def ennreal_power) have G'_sqr: "G'\<^sup>2 = G" using False by (cases G) (auto simp: G'_def ennreal_power) have nz': "F' \ 0" "G' \ 0" and fin': "F' \ \" "G' \ \" using F'_sqr G'_sqr nz fin by auto from fin' have fin'': "F' < top" "G' < top" by (auto simp: top.not_eq_extremum) have "2 * (F' / F') * (G' / G') * (\\<^sup>+x. f x * g x \M) = F' * G' * (\\<^sup>+x. 2 * (f x / F') * (g x / G') \M)" using nz' fin'' by (simp add: divide_ennreal_def algebra_simps ennreal_inverse_mult flip: nn_integral_cmult) also have "F'/ F' = 1" using nz' fin'' by simp also have "G'/ G' = 1" using nz' fin'' by simp also have "2 * 1 * 1 = (2 :: ennreal)" by simp also have "F' * G' * (\\<^sup>+ x. 2 * (f x / F') * (g x / G') \M) \ F' * G' * (\\<^sup>+x. (f x / F')\<^sup>2 + (g x / G')\<^sup>2 \M)" by (intro mult_left_mono nn_integral_mono sum_of_squares_ge_ennreal) auto also have "\ = F' * G' * (F / F'\<^sup>2 + G / G'\<^sup>2)" using nz by (auto simp: nn_integral_add algebra_simps nn_integral_divide F_def G_def) also have "F / F'\<^sup>2 = 1" using nz F'_sqr fin by simp also have "G / G'\<^sup>2 = 1" using nz G'_sqr fin by simp also have "F' * G' * (1 + 1) = 2 * (F' * G')" by (simp add: mult_ac) finally have "(\\<^sup>+x. f x * g x \M) \ F' * G'" by (subst (asm) ennreal_mult_le_mult_iff) auto hence "(\\<^sup>+x. f x * g x \M)\<^sup>2 \ (F' * G')\<^sup>2" by (intro power_mono_ennreal) also have "\ = F * G" by (simp add: algebra_simps F'_sqr G'_sqr) finally show ?thesis by (simp add: F_def G_def) qed qed auto (* TODO: rename? *) subsection \Integral under concrete measures\ lemma nn_integral_mono_measure: assumes "sets M = sets N" "M \ N" shows "nn_integral M f \ nn_integral N f" unfolding nn_integral_def proof (intro SUP_subset_mono) note \sets M = sets N\[simp] \sets M = sets N\[THEN sets_eq_imp_space_eq, simp] show "{g. simple_function M g \ g \ f} \ {g. simple_function N g \ g \ f}" by (simp add: simple_function_def) show "integral\<^sup>S M x \ integral\<^sup>S N x" for x using le_measureD3[OF \M \ N\] by (auto simp add: simple_integral_def intro!: sum_mono mult_mono) qed lemma nn_integral_empty: assumes "space M = {}" shows "nn_integral M f = 0" proof - have "(\\<^sup>+ x. f x \M) = (\\<^sup>+ x. 0 \M)" by(rule nn_integral_cong)(simp add: assms) thus ?thesis by simp qed lemma nn_integral_bot[simp]: "nn_integral bot f = 0" by (simp add: nn_integral_empty) subsubsection\<^marker>\tag unimportant\ \Distributions\ lemma nn_integral_distr: assumes T: "T \ measurable M M'" and f: "f \ borel_measurable (distr M M' T)" shows "integral\<^sup>N (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" using f proof induct case (cong f g) with T show ?case apply (subst nn_integral_cong[of _ f g]) apply simp apply (subst nn_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) apply (simp add: measurable_def Pi_iff) apply simp done next case (set A) then have eq: "\x. x \ space M \ indicator A (T x) = indicator (T -` A \ space M) x" by (auto simp: indicator_def) from set T show ?case by (subst nn_integral_cong[OF eq]) (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets) qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add nn_integral_monotone_convergence_SUP le_fun_def incseq_def image_comp) subsubsection\<^marker>\tag unimportant\ \Counting space\ lemma simple_function_count_space[simp]: "simple_function (count_space A) f \ finite (f ` A)" unfolding simple_function_def by simp lemma nn_integral_count_space: assumes A: "finite {a\A. 0 < f a}" shows "integral\<^sup>N (count_space A) f = (\a|a\A \ 0 < f a. f a)" proof - have *: "(\\<^sup>+x. max 0 (f x) \count_space A) = (\\<^sup>+ x. (\a|a\A \ 0 < f a. f a * indicator {a} x) \count_space A)" by (auto intro!: nn_integral_cong simp add: indicator_def of_bool_def if_distrib sum.If_cases[OF A] max_def le_less) also have "\ = (\a|a\A \ 0 < f a. \\<^sup>+ x. f a * indicator {a} x \count_space A)" by (subst nn_integral_sum) (simp_all add: AE_count_space less_imp_le) also have "\ = (\a|a\A \ 0 < f a. f a)" by (auto intro!: sum.cong simp: one_ennreal_def[symmetric] max_def) finally show ?thesis by (simp add: max.absorb2) qed lemma nn_integral_count_space_finite: "finite A \ (\\<^sup>+x. f x \count_space A) = (\a\A. f a)" by (auto intro!: sum.mono_neutral_left simp: nn_integral_count_space less_le) lemma nn_integral_count_space': assumes "finite A" "\x. x \ B \ x \ A \ f x = 0" "A \ B" shows "(\\<^sup>+x. f x \count_space B) = (\x\A. f x)" proof - have "(\\<^sup>+x. f x \count_space B) = (\a | a \ B \ 0 < f a. f a)" using assms(2,3) by (intro nn_integral_count_space finite_subset[OF _ \finite A\]) (auto simp: less_le) also have "\ = (\a\A. f a)" using assms by (intro sum.mono_neutral_cong_left) (auto simp: less_le) finally show ?thesis . qed lemma nn_integral_bij_count_space: assumes g: "bij_betw g A B" shows "(\\<^sup>+x. f (g x) \count_space A) = (\\<^sup>+x. f x \count_space B)" using g[THEN bij_betw_imp_funcset] by (subst distr_bij_count_space[OF g, symmetric]) (auto intro!: nn_integral_distr[symmetric]) lemma nn_integral_indicator_finite: fixes f :: "'a \ ennreal" assumes f: "finite A" and [measurable]: "\a. a \ A \ {a} \ sets M" shows "(\\<^sup>+x. f x * indicator A x \M) = (\x\A. f x * emeasure M {x})" proof - from f have "(\\<^sup>+x. f x * indicator A x \M) = (\\<^sup>+x. (\a\A. f a * indicator {a} x) \M)" by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\a. x * a" for x] sum.If_cases) also have "\ = (\a\A. f a * emeasure M {a})" by (subst nn_integral_sum) auto finally show ?thesis . qed lemma nn_integral_count_space_nat: fixes f :: "nat \ ennreal" shows "(\\<^sup>+i. f i \count_space UNIV) = (\i. f i)" proof - have "(\\<^sup>+i. f i \count_space UNIV) = (\\<^sup>+i. (\j. f j * indicator {j} i) \count_space UNIV)" proof (intro nn_integral_cong) fix i have "f i = (\j\{i}. f j * indicator {j} i)" by simp also have "\ = (\j. f j * indicator {j} i)" by (rule suminf_finite[symmetric]) auto finally show "f i = (\j. f j * indicator {j} i)" . qed also have "\ = (\j. (\\<^sup>+i. f j * indicator {j} i \count_space UNIV))" by (rule nn_integral_suminf) auto finally show ?thesis by simp qed lemma nn_integral_enat_function: assumes f: "f \ measurable M (count_space UNIV)" shows "(\\<^sup>+ x. ennreal_of_enat (f x) \M) = (\t. emeasure M {x \ space M. t < f x})" proof - define F where "F i = {x\space M. i < f x}" for i :: nat with assms have [measurable]: "\i. F i \ sets M" by auto { fix x assume "x \ space M" have "(\i::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)" using sums_If_finite[of "\r. r < f x" "\_. 1 :: ennreal"] by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal) also have "(\i. (if i < f x then 1 else 0)) = (\i. indicator (F i) x)" using \x \ space M\ by (simp add: one_ennreal_def F_def fun_eq_iff) finally have "ennreal_of_enat (f x) = (\i. indicator (F i) x)" by (simp add: sums_iff) } then have "(\\<^sup>+x. ennreal_of_enat (f x) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" by (simp cong: nn_integral_cong) also have "\ = (\i. emeasure M (F i))" by (simp add: nn_integral_suminf) finally show ?thesis by (simp add: F_def) qed lemma nn_integral_count_space_nn_integral: fixes f :: "'i \ 'a \ ennreal" assumes "countable I" and [measurable]: "\i. i \ I \ f i \ borel_measurable M" shows "(\\<^sup>+x. \\<^sup>+i. f i x \count_space I \M) = (\\<^sup>+i. \\<^sup>+x. f i x \M \count_space I)" proof cases assume "finite I" then show ?thesis by (simp add: nn_integral_count_space_finite nn_integral_sum) next assume "infinite I" then have [simp]: "I \ {}" by auto note * = bij_betw_from_nat_into[OF \countable I\ \infinite I\] have **: "\f. (\i. 0 \ f i) \ (\\<^sup>+i. f i \count_space I) = (\n. f (from_nat_into I n))" by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat) show ?thesis by (simp add: ** nn_integral_suminf from_nat_into) qed lemma of_bool_Bex_eq_nn_integral: assumes unique: "\x y. x \ X \ y \ X \ P x \ P y \ x = y" shows "of_bool (\y\X. P y) = (\\<^sup>+y. of_bool (P y) \count_space X)" proof cases assume "\y\X. P y" then obtain y where "P y" "y \ X" by auto then show ?thesis by (subst nn_integral_count_space'[where A="{y}"]) (auto dest: unique) qed (auto cong: nn_integral_cong_simp) lemma emeasure_UN_countable: assumes sets[measurable]: "\i. i \ I \ X i \ sets M" and I[simp]: "countable I" assumes disj: "disjoint_family_on X I" shows "emeasure M (\(X ` I)) = (\\<^sup>+i. emeasure M (X i) \count_space I)" proof - have eq: "\x. indicator (\(X ` I)) x = \\<^sup>+ i. indicator (X i) x \count_space I" proof cases fix x assume x: "x \ \(X ` I)" then obtain j where j: "x \ X j" "j \ I" by auto with disj have "\i. i \ I \ indicator (X i) x = (indicator {j} i::ennreal)" by (auto simp: disjoint_family_on_def split: split_indicator) with x j show "?thesis x" by (simp cong: nn_integral_cong_simp) qed (auto simp: nn_integral_0_iff_AE) note sets.countable_UN'[unfolded subset_eq, measurable] have "emeasure M (\(X ` I)) = (\\<^sup>+x. indicator (\(X ` I)) x \M)" by simp also have "\ = (\\<^sup>+i. \\<^sup>+x. indicator (X i) x \M \count_space I)" by (simp add: eq nn_integral_count_space_nn_integral) finally show ?thesis by (simp cong: nn_integral_cong_simp) qed lemma emeasure_countable_singleton: assumes sets: "\x. x \ X \ {x} \ sets M" and X: "countable X" shows "emeasure M X = (\\<^sup>+x. emeasure M {x} \count_space X)" proof - have "emeasure M (\i\X. {i}) = (\\<^sup>+x. emeasure M {x} \count_space X)" using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def) also have "(\i\X. {i}) = X" by auto finally show ?thesis . qed lemma measure_eqI_countable: assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A" assumes eq: "\a. a \ A \ emeasure M {a} = emeasure N {a}" shows "M = N" proof (rule measure_eqI) fix X assume "X \ sets M" then have X: "X \ A" by auto moreover from A X have "countable X" by (auto dest: countable_subset) ultimately have "emeasure M X = (\\<^sup>+a. emeasure M {a} \count_space X)" "emeasure N X = (\\<^sup>+a. emeasure N {a} \count_space X)" by (auto intro!: emeasure_countable_singleton) moreover have "(\\<^sup>+a. emeasure M {a} \count_space X) = (\\<^sup>+a. emeasure N {a} \count_space X)" using X by (intro nn_integral_cong eq) auto ultimately show "emeasure M X = emeasure N X" by simp qed simp lemma measure_eqI_countable_AE: assumes [simp]: "sets M = UNIV" "sets N = UNIV" assumes ae: "AE x in M. x \ \" "AE x in N. x \ \" and [simp]: "countable \" assumes eq: "\x. x \ \ \ emeasure M {x} = emeasure N {x}" shows "M = N" proof (rule measure_eqI) fix A have "emeasure N A = emeasure N {x\\. x \ A}" using ae by (intro emeasure_eq_AE) auto also have "\ = (\\<^sup>+x. emeasure N {x} \count_space {x\\. x \ A})" by (intro emeasure_countable_singleton) auto also have "\ = (\\<^sup>+x. emeasure M {x} \count_space {x\\. x \ A})" by (intro nn_integral_cong eq[symmetric]) auto also have "\ = emeasure M {x\\. x \ A}" by (intro emeasure_countable_singleton[symmetric]) auto also have "\ = emeasure M A" using ae by (intro emeasure_eq_AE) auto finally show "emeasure M A = emeasure N A" .. qed simp lemma nn_integral_monotone_convergence_SUP_nat: fixes f :: "'a \ nat \ ennreal" assumes chain: "Complete_Partial_Order.chain (\) (f ` Y)" and nonempty: "Y \ {}" shows "(\\<^sup>+ x. (SUP i\Y. f i x) \count_space UNIV) = (SUP i\Y. (\\<^sup>+ x. f i x \count_space UNIV))" (is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _") proof (rule order_class.order.antisym) show "?rhs \ ?lhs" by (auto intro!: SUP_least SUP_upper nn_integral_mono) next have "\g. incseq g \ range g \ (\i. f i x) ` Y \ (SUP i\Y. f i x) = (SUP i. g i)" for x by (rule ennreal_Sup_countable_SUP) (simp add: nonempty) then obtain g where incseq: "\x. incseq (g x)" and range: "\x. range (g x) \ (\i. f i x) ` Y" and sup: "\x. (SUP i\Y. f i x) = (SUP i. g x i)" by moura from incseq have incseq': "incseq (\i x. g x i)" by(blast intro: incseq_SucI le_funI dest: incseq_SucD) have "?lhs = \\<^sup>+ x. (SUP i. g x i) \?M" by(simp add: sup) also have "\ = (SUP i. \\<^sup>+ x. g x i \?M)" using incseq' by(rule nn_integral_monotone_convergence_SUP) simp also have "\ \ (SUP i\Y. \\<^sup>+ x. f i x \?M)" proof(rule SUP_least) fix n have "\x. \i. g x n = f i x \ i \ Y" using range by blast then obtain I where I: "\x. g x n = f (I x) x" "\x. I x \ Y" by moura have "(\\<^sup>+ x. g x n \count_space UNIV) = (\x. g x n)" by(rule nn_integral_count_space_nat) also have "\ = (SUP m. \x \ (SUP i\Y. \\<^sup>+ x. f i x \?M)" proof(rule SUP_mono) fix m show "\m'\Y. (\x (\\<^sup>+ x. f m' x \?M)" proof(cases "m > 0") case False thus ?thesis using nonempty by auto next case True let ?Y = "I ` {.. f ` Y" using I by auto with chain have chain': "Complete_Partial_Order.chain (\) (f ` ?Y)" by(rule chain_subset) hence "Sup (f ` ?Y) \ f ` ?Y" by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff) then obtain m' where "m' < m" and m': "(SUP i\?Y. f i) = f (I m')" by auto have "I m' \ Y" using I by blast have "(\x (\x {.. \ (SUP i\?Y. f i) x" unfolding Sup_fun_def image_image using \x \ {.. by (rule Sup_upper [OF imageI]) also have "\ = f (I m') x" unfolding m' by simp finally show "g x n \ f (I m') x" . qed also have "\ \ (SUP m. (\x = (\x. f (I m') x)" by(rule suminf_eq_SUP[symmetric]) also have "\ = (\\<^sup>+ x. f (I m') x \?M)" by(rule nn_integral_count_space_nat[symmetric]) finally show ?thesis using \I m' \ Y\ by blast qed qed finally show "(\\<^sup>+ x. g x n \count_space UNIV) \ \" . qed finally show "?lhs \ ?rhs" . qed lemma power_series_tendsto_at_left: assumes nonneg: "\i. 0 \ f i" and summable: "\z. 0 \ z \ z < 1 \ summable (\n. f n * z^n)" shows "((\z. ennreal (\n. f n * z^n)) \ (\n. ennreal (f n))) (at_left (1::real))" proof (intro tendsto_at_left_sequentially) show "0 < (1::real)" by simp fix S :: "nat \ real" assume S: "\n. S n < 1" "\n. 0 < S n" "S \ 1" "incseq S" then have S_nonneg: "\i. 0 \ S i" by (auto intro: less_imp_le) have "(\i. (\\<^sup>+n. f n * S i^n \count_space UNIV)) \ (\\<^sup>+n. ennreal (f n) \count_space UNIV)" proof (rule nn_integral_LIMSEQ) show "incseq (\i n. ennreal (f n * S i^n))" using S by (auto intro!: mult_mono power_mono nonneg ennreal_leI simp: incseq_def le_fun_def less_imp_le) fix n have "(\i. ennreal (f n * S i^n)) \ ennreal (f n * 1^n)" by (intro tendsto_intros tendsto_ennrealI S) then show "(\i. ennreal (f n * S i^n)) \ ennreal (f n)" by simp qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg) also have "(\i. (\\<^sup>+n. f n * S i^n \count_space UNIV)) = (\i. \n. f n * S i^n)" by (subst nn_integral_count_space_nat) (intro ext suminf_ennreal2 mult_nonneg_nonneg nonneg S_nonneg zero_le_power summable S)+ also have "(\\<^sup>+n. ennreal (f n) \count_space UNIV) = (\n. ennreal (f n))" by (simp add: nn_integral_count_space_nat nonneg) finally show "(\n. ennreal (\na. f na * S n ^ na)) \ (\n. ennreal (f n))" . qed subsubsection \Measures with Restricted Space\ lemma simple_function_restrict_space_ennreal: fixes f :: "'a \ ennreal" assumes "\ \ space M \ sets M" shows "simple_function (restrict_space M \) f \ simple_function M (\x. f x * indicator \ x)" proof - { assume "finite (f ` space (restrict_space M \))" then have "finite (f ` space (restrict_space M \) \ {0})" by simp then have "finite ((\x. f x * indicator \ x) ` space M)" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } moreover { assume "finite ((\x. f x * indicator \ x) ` space M)" then have "finite (f ` space (restrict_space M \))" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } ultimately show ?thesis unfolding simple_function_iff_borel_measurable borel_measurable_restrict_space_iff_ennreal[OF assms] by auto qed lemma simple_function_restrict_space: fixes f :: "'a \ 'b::real_normed_vector" assumes "\ \ space M \ sets M" shows "simple_function (restrict_space M \) f \ simple_function M (\x. indicator \ x *\<^sub>R f x)" proof - { assume "finite (f ` space (restrict_space M \))" then have "finite (f ` space (restrict_space M \) \ {0})" by simp then have "finite ((\x. indicator \ x *\<^sub>R f x) ` space M)" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } moreover { assume "finite ((\x. indicator \ x *\<^sub>R f x) ` space M)" then have "finite (f ` space (restrict_space M \))" by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } ultimately show ?thesis unfolding simple_function_iff_borel_measurable borel_measurable_restrict_space_iff[OF assms] by auto qed lemma simple_integral_restrict_space: assumes \: "\ \ space M \ sets M" "simple_function (restrict_space M \) f" shows "simple_integral (restrict_space M \) f = simple_integral M (\x. f x * indicator \ x)" using simple_function_restrict_space_ennreal[THEN iffD1, OF \, THEN simple_functionD(1)] by (auto simp add: space_restrict_space emeasure_restrict_space[OF \(1)] le_infI2 simple_integral_def split: split_indicator split_indicator_asm intro!: sum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure]) lemma nn_integral_restrict_space: assumes \[simp]: "\ \ space M \ sets M" shows "nn_integral (restrict_space M \) f = nn_integral M (\x. f x * indicator \ x)" proof - let ?R = "restrict_space M \" and ?X = "\M f. {s. simple_function M s \ s \ f \ (\x. s x < top)}" have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\x. f x * indicator \ x)" proof (safe intro!: image_eqI) fix s assume s: "simple_function ?R s" "s \ f" "\x. s x < top" from s show "integral\<^sup>S (restrict_space M \) s = integral\<^sup>S M (\x. s x * indicator \ x)" by (intro simple_integral_restrict_space) auto from s show "simple_function M (\x. s x * indicator \ x)" by (simp add: simple_function_restrict_space_ennreal) from s show "(\x. s x * indicator \ x) \ (\x. f x * indicator \ x)" "\x. s x * indicator \ x < top" by (auto split: split_indicator simp: le_fun_def image_subset_iff) next fix s assume s: "simple_function M s" "s \ (\x. f x * indicator \ x)" "\x. s x < top" then have "simple_function M (\x. s x * indicator (\ \ space M) x)" (is ?s') by (intro simple_function_mult simple_function_indicator) auto also have "?s' \ simple_function M (\x. s x * indicator \ x)" by (rule simple_function_cong) (auto split: split_indicator) finally show sf: "simple_function (restrict_space M \) s" by (simp add: simple_function_restrict_space_ennreal) from s have s_eq: "s = (\x. s x * indicator \ x)" by (auto simp add: fun_eq_iff le_fun_def image_subset_iff split: split_indicator split_indicator_asm intro: antisym) show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \) s" by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \ sf]) show "\x. s x < top" using s by (auto simp: image_subset_iff) from s show "s \ f" by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) qed then show ?thesis unfolding nn_integral_def_finite by (simp cong del: SUP_cong_simp) qed lemma nn_integral_count_space_indicator: assumes "NO_MATCH (UNIV::'a set) (X::'a set)" shows "(\\<^sup>+x. f x \count_space X) = (\\<^sup>+x. f x * indicator X x \count_space UNIV)" by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space) lemma nn_integral_count_space_eq: "(\x. x \ A - B \ f x = 0) \ (\x. x \ B - A \ f x = 0) \ (\\<^sup>+x. f x \count_space A) = (\\<^sup>+x. f x \count_space B)" by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) lemma nn_integral_ge_point: assumes "x \ A" shows "p x \ \\<^sup>+ x. p x \count_space A" proof - from assms have "p x \ \\<^sup>+ x. p x \count_space {x}" by(auto simp add: nn_integral_count_space_finite max_def) also have "\ = \\<^sup>+ x'. p x' * indicator {x} x' \count_space A" using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) also have "\ \ \\<^sup>+ x. p x \count_space A" by(rule nn_integral_mono)(simp add: indicator_def) finally show ?thesis . qed subsubsection \Measure spaces with an associated density\ definition\<^marker>\tag important\ density :: "'a measure \ ('a \ ennreal) \ 'a measure" where "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" lemma shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M" and space_density[simp]: "space (density M f) = space M" by (auto simp: density_def) (* FIXME: add conversion to simplify space, sets and measurable *) lemma space_density_imp[measurable_dest]: "\x M f. x \ space (density M f) \ x \ space M" by auto lemma shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" unfolding measurable_def simple_function_def by simp_all lemma density_cong: "f \ borel_measurable M \ f' \ borel_measurable M \ (AE x in M. f x = f' x) \ density M f = density M f'" unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed) lemma emeasure_density: assumes f[measurable]: "f \ borel_measurable M" and A[measurable]: "A \ sets M" shows "emeasure (density M f) A = (\\<^sup>+ x. f x * indicator A x \M)" (is "_ = ?\ A") unfolding density_def proof (rule emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) ?\" using f by (auto simp: positive_def) show "countably_additive (sets M) ?\" proof (intro countably_additiveI) fix A :: "nat \ 'a set" assume "range A \ sets M" then have "\i. A i \ sets M" by auto then have *: "\i. (\x. f x * indicator (A i) x) \ borel_measurable M" by auto assume disj: "disjoint_family A" then have "(\n. ?\ (A n)) = (\\<^sup>+ x. (\n. f x * indicator (A n) x) \M)" using f * by (subst nn_integral_suminf) auto also have "(\\<^sup>+ x. (\n. f x * indicator (A n) x) \M) = (\\<^sup>+ x. f x * (\n. indicator (A n) x) \M)" using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE) also have "\ = (\\<^sup>+ x. f x * indicator (\n. A n) x \M)" unfolding suminf_indicator[OF disj] .. finally show "(\i. \\<^sup>+ x. f x * indicator (A i) x \M) = \\<^sup>+ x. f x * indicator (\i. A i) x \M" . qed qed fact lemma null_sets_density_iff: assumes f: "f \ borel_measurable M" shows "A \ null_sets (density M f) \ A \ sets M \ (AE x in M. x \ A \ f x = 0)" proof - { assume "A \ sets M" have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ emeasure M {x \ space M. f x * indicator A x \ 0} = 0" using f \A \ sets M\ by (intro nn_integral_0_iff) auto also have "\ \ (AE x in M. f x * indicator A x = 0)" using f \A \ sets M\ by (intro AE_iff_measurable[OF _ refl, symmetric]) auto also have "(AE x in M. f x * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" by (auto simp add: indicator_def max_def split: if_split_asm) finally have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ (AE x in M. x \ A \ f x \ 0)" . } with f show ?thesis by (simp add: null_sets_def emeasure_density cong: conj_cong) qed lemma AE_density: assumes f: "f \ borel_measurable M" shows "(AE x in density M f. P x) \ (AE x in M. 0 < f x \ P x)" proof assume "AE x in density M f. P x" with f obtain N where "{x \ space M. \ P x} \ N" "N \ sets M" and ae: "AE x in M. x \ N \ f x = 0" by (auto simp: eventually_ae_filter null_sets_density_iff) then have "AE x in M. x \ N \ P x" by auto with ae show "AE x in M. 0 < f x \ P x" by (rule eventually_elim2) auto next fix N assume ae: "AE x in M. 0 < f x \ P x" then obtain N where "{x \ space M. \ (0 < f x \ P x)} \ N" "N \ null_sets M" by (auto simp: eventually_ae_filter) then have *: "{x \ space (density M f). \ P x} \ N \ {x\space M. f x = 0}" "N \ {x\space M. f x = 0} \ sets M" and ae2: "AE x in M. x \ N" using f by (auto simp: subset_eq zero_less_iff_neq_zero intro!: AE_not_in) show "AE x in density M f. P x" using ae2 unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] by (intro exI[of _ "N \ {x\space M. f x = 0}"] conjI *) (auto elim: eventually_elim2) qed lemma\<^marker>\tag important\ nn_integral_density: assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "integral\<^sup>N (density M f) g = (\\<^sup>+ x. f x * g x \M)" using g proof induct case (cong u v) then show ?case apply (subst nn_integral_cong[OF cong(3)]) apply (simp_all cong: nn_integral_cong) done next case (set A) then show ?case by (simp add: emeasure_density f) next case (mult u c) moreover have "\x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) ultimately show ?case using f by (simp add: nn_integral_cmult) next case (add u v) then have "\x. f x * (v x + u x) = f x * v x + f x * u x" by (simp add: distrib_left) with add f show ?case by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric]) next case (seq U) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" by eventually_elim (simp add: SUP_mult_left_ennreal seq) from seq f show ?case apply (simp add: nn_integral_monotone_convergence_SUP image_comp) apply (subst nn_integral_cong_AE[OF eq]) apply (subst nn_integral_monotone_convergence_SUP_AE) apply (auto simp: incseq_def le_fun_def intro!: mult_left_mono) done qed lemma density_distr: assumes [measurable]: "f \ borel_measurable N" "X \ measurable M N" shows "density (distr M N X) f = distr (density M (\x. f (X x))) N X" by (intro measure_eqI) (auto simp add: emeasure_density nn_integral_distr emeasure_distr split: split_indicator intro!: nn_integral_cong) lemma emeasure_restricted: assumes S: "S \ sets M" and X: "X \ sets M" shows "emeasure (density M (indicator S)) X = emeasure M (S \ X)" proof - have "emeasure (density M (indicator S)) X = (\\<^sup>+x. indicator S x * indicator X x \M)" using S X by (simp add: emeasure_density) also have "\ = (\\<^sup>+x. indicator (S \ X) x \M)" by (auto intro!: nn_integral_cong simp: indicator_def) also have "\ = emeasure M (S \ X)" using S X by (simp add: sets.Int) finally show ?thesis . qed lemma measure_restricted: "S \ sets M \ X \ sets M \ measure (density M (indicator S)) X = measure M (S \ X)" by (simp add: emeasure_restricted measure_def) lemma (in finite_measure) finite_measure_restricted: "S \ sets M \ finite_measure (density M (indicator S))" by standard (simp add: emeasure_restricted) lemma emeasure_density_const: "A \ sets M \ emeasure (density M (\_. c)) A = c * emeasure M A" by (auto simp: nn_integral_cmult_indicator emeasure_density) lemma measure_density_const: "A \ sets M \ c \ \ \ measure (density M (\_. c)) A = enn2real c * measure M A" by (auto simp: emeasure_density_const measure_def enn2real_mult) lemma density_density_eq: "f \ borel_measurable M \ g \ borel_measurable M \ density (density M f) g = density M (\x. f x * g x)" by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps) lemma distr_density_distr: assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" and inv: "\x\space M. T' (T x) = x" assumes f: "f \ borel_measurable M'" shows "distr (density (distr M M' T) f) M T' = density M (f \ T)" (is "?R = ?L") proof (rule measure_eqI) fix A assume A: "A \ sets ?R" { fix x assume "x \ space M" with sets.sets_into_space[OF A] have "indicator (T' -` A \ space M') (T x) = (indicator A x :: ennreal)" using T inv by (auto simp: indicator_def measurable_space) } with A T T' f show "emeasure ?R A = emeasure ?L A" by (simp add: measurable_comp emeasure_density emeasure_distr nn_integral_distr measurable_sets cong: nn_integral_cong) qed simp lemma density_density_divide: fixes f g :: "'a \ real" assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" assumes ac: "AE x in M. f x = 0 \ g x = 0" shows "density (density M f) (\x. g x / f x) = density M g" proof - have "density M g = density M (\x. ennreal (f x) * ennreal (g x / f x))" using f g ac by (auto intro!: density_cong measurable_If simp: ennreal_mult[symmetric]) then show ?thesis using f g by (subst density_density_eq) auto qed lemma density_1: "density M (\_. 1) = M" by (intro measure_eqI) (auto simp: emeasure_density) lemma emeasure_density_add: assumes X: "X \ sets M" assumes Mf[measurable]: "f \ borel_measurable M" assumes Mg[measurable]: "g \ borel_measurable M" shows "emeasure (density M f) X + emeasure (density M g) X = emeasure (density M (\x. f x + g x)) X" using assms apply (subst (1 2 3) emeasure_density, simp_all) [] apply (subst nn_integral_add[symmetric], simp_all) [] apply (intro nn_integral_cong, simp split: split_indicator) done subsubsection \Point measure\ definition\<^marker>\tag important\ point_measure :: "'a set \ ('a \ ennreal) \ 'a measure" where "point_measure A f = density (count_space A) f" lemma shows space_point_measure: "space (point_measure A f) = A" and sets_point_measure: "sets (point_measure A f) = Pow A" by (auto simp: point_measure_def) lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)" by (simp add: sets_point_measure) lemma measurable_point_measure_eq1[simp]: "g \ measurable (point_measure A f) M \ g \ A \ space M" unfolding point_measure_def by simp lemma measurable_point_measure_eq2_finite[simp]: "finite A \ g \ measurable M (point_measure A f) \ (g \ space M \ A \ (\a\A. g -` {a} \ space M \ sets M))" unfolding point_measure_def by (simp add: measurable_count_space_eq2) lemma simple_function_point_measure[simp]: "simple_function (point_measure A f) g \ finite (g ` A)" by (simp add: point_measure_def) lemma emeasure_point_measure: assumes A: "finite {a\X. 0 < f a}" "X \ A" shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" proof - have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" using \X \ A\ by auto with A show ?thesis by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def of_bool_def) qed lemma emeasure_point_measure_finite: "finite A \ X \ A \ emeasure (point_measure A f) X = (\a\X. f a)" by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le) lemma emeasure_point_measure_finite2: "X \ A \ finite X \ emeasure (point_measure A f) X = (\a\X. f a)" by (subst emeasure_point_measure) (auto dest: finite_subset intro!: sum.mono_neutral_left simp: less_le) lemma null_sets_point_measure_iff: "X \ null_sets (point_measure A f) \ X \ A \ (\x\X. f x = 0)" by (auto simp: AE_count_space null_sets_density_iff point_measure_def) lemma AE_point_measure: "(AE x in point_measure A f. P x) \ (\x\A. 0 < f x \ P x)" unfolding point_measure_def by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) lemma nn_integral_point_measure: "finite {a\A. 0 < f a \ 0 < g a} \ integral\<^sup>N (point_measure A f) g = (\a|a\A \ 0 < f a \ 0 < g a. f a * g a)" unfolding point_measure_def by (subst nn_integral_density) (simp_all add: nn_integral_density nn_integral_count_space ennreal_zero_less_mult_iff) lemma nn_integral_point_measure_finite: "finite A \ integral\<^sup>N (point_measure A f) g = (\a\A. f a * g a)" by (subst nn_integral_point_measure) (auto intro!: sum.mono_neutral_left simp: less_le) subsubsection \Uniform measure\ definition\<^marker>\tag important\ "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" lemma shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M" and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" by (auto simp: uniform_measure_def) lemma emeasure_uniform_measure[simp]: assumes A: "A \ sets M" and B: "B \ sets M" shows "emeasure (uniform_measure M A) B = emeasure M (A \ B) / emeasure M A" proof - from A B have "emeasure (uniform_measure M A) B = (\\<^sup>+x. (1 / emeasure M A) * indicator (A \ B) x \M)" by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator intro!: nn_integral_cong) also have "\ = emeasure M (A \ B) / emeasure M A" using A B by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute) finally show ?thesis . qed lemma measure_uniform_measure[simp]: assumes A: "emeasure M A \ 0" "emeasure M A \ \" and B: "B \ sets M" shows "measure (uniform_measure M A) B = measure M (A \ B) / measure M A" using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A by (cases "emeasure M A" "emeasure M (A \ B)" rule: ennreal2_cases) (simp_all add: measure_def divide_ennreal top_ennreal.rep_eq top_ereal_def ennreal_top_divide) lemma AE_uniform_measureI: "A \ sets M \ (AE x in M. x \ A \ P x) \ (AE x in uniform_measure M A. P x)" unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def) lemma emeasure_uniform_measure_1: "emeasure M S \ 0 \ emeasure M S \ \ \ emeasure (uniform_measure M S) S = 1" by (subst emeasure_uniform_measure) (simp_all add: emeasure_neq_0_sets emeasure_eq_ennreal_measure divide_ennreal zero_less_iff_neq_zero[symmetric]) lemma nn_integral_uniform_measure: assumes f[measurable]: "f \ borel_measurable M" and S[measurable]: "S \ sets M" shows "(\\<^sup>+x. f x \uniform_measure M S) = (\\<^sup>+x. f x * indicator S x \M) / emeasure M S" proof - { assume "emeasure M S = \" then have ?thesis by (simp add: uniform_measure_def nn_integral_density f) } moreover { assume [simp]: "emeasure M S = 0" then have ae: "AE x in M. x \ S" using sets.sets_into_space[OF S] by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong) from ae have "(\\<^sup>+ x. indicator S x / 0 * f x \M) = 0" by (subst nn_integral_0_iff_AE) auto moreover from ae have "(\\<^sup>+ x. f x * indicator S x \M) = 0" by (subst nn_integral_0_iff_AE) auto ultimately have ?thesis by (simp add: uniform_measure_def nn_integral_density f) } moreover have "emeasure M S \ 0 \ emeasure M S \ \ \ ?thesis" unfolding uniform_measure_def by (subst nn_integral_density) (auto simp: ennreal_times_divide f nn_integral_divide[symmetric] mult.commute) ultimately show ?thesis by blast qed lemma AE_uniform_measure: assumes "emeasure M A \ 0" "emeasure M A < \" shows "(AE x in uniform_measure M A. P x) \ (AE x in M. x \ A \ P x)" proof - have "A \ sets M" using \emeasure M A \ 0\ by (metis emeasure_notin_sets) moreover have "\x. 0 < indicator A x / emeasure M A \ x \ A" using assms by (cases "emeasure M A") (auto split: split_indicator simp: ennreal_zero_less_divide) ultimately show ?thesis unfolding uniform_measure_def by (simp add: AE_density) qed subsubsection\<^marker>\tag unimportant\ \Null measure\ lemma null_measure_eq_density: "null_measure M = density M (\_. 0)" by (intro measure_eqI) (simp_all add: emeasure_density) lemma nn_integral_null_measure[simp]: "(\\<^sup>+x. f x \null_measure M) = 0" by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def intro!: exI[of _ "\x. 0"]) lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M" proof (intro measure_eqI) fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A" by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure) qed simp subsubsection \Uniform count measure\ definition\<^marker>\tag important\ "uniform_count_measure A = point_measure A (\x. 1 / card A)" lemma shows space_uniform_count_measure: "space (uniform_count_measure A) = A" and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) lemma sets_uniform_count_measure_count_space[measurable_cong]: "sets (uniform_count_measure A) = sets (count_space A)" by (simp add: sets_uniform_count_measure) lemma emeasure_uniform_count_measure: "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult ennreal_of_nat_eq_real_of_nat) lemma measure_uniform_count_measure: "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult) lemma space_uniform_count_measure_empty_iff [simp]: "space (uniform_count_measure X) = {} \ X = {}" by(simp add: space_uniform_count_measure) lemma sets_uniform_count_measure_eq_UNIV [simp]: "sets (uniform_count_measure UNIV) = UNIV \ True" "UNIV = sets (uniform_count_measure UNIV) \ True" by(simp_all add: sets_uniform_count_measure) subsubsection\<^marker>\tag unimportant\ \Scaled measure\ lemma nn_integral_scale_measure: assumes f: "f \ borel_measurable M" shows "nn_integral (scale_measure r M) f = r * nn_integral M f" using f proof induction case (cong f g) thus ?case by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp) next case (mult f c) thus ?case by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute) next case (add f g) thus ?case by(simp add: nn_integral_add distrib_left) next case (seq U) thus ?case by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal image_comp) qed simp end diff --git a/src/HOL/Analysis/Radon_Nikodym.thy b/src/HOL/Analysis/Radon_Nikodym.thy --- a/src/HOL/Analysis/Radon_Nikodym.thy +++ b/src/HOL/Analysis/Radon_Nikodym.thy @@ -1,1120 +1,1129 @@ (* Title: HOL/Analysis/Radon_Nikodym.thy Author: Johannes Hölzl, TU München *) section \Radon-Nikod{\'y}m Derivative\ theory Radon_Nikodym imports Bochner_Integration begin definition\<^marker>\tag important\ diff_measure :: "'a measure \ 'a measure \ 'a measure" where "diff_measure M N = measure_of (space M) (sets M) (\A. emeasure M A - emeasure N A)" lemma shows space_diff_measure[simp]: "space (diff_measure M N) = space M" and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M" by (auto simp: diff_measure_def) lemma emeasure_diff_measure: assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N" assumes pos: "\A. A \ sets M \ emeasure N A \ emeasure M A" and A: "A \ sets M" shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\ A") unfolding diff_measure_def proof (rule emeasure_measure_of_sigma) show "sigma_algebra (space M) (sets M)" .. show "positive (sets M) ?\" using pos by (simp add: positive_def) show "countably_additive (sets M) ?\" proof (rule countably_additiveI) fix A :: "nat \ _" assume A: "range A \ sets M" and "disjoint_family A" then have suminf: "(\i. emeasure M (A i)) = emeasure M (\i. A i)" "(\i. emeasure N (A i)) = emeasure N (\i. A i)" by (simp_all add: suminf_emeasure sets_eq) with A have "(\i. emeasure M (A i) - emeasure N (A i)) = (\i. emeasure M (A i)) - (\i. emeasure N (A i))" using fin pos[of "A _"] by (intro ennreal_suminf_minus) (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure) then show "(\i. emeasure M (A i) - emeasure N (A i)) = emeasure M (\i. A i) - emeasure N (\i. A i) " by (simp add: suminf) qed qed fact text \An equivalent characterization of sigma-finite spaces is the existence of integrable positive functions (or, still equivalently, the existence of a probability measure which is in the same measure class as the original measure).\ proposition (in sigma_finite_measure) obtain_positive_integrable_function: obtains f::"'a \ real" where "f \ borel_measurable M" "\x. f x > 0" "\x. f x \ 1" "integrable M f" proof - obtain A :: "nat \ 'a set" where "range A \ sets M" "(\i. A i) = space M" "\i. emeasure M (A i) \ \" using sigma_finite by auto then have [measurable]:"A n \ sets M" for n by auto define g where "g = (\x. (\n. (1/2)^(Suc n) * indicator (A n) x / (1+ measure M (A n))))" have [measurable]: "g \ borel_measurable M" unfolding g_def by auto have *: "summable (\n. (1/2)^(Suc n) * indicator (A n) x / (1+ measure M (A n)))" for x apply (rule summable_comparison_test'[of "\n. (1/2)^(Suc n)" 0]) using power_half_series summable_def by (auto simp add: indicator_def divide_simps) have "g x \ (\n. (1/2)^(Suc n))" for x unfolding g_def apply (rule suminf_le) using * power_half_series summable_def by (auto simp add: indicator_def divide_simps) then have g_le_1: "g x \ 1" for x using power_half_series sums_unique by fastforce have g_pos: "g x > 0" if "x \ space M" for x unfolding g_def proof (subst suminf_pos_iff[OF *[of x]], auto) obtain i where "x \ A i" using \(\i. A i) = space M\ \x \ space M\ by auto then have "0 < (1 / 2) ^ Suc i * indicator (A i) x / (1 + Sigma_Algebra.measure M (A i))" unfolding indicator_def apply (auto simp add: divide_simps) using measure_nonneg[of M "A i"] by (auto, meson add_nonneg_nonneg linorder_not_le mult_nonneg_nonneg zero_le_numeral zero_le_one zero_le_power) then show "\i. 0 < (1 / 2) ^ i * indicator (A i) x / (2 + 2 * Sigma_Algebra.measure M (A i))" by auto qed have "integrable M g" unfolding g_def proof (rule integrable_suminf) fix n show "integrable M (\x. (1 / 2) ^ Suc n * indicator (A n) x / (1 + Sigma_Algebra.measure M (A n)))" using \emeasure M (A n) \ \\ by (auto intro!: integrable_mult_right integrable_divide_zero integrable_real_indicator simp add: top.not_eq_extremum) next show "AE x in M. summable (\n. norm ((1 / 2) ^ Suc n * indicator (A n) x / (1 + Sigma_Algebra.measure M (A n))))" using * by auto show "summable (\n. (\x. norm ((1 / 2) ^ Suc n * indicator (A n) x / (1 + Sigma_Algebra.measure M (A n))) \M))" apply (rule summable_comparison_test'[of "\n. (1/2)^(Suc n)" 0], auto) using power_half_series summable_def apply auto[1] apply (auto simp add: field_split_simps) using measure_nonneg[of M] not_less by fastforce qed define f where "f = (\x. if x \ space M then g x else 1)" have "f x > 0" for x unfolding f_def using g_pos by auto moreover have "f x \ 1" for x unfolding f_def using g_le_1 by auto moreover have [measurable]: "f \ borel_measurable M" unfolding f_def by auto moreover have "integrable M f" apply (subst integrable_cong[of _ _ _ g]) unfolding f_def using \integrable M g\ by auto ultimately show "(\f. f \ borel_measurable M \ (\x. 0 < f x) \ (\x. f x \ 1) \ integrable M f \ thesis) \ thesis" by (meson that) qed lemma (in sigma_finite_measure) Ex_finite_integrable_function: "\h\borel_measurable M. integral\<^sup>N M h \ \ \ (\x\space M. 0 < h x \ h x < \)" proof - obtain A :: "nat \ 'a set" where range[measurable]: "range A \ sets M" and space: "(\i. A i) = space M" and measure: "\i. emeasure M (A i) \ \" and disjoint: "disjoint_family A" using sigma_finite_disjoint by blast let ?B = "\i. 2^Suc i * emeasure M (A i)" have [measurable]: "\i. A i \ sets M" using range by fastforce+ have "\i. \x. 0 < x \ x < inverse (?B i)" proof fix i show "\x. 0 < x \ x < inverse (?B i)" using measure[of i] by (auto intro!: dense simp: ennreal_inverse_positive ennreal_mult_eq_top_iff power_eq_top_ennreal) qed from choice[OF this] obtain n where n: "\i. 0 < n i" "\i. n i < inverse (2^Suc i * emeasure M (A i))" by auto { fix i have "0 \ n i" using n(1)[of i] by auto } note pos = this let ?h = "\x. \i. n i * indicator (A i) x" show ?thesis proof (safe intro!: bexI[of _ ?h] del: notI) have "integral\<^sup>N M ?h = (\i. n i * emeasure M (A i))" using pos by (simp add: nn_integral_suminf nn_integral_cmult_indicator) also have "\ \ (\i. ennreal ((1/2)^Suc i))" proof (intro suminf_le allI) fix N have "n N * emeasure M (A N) \ inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)" using n[of N] by (intro mult_right_mono) auto also have "\ = (1/2)^Suc N * (inverse (emeasure M (A N)) * emeasure M (A N))" using measure[of N] by (simp add: ennreal_inverse_power divide_ennreal_def ennreal_inverse_mult power_eq_top_ennreal less_top[symmetric] mult_ac del: power_Suc) also have "\ \ inverse (ennreal 2) ^ Suc N" using measure[of N] by (cases "emeasure M (A N)"; cases "emeasure M (A N) = 0") (auto simp: inverse_ennreal ennreal_mult[symmetric] divide_ennreal_def simp del: power_Suc) also have "\ = ennreal (inverse 2 ^ Suc N)" by (subst ennreal_power[symmetric], simp) (simp add: inverse_ennreal) finally show "n N * emeasure M (A N) \ ennreal ((1/2)^Suc N)" by simp qed auto also have "\ < top" unfolding less_top[symmetric] by (rule ennreal_suminf_neq_top) (auto simp: summable_geometric summable_Suc_iff simp del: power_Suc) finally show "integral\<^sup>N M ?h \ \" by (auto simp: top_unique) next { fix x assume "x \ space M" then obtain i where "x \ A i" using space[symmetric] by auto with disjoint n have "?h x = n i" by (auto intro!: suminf_cmult_indicator intro: less_imp_le) then show "0 < ?h x" and "?h x < \" using n[of i] by (auto simp: less_top[symmetric]) } note pos = this qed measurable qed subsection "Absolutely continuous" definition\<^marker>\tag important\ absolutely_continuous :: "'a measure \ 'a measure \ bool" where "absolutely_continuous M N \ null_sets M \ null_sets N" lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M" unfolding absolutely_continuous_def by (auto simp: null_sets_count_space) lemma absolutely_continuousI_density: "f \ borel_measurable M \ absolutely_continuous M (density M f)" by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in) lemma absolutely_continuousI_point_measure_finite: "(\x. \ x \ A ; f x \ 0 \ \ g x \ 0) \ absolutely_continuous (point_measure A f) (point_measure A g)" unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff) lemma absolutely_continuousD: "absolutely_continuous M N \ A \ sets M \ emeasure M A = 0 \ emeasure N A = 0" by (auto simp: absolutely_continuous_def null_sets_def) lemma absolutely_continuous_AE: assumes sets_eq: "sets M' = sets M" and "absolutely_continuous M M'" "AE x in M. P x" shows "AE x in M'. P x" proof - from \AE x in M. P x\ obtain N where N: "N \ null_sets M" "{x\space M. \ P x} \ N" unfolding eventually_ae_filter by auto show "AE x in M'. P x" proof (rule AE_I') show "{x\space M'. \ P x} \ N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp from \absolutely_continuous M M'\ show "N \ null_sets M'" using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto qed qed subsection "Existence of the Radon-Nikodym derivative" proposition (in finite_measure) Radon_Nikodym_finite_measure: assumes "finite_measure N" and sets_eq[simp]: "sets N = sets M" assumes "absolutely_continuous M N" shows "\f \ borel_measurable M. density M f = N" proof - interpret N: finite_measure N by fact define G where "G = {g \ borel_measurable M. \A\sets M. (\\<^sup>+x. g x * indicator A x \M) \ N A}" have [measurable_dest]: "f \ G \ f \ borel_measurable M" and G_D: "\A. f \ G \ A \ sets M \ (\\<^sup>+x. f x * indicator A x \M) \ N A" for f by (auto simp: G_def) note this[measurable_dest] have "(\x. 0) \ G" unfolding G_def by auto hence "G \ {}" by auto { fix f g assume f[measurable]: "f \ G" and g[measurable]: "g \ G" have "(\x. max (g x) (f x)) \ G" (is "?max \ G") unfolding G_def proof safe let ?A = "{x \ space M. f x \ g x}" have "?A \ sets M" using f g unfolding G_def by auto fix A assume [measurable]: "A \ sets M" have union: "((?A \ A) \ ((space M - ?A) \ A)) = A" using sets.sets_into_space[OF \A \ sets M\] by auto have "\x. x \ space M \ max (g x) (f x) * indicator A x = g x * indicator (?A \ A) x + f x * indicator ((space M - ?A) \ A) x" by (auto simp: indicator_def max_def) hence "(\\<^sup>+x. max (g x) (f x) * indicator A x \M) = (\\<^sup>+x. g x * indicator (?A \ A) x \M) + (\\<^sup>+x. f x * indicator ((space M - ?A) \ A) x \M)" by (auto cong: nn_integral_cong intro!: nn_integral_add) also have "\ \ N (?A \ A) + N ((space M - ?A) \ A)" using f g unfolding G_def by (auto intro!: add_mono) also have "\ = N A" using union by (subst plus_emeasure) auto finally show "(\\<^sup>+x. max (g x) (f x) * indicator A x \M) \ N A" . qed auto } note max_in_G = this { fix f assume "incseq f" and f: "\i. f i \ G" then have [measurable]: "\i. f i \ borel_measurable M" by (auto simp: G_def) have "(\x. SUP i. f i x) \ G" unfolding G_def proof safe show "(\x. SUP i. f i x) \ borel_measurable M" by measurable next fix A assume "A \ sets M" have "(\\<^sup>+x. (SUP i. f i x) * indicator A x \M) = (\\<^sup>+x. (SUP i. f i x * indicator A x) \M)" by (intro nn_integral_cong) (simp split: split_indicator) also have "\ = (SUP i. (\\<^sup>+x. f i x * indicator A x \M))" using \incseq f\ f \A \ sets M\ by (intro nn_integral_monotone_convergence_SUP) (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator) finally show "(\\<^sup>+x. (SUP i. f i x) * indicator A x \M) \ N A" using f \A \ sets M\ by (auto intro!: SUP_least simp: G_D) qed } note SUP_in_G = this let ?y = "SUP g \ G. integral\<^sup>N M g" have y_le: "?y \ N (space M)" unfolding G_def proof (safe intro!: SUP_least) fix g assume "\A\sets M. (\\<^sup>+x. g x * indicator A x \M) \ N A" from this[THEN bspec, OF sets.top] show "integral\<^sup>N M g \ N (space M)" by (simp cong: nn_integral_cong) qed - from ennreal_SUP_countable_SUP [OF \G \ {}\, of "integral\<^sup>N M"] guess ys .. note ys = this + from ennreal_SUP_countable_SUP [OF \G \ {}\, of "integral\<^sup>N M"] + obtain ys :: "nat \ ennreal" + where ys: "range ys \ integral\<^sup>N M ` G \ Sup (integral\<^sup>N M ` G) = Sup (range ys)" + by auto then have "\n. \g. g\G \ integral\<^sup>N M g = ys n" proof safe fix n assume "range ys \ integral\<^sup>N M ` G" hence "ys n \ integral\<^sup>N M ` G" by auto thus "\g. g\G \ integral\<^sup>N M g = ys n" by auto qed from choice[OF this] obtain gs where "\i. gs i \ G" "\n. integral\<^sup>N M (gs n) = ys n" by auto hence y_eq: "?y = (SUP i. integral\<^sup>N M (gs i))" using ys by auto let ?g = "\i x. Max ((\n. gs n x) ` {..i})" define f where [abs_def]: "f x = (SUP i. ?g i x)" for x let ?F = "\A x. f x * indicator A x" have gs_not_empty: "\i x. (\n. gs n x) ` {..i} \ {}" by auto { fix i have "?g i \ G" proof (induct i) case 0 thus ?case by simp fact next case (Suc i) with Suc gs_not_empty \gs (Suc i) \ G\ show ?case by (auto simp add: atMost_Suc intro!: max_in_G) qed } note g_in_G = this have "incseq ?g" using gs_not_empty by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc) from SUP_in_G[OF this g_in_G] have [measurable]: "f \ G" unfolding f_def . then have [measurable]: "f \ borel_measurable M" unfolding G_def by auto have "integral\<^sup>N M f = (SUP i. integral\<^sup>N M (?g i))" unfolding f_def using g_in_G \incseq ?g\ by (auto intro!: nn_integral_monotone_convergence_SUP simp: G_def) also have "\ = ?y" proof (rule antisym) show "(SUP i. integral\<^sup>N M (?g i)) \ ?y" using g_in_G by (auto intro: SUP_mono) show "?y \ (SUP i. integral\<^sup>N M (?g i))" unfolding y_eq by (auto intro!: SUP_mono nn_integral_mono Max_ge) qed finally have int_f_eq_y: "integral\<^sup>N M f = ?y" . have upper_bound: "\A\sets M. N A \ density M f A" proof (rule ccontr) assume "\ ?thesis" then obtain A where A[measurable]: "A \ sets M" and f_less_N: "density M f A < N A" by (auto simp: not_le) then have pos_A: "0 < M A" using \absolutely_continuous M N\[THEN absolutely_continuousD, OF A] by (auto simp: zero_less_iff_neq_zero) define b where "b = (N A - density M f A) / M A / 2" with f_less_N pos_A have "0 < b" "b \ top" by (auto intro!: diff_gr0_ennreal simp: zero_less_iff_neq_zero diff_eq_0_iff_ennreal ennreal_divide_eq_top_iff) let ?f = "\x. f x + b" have "nn_integral M f \ top" using \f \ G\[THEN G_D, of "space M"] by (auto simp: top_unique cong: nn_integral_cong) with \b \ top\ interpret Mf: finite_measure "density M ?f" by (intro finite_measureI) (auto simp: field_simps mult_indicator_subset ennreal_mult_eq_top_iff emeasure_density nn_integral_cmult_indicator nn_integral_add cong: nn_integral_cong) from unsigned_Hahn_decomposition[of "density M ?f" N A] obtain Y where [measurable]: "Y \ sets M" and [simp]: "Y \ A" and Y1: "\C. C \ sets M \ C \ Y \ density M ?f C \ N C" and Y2: "\C. C \ sets M \ C \ A \ C \ Y = {} \ N C \ density M ?f C" by auto let ?f' = "\x. f x + b * indicator Y x" have "M Y \ 0" proof assume "M Y = 0" then have "N Y = 0" using \absolutely_continuous M N\[THEN absolutely_continuousD, of Y] by auto then have "N A = N (A - Y)" by (subst emeasure_Diff) auto also have "\ \ density M ?f (A - Y)" by (rule Y2) auto also have "\ \ density M ?f A - density M ?f Y" by (subst emeasure_Diff) auto also have "\ \ density M ?f A - 0" by (intro ennreal_minus_mono) auto also have "density M ?f A = b * M A + density M f A" by (simp add: emeasure_density field_simps mult_indicator_subset nn_integral_add nn_integral_cmult_indicator) also have "\ < N A" using f_less_N pos_A by (cases "density M f A"; cases "M A"; cases "N A") (auto simp: b_def ennreal_less_iff ennreal_minus divide_ennreal ennreal_numeral[symmetric] ennreal_plus[symmetric] ennreal_mult[symmetric] field_simps simp del: ennreal_numeral ennreal_plus) finally show False by simp qed then have "nn_integral M f < nn_integral M ?f'" using \0 < b\ \nn_integral M f \ top\ by (simp add: nn_integral_add nn_integral_cmult_indicator ennreal_zero_less_mult_iff zero_less_iff_neq_zero) moreover have "?f' \ G" unfolding G_def proof safe fix X assume [measurable]: "X \ sets M" have "(\\<^sup>+ x. ?f' x * indicator X x \M) = density M f (X - Y) + density M ?f (X \ Y)" by (auto simp add: emeasure_density nn_integral_add[symmetric] split: split_indicator intro!: nn_integral_cong) also have "\ \ N (X - Y) + N (X \ Y)" using G_D[OF \f \ G\] by (intro add_mono Y1) (auto simp: emeasure_density) also have "\ = N X" by (subst plus_emeasure) (auto intro!: arg_cong2[where f=emeasure]) finally show "(\\<^sup>+ x. ?f' x * indicator X x \M) \ N X" . qed simp then have "nn_integral M ?f' \ ?y" by (rule SUP_upper) ultimately show False by (simp add: int_f_eq_y) qed show ?thesis proof (intro bexI[of _ f] measure_eqI conjI antisym) fix A assume "A \ sets (density M f)" then show "emeasure (density M f) A \ emeasure N A" by (auto simp: emeasure_density intro!: G_D[OF \f \ G\]) next fix A assume A: "A \ sets (density M f)" then show "emeasure N A \ emeasure (density M f) A" using upper_bound by auto qed auto qed lemma (in finite_measure) split_space_into_finite_sets_and_rest: assumes ac: "absolutely_continuous M N" and sets_eq[simp]: "sets N = sets M" shows "\B::nat\'a set. disjoint_family B \ range B \ sets M \ (\i. N (B i) \ \) \ (\A\sets M. A \ (\i. B i) = {} \ (emeasure M A = 0 \ N A = 0) \ (emeasure M A > 0 \ N A = \))" proof - let ?Q = "{Q\sets M. N Q \ \}" let ?a = "SUP Q\?Q. emeasure M Q" have "{} \ ?Q" by auto then have Q_not_empty: "?Q \ {}" by blast have "?a \ emeasure M (space M)" using sets.sets_into_space by (auto intro!: SUP_least emeasure_mono) then have "?a \ \" using finite_emeasure_space by (auto simp: less_top[symmetric] top_unique simp del: SUP_eq_top_iff Sup_eq_top_iff) from ennreal_SUP_countable_SUP [OF Q_not_empty, of "emeasure M"] obtain Q'' where "range Q'' \ emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)" by auto then have "\i. \Q'. Q'' i = emeasure M Q' \ Q' \ ?Q" by auto from choice[OF this] obtain Q' where Q': "\i. Q'' i = emeasure M (Q' i)" "\i. Q' i \ ?Q" by auto then have a_Lim: "?a = (SUP i. emeasure M (Q' i))" using a by simp let ?O = "\n. \i\n. Q' i" have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\i. ?O i)" proof (rule SUP_emeasure_incseq[of ?O]) show "range ?O \ sets M" using Q' by auto show "incseq ?O" by (fastforce intro!: incseq_SucI) qed have Q'_sets[measurable]: "\i. Q' i \ sets M" using Q' by auto have O_sets: "\i. ?O i \ sets M" using Q' by auto then have O_in_G: "\i. ?O i \ ?Q" proof (safe del: notI) fix i have "Q' ` {..i} \ sets M" using Q' by auto then have "N (?O i) \ (\i\i. N (Q' i))" by (simp add: emeasure_subadditive_finite) also have "\ < \" using Q' by (simp add: less_top) finally show "N (?O i) \ \" by simp qed auto have O_mono: "\n. ?O n \ ?O (Suc n)" by fastforce have a_eq: "?a = emeasure M (\i. ?O i)" unfolding Union[symmetric] proof (rule antisym) show "?a \ (SUP i. emeasure M (?O i))" unfolding a_Lim using Q' by (auto intro!: SUP_mono emeasure_mono) show "(SUP i. emeasure M (?O i)) \ ?a" proof (safe intro!: Sup_mono, unfold bex_simps) fix i have *: "(\(Q' ` {..i})) = ?O i" by auto then show "\x. (x \ sets M \ N x \ \) \ emeasure M (\(Q' ` {..i})) \ emeasure M x" using O_in_G[of i] by (auto intro!: exI[of _ "?O i"]) qed qed let ?O_0 = "(\i. ?O i)" have "?O_0 \ sets M" using Q' by auto have "disjointed Q' i \ sets M" for i using sets.range_disjointed_sets[of Q' M] using Q'_sets by (auto simp: subset_eq) note Q_sets = this show ?thesis proof (intro bexI exI conjI ballI impI allI) show "disjoint_family (disjointed Q')" by (rule disjoint_family_disjointed) show "range (disjointed Q') \ sets M" using Q'_sets by (intro sets.range_disjointed_sets) auto { fix A assume A: "A \ sets M" "A \ (\i. disjointed Q' i) = {}" then have A1: "A \ (\i. Q' i) = {}" unfolding UN_disjointed_eq by auto show "emeasure M A = 0 \ N A = 0 \ 0 < emeasure M A \ N A = \" proof (rule disjCI, simp) assume *: "emeasure M A = 0 \ N A \ top" show "emeasure M A = 0 \ N A = 0" proof (cases "emeasure M A = 0") case True with ac A have "N A = 0" unfolding absolutely_continuous_def by auto with True show ?thesis by simp next case False with * have "N A \ \" by auto with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \ A)" using Q' A1 by (auto intro!: plus_emeasure simp: set_eq_iff) also have "\ = (SUP i. emeasure M (?O i \ A))" proof (rule SUP_emeasure_incseq[of "\i. ?O i \ A", symmetric, simplified]) show "range (\i. ?O i \ A) \ sets M" using \N A \ \\ O_sets A by auto qed (fastforce intro!: incseq_SucI) also have "\ \ ?a" proof (safe intro!: SUP_least) fix i have "?O i \ A \ ?Q" proof (safe del: notI) show "?O i \ A \ sets M" using O_sets A by auto from O_in_G[of i] have "N (?O i \ A) \ N (?O i) + N A" using emeasure_subadditive[of "?O i" N A] A O_sets by auto with O_in_G[of i] show "N (?O i \ A) \ \" using \N A \ \\ by (auto simp: top_unique) qed then show "emeasure M (?O i \ A) \ ?a" by (rule SUP_upper) qed finally have "emeasure M A = 0" unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure) with \emeasure M A \ 0\ show ?thesis by auto qed qed } { fix i have "N (disjointed Q' i) \ N (Q' i)" by (auto intro!: emeasure_mono simp: disjointed_def) then show "N (disjointed Q' i) \ \" using Q'(2)[of i] by (auto simp: top_unique) } qed qed proposition (in finite_measure) Radon_Nikodym_finite_measure_infinite: assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M" shows "\f\borel_measurable M. density M f = N" proof - from split_space_into_finite_sets_and_rest[OF assms] obtain Q :: "nat \ 'a set" where Q: "disjoint_family Q" "range Q \ sets M" and in_Q0: "\A. A \ sets M \ A \ (\i. Q i) = {} \ emeasure M A = 0 \ N A = 0 \ 0 < emeasure M A \ N A = \" and Q_fin: "\i. N (Q i) \ \" by force from Q have Q_sets: "\i. Q i \ sets M" by auto let ?N = "\i. density N (indicator (Q i))" and ?M = "\i. density M (indicator (Q i))" have "\i. \f\borel_measurable (?M i). density (?M i) f = ?N i" proof (intro allI finite_measure.Radon_Nikodym_finite_measure) fix i from Q show "finite_measure (?M i)" by (auto intro!: finite_measureI cong: nn_integral_cong simp add: emeasure_density subset_eq sets_eq) from Q have "emeasure (?N i) (space N) = emeasure N (Q i)" by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: nn_integral_cong) with Q_fin show "finite_measure (?N i)" by (auto intro!: finite_measureI) show "sets (?N i) = sets (?M i)" by (simp add: sets_eq) have [measurable]: "\A. A \ sets M \ A \ sets N" by (simp add: sets_eq) show "absolutely_continuous (?M i) (?N i)" using \absolutely_continuous M N\ \Q i \ sets M\ by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq intro!: absolutely_continuous_AE[OF sets_eq]) qed from choice[OF this[unfolded Bex_def]] obtain f where borel: "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" and f_density: "\i. density (?M i) (f i) = ?N i" by force { fix A i assume A: "A \ sets M" with Q borel have "(\\<^sup>+x. f i x * indicator (Q i \ A) x \M) = emeasure (density (?M i) (f i)) A" by (auto simp add: emeasure_density nn_integral_density subset_eq intro!: nn_integral_cong split: split_indicator) also have "\ = emeasure N (Q i \ A)" using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq) finally have "emeasure N (Q i \ A) = (\\<^sup>+x. f i x * indicator (Q i \ A) x \M)" .. } note integral_eq = this let ?f = "\x. (\i. f i x * indicator (Q i) x) + \ * indicator (space M - (\i. Q i)) x" show ?thesis proof (safe intro!: bexI[of _ ?f]) show "?f \ borel_measurable M" using borel Q_sets by (auto intro!: measurable_If) show "density M ?f = N" proof (rule measure_eqI) fix A assume "A \ sets (density M ?f)" then have "A \ sets M" by simp have Qi: "\i. Q i \ sets M" using Q by auto have [intro,simp]: "\i. (\x. f i x * indicator (Q i \ A) x) \ borel_measurable M" "\i. AE x in M. 0 \ f i x * indicator (Q i \ A) x" using borel Qi \A \ sets M\ by auto have "(\\<^sup>+x. ?f x * indicator A x \M) = (\\<^sup>+x. (\i. f i x * indicator (Q i \ A) x) + \ * indicator ((space M - (\i. Q i)) \ A) x \M)" using borel by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x. (\i. f i x * indicator (Q i \ A) x) \M) + \ * emeasure M ((space M - (\i. Q i)) \ A)" using borel Qi \A \ sets M\ by (subst nn_integral_add) (auto simp add: nn_integral_cmult_indicator sets.Int intro!: suminf_0_le) also have "\ = (\i. N (Q i \ A)) + \ * emeasure M ((space M - (\i. Q i)) \ A)" by (subst integral_eq[OF \A \ sets M\], subst nn_integral_suminf) auto finally have "(\\<^sup>+x. ?f x * indicator A x \M) = (\i. N (Q i \ A)) + \ * emeasure M ((space M - (\i. Q i)) \ A)" . moreover have "(\i. N (Q i \ A)) = N ((\i. Q i) \ A)" using Q Q_sets \A \ sets M\ by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq) moreover have "(space M - (\x. Q x)) \ A \ (\x. Q x) = {}" by auto then have "\ * emeasure M ((space M - (\i. Q i)) \ A) = N ((space M - (\i. Q i)) \ A)" using in_Q0[of "(space M - (\i. Q i)) \ A"] \A \ sets M\ Q by (auto simp: ennreal_top_mult) moreover have "(space M - (\i. Q i)) \ A \ sets M" "((\i. Q i) \ A) \ sets M" using Q_sets \A \ sets M\ by auto moreover have "((\i. Q i) \ A) \ ((space M - (\i. Q i)) \ A) = A" "((\i. Q i) \ A) \ ((space M - (\i. Q i)) \ A) = {}" using \A \ sets M\ sets.sets_into_space by auto ultimately have "N A = (\\<^sup>+x. ?f x * indicator A x \M)" using plus_emeasure[of "(\i. Q i) \ A" N "(space M - (\i. Q i)) \ A"] by (simp add: sets_eq) with \A \ sets M\ borel Q show "emeasure (density M ?f) A = N A" by (auto simp: subset_eq emeasure_density) qed (simp add: sets_eq) qed qed theorem (in sigma_finite_measure) Radon_Nikodym: assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M" shows "\f \ borel_measurable M. density M f = N" proof - from Ex_finite_integrable_function obtain h where finite: "integral\<^sup>N M h \ \" and borel: "h \ borel_measurable M" and nn: "\x. 0 \ h x" and pos: "\x. x \ space M \ 0 < h x" and "\x. x \ space M \ h x < \" by auto let ?T = "\A. (\\<^sup>+x. h x * indicator A x \M)" let ?MT = "density M h" from borel finite nn interpret T: finite_measure ?MT by (auto intro!: finite_measureI cong: nn_integral_cong simp: emeasure_density) have "absolutely_continuous ?MT N" "sets N = sets ?MT" proof (unfold absolutely_continuous_def, safe) fix A assume "A \ null_sets ?MT" with borel have "A \ sets M" "AE x in M. x \ A \ h x \ 0" by (auto simp add: null_sets_density_iff) with pos sets.sets_into_space have "AE x in M. x \ A" by (elim eventually_mono) (auto simp: not_le[symmetric]) then have "A \ null_sets M" using \A \ sets M\ by (simp add: AE_iff_null_sets) with ac show "A \ null_sets N" by (auto simp: absolutely_continuous_def) qed (auto simp add: sets_eq) from T.Radon_Nikodym_finite_measure_infinite[OF this] obtain f where f_borel: "f \ borel_measurable M" "density ?MT f = N" by auto with nn borel show ?thesis by (auto intro!: bexI[of _ "\x. h x * f x"] simp: density_density_eq) qed subsection \Uniqueness of densities\ lemma finite_density_unique: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" assumes pos: "AE x in M. 0 \ f x" "AE x in M. 0 \ g x" and fin: "integral\<^sup>N M f \ \" shows "density M f = density M g \ (AE x in M. f x = g x)" proof (intro iffI ballI) fix A assume eq: "AE x in M. f x = g x" with borel show "density M f = density M g" by (auto intro: density_cong) next let ?P = "\f A. \\<^sup>+ x. f x * indicator A x \M" assume "density M f = density M g" with borel have eq: "\A\sets M. ?P f A = ?P g A" by (simp add: emeasure_density[symmetric]) from this[THEN bspec, OF sets.top] fin have g_fin: "integral\<^sup>N M g \ \" by (simp cong: nn_integral_cong) { fix f g assume borel: "f \ borel_measurable M" "g \ borel_measurable M" and pos: "AE x in M. 0 \ f x" "AE x in M. 0 \ g x" and g_fin: "integral\<^sup>N M g \ \" and eq: "\A\sets M. ?P f A = ?P g A" let ?N = "{x\space M. g x < f x}" have N: "?N \ sets M" using borel by simp have "?P g ?N \ integral\<^sup>N M g" using pos by (intro nn_integral_mono_AE) (auto split: split_indicator) then have Pg_fin: "?P g ?N \ \" using g_fin by (auto simp: top_unique) have "?P (\x. (f x - g x)) ?N = (\\<^sup>+x. f x * indicator ?N x - g x * indicator ?N x \M)" by (auto intro!: nn_integral_cong simp: indicator_def) also have "\ = ?P f ?N - ?P g ?N" proof (rule nn_integral_diff) show "(\x. f x * indicator ?N x) \ borel_measurable M" "(\x. g x * indicator ?N x) \ borel_measurable M" using borel N by auto show "AE x in M. g x * indicator ?N x \ f x * indicator ?N x" using pos by (auto split: split_indicator) qed fact also have "\ = 0" unfolding eq[THEN bspec, OF N] using Pg_fin by auto finally have "AE x in M. f x \ g x" using pos borel nn_integral_PInf_AE[OF borel(2) g_fin] by (subst (asm) nn_integral_0_iff_AE) (auto split: split_indicator simp: not_less ennreal_minus_eq_0) } from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq show "AE x in M. f x = g x" by auto qed lemma (in finite_measure) density_unique_finite_measure: assumes borel: "f \ borel_measurable M" "f' \ borel_measurable M" assumes pos: "AE x in M. 0 \ f x" "AE x in M. 0 \ f' x" assumes f: "\A. A \ sets M \ (\\<^sup>+x. f x * indicator A x \M) = (\\<^sup>+x. f' x * indicator A x \M)" (is "\A. A \ sets M \ ?P f A = ?P f' A") shows "AE x in M. f x = f' x" proof - let ?D = "\f. density M f" let ?N = "\A. ?P f A" and ?N' = "\A. ?P f' A" let ?f = "\A x. f x * indicator A x" and ?f' = "\A x. f' x * indicator A x" have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M" using borel by (auto intro!: absolutely_continuousI_density) from split_space_into_finite_sets_and_rest[OF this] obtain Q :: "nat \ 'a set" where Q: "disjoint_family Q" "range Q \ sets M" and in_Q0: "\A. A \ sets M \ A \ (\i. Q i) = {} \ emeasure M A = 0 \ ?D f A = 0 \ 0 < emeasure M A \ ?D f A = \" and Q_fin: "\i. ?D f (Q i) \ \" by force with borel pos have in_Q0: "\A. A \ sets M \ A \ (\i. Q i) = {} \ emeasure M A = 0 \ ?N A = 0 \ 0 < emeasure M A \ ?N A = \" and Q_fin: "\i. ?N (Q i) \ \" by (auto simp: emeasure_density subset_eq) from Q have Q_sets[measurable]: "\i. Q i \ sets M" by auto let ?D = "{x\space M. f x \ f' x}" have "?D \ sets M" using borel by auto have *: "\i x A. \y::ennreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \ A) x" unfolding indicator_def by auto have "\i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos by (intro finite_density_unique[THEN iffD1] allI) (auto intro!: f measure_eqI simp: emeasure_density * subset_eq) moreover have "AE x in M. ?f (space M - (\i. Q i)) x = ?f' (space M - (\i. Q i)) x" proof (rule AE_I') { fix f :: "'a \ ennreal" assume borel: "f \ borel_measurable M" and eq: "\A. A \ sets M \ ?N A = (\\<^sup>+x. f x * indicator A x \M)" let ?A = "\i. (space M - (\i. Q i)) \ {x \ space M. f x < (i::nat)}" have "(\i. ?A i) \ null_sets M" proof (rule null_sets_UN) fix i ::nat have "?A i \ sets M" using borel by auto have "?N (?A i) \ (\\<^sup>+x. (i::ennreal) * indicator (?A i) x \M)" unfolding eq[OF \?A i \ sets M\] by (auto intro!: nn_integral_mono simp: indicator_def) also have "\ = i * emeasure M (?A i)" using \?A i \ sets M\ by (auto intro!: nn_integral_cmult_indicator) also have "\ < \" using emeasure_real[of "?A i"] by (auto simp: ennreal_mult_less_top of_nat_less_top) finally have "?N (?A i) \ \" by simp then show "?A i \ null_sets M" using in_Q0[OF \?A i \ sets M\] \?A i \ sets M\ by auto qed also have "(\i. ?A i) = (space M - (\i. Q i)) \ {x\space M. f x \ \}" by (auto simp: ennreal_Ex_less_of_nat less_top[symmetric]) finally have "(space M - (\i. Q i)) \ {x\space M. f x \ \} \ null_sets M" by simp } from this[OF borel(1) refl] this[OF borel(2) f] have "(space M - (\i. Q i)) \ {x\space M. f x \ \} \ null_sets M" "(space M - (\i. Q i)) \ {x\space M. f' x \ \} \ null_sets M" by simp_all then show "((space M - (\i. Q i)) \ {x\space M. f x \ \}) \ ((space M - (\i. Q i)) \ {x\space M. f' x \ \}) \ null_sets M" by (rule null_sets.Un) show "{x \ space M. ?f (space M - (\i. Q i)) x \ ?f' (space M - (\i. Q i)) x} \ ((space M - (\i. Q i)) \ {x\space M. f x \ \}) \ ((space M - (\i. Q i)) \ {x\space M. f' x \ \})" by (auto simp: indicator_def) qed moreover have "AE x in M. (?f (space M - (\i. Q i)) x = ?f' (space M - (\i. Q i)) x) \ (\i. ?f (Q i) x = ?f' (Q i) x) \ ?f (space M) x = ?f' (space M) x" by (auto simp: indicator_def) ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x" unfolding AE_all_countable[symmetric] by eventually_elim (auto split: if_split_asm simp: indicator_def) then show "AE x in M. f x = f' x" by auto qed proposition (in sigma_finite_measure) density_unique: assumes f: "f \ borel_measurable M" assumes f': "f' \ borel_measurable M" assumes density_eq: "density M f = density M f'" shows "AE x in M. f x = f' x" proof - obtain h where h_borel: "h \ borel_measurable M" and fin: "integral\<^sup>N M h \ \" and pos: "\x. x \ space M \ 0 < h x \ h x < \" "\x. 0 \ h x" using Ex_finite_integrable_function by auto then have h_nn: "AE x in M. 0 \ h x" by auto let ?H = "density M h" interpret h: finite_measure ?H using fin h_borel pos by (intro finite_measureI) (simp cong: nn_integral_cong emeasure_density add: fin) let ?fM = "density M f" let ?f'M = "density M f'" { fix A assume "A \ sets M" then have "{x \ space M. h x * indicator A x \ 0} = A" using pos(1) sets.sets_into_space by (force simp: indicator_def) then have "(\\<^sup>+x. h x * indicator A x \M) = 0 \ A \ null_sets M" using h_borel \A \ sets M\ h_nn by (subst nn_integral_0_iff) auto } note h_null_sets = this { fix A assume "A \ sets M" have "(\\<^sup>+x. f x * (h x * indicator A x) \M) = (\\<^sup>+x. h x * indicator A x \?fM)" using \A \ sets M\ h_borel h_nn f f' by (intro nn_integral_density[symmetric]) auto also have "\ = (\\<^sup>+x. h x * indicator A x \?f'M)" by (simp_all add: density_eq) also have "\ = (\\<^sup>+x. f' x * (h x * indicator A x) \M)" using \A \ sets M\ h_borel h_nn f f' by (intro nn_integral_density) auto finally have "(\\<^sup>+x. h x * (f x * indicator A x) \M) = (\\<^sup>+x. h x * (f' x * indicator A x) \M)" by (simp add: ac_simps) then have "(\\<^sup>+x. (f x * indicator A x) \?H) = (\\<^sup>+x. (f' x * indicator A x) \?H)" using \A \ sets M\ h_borel h_nn f f' by (subst (asm) (1 2) nn_integral_density[symmetric]) auto } then have "AE x in ?H. f x = f' x" using h_borel h_nn f f' by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M]) auto with AE_space[of M] pos show "AE x in M. f x = f' x" unfolding AE_density[OF h_borel] by auto qed lemma (in sigma_finite_measure) density_unique_iff: assumes f: "f \ borel_measurable M" and f': "f' \ borel_measurable M" shows "density M f = density M f' \ (AE x in M. f x = f' x)" using density_unique[OF assms] density_cong[OF f f'] by auto lemma sigma_finite_density_unique: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" and fin: "sigma_finite_measure (density M f)" shows "density M f = density M g \ (AE x in M. f x = g x)" proof assume "AE x in M. f x = g x" with borel show "density M f = density M g" by (auto intro: density_cong) next assume eq: "density M f = density M g" interpret f: sigma_finite_measure "density M f" by fact - from f.sigma_finite_incseq guess A . note cover = this - + from f.sigma_finite_incseq obtain A where cover: "range A \ sets (density M f)" + "\ (range A) = space (density M f)" + "\i. emeasure (density M f) (A i) \ \" + "incseq A" + by auto have "AE x in M. \i. x \ A i \ f x = g x" unfolding AE_all_countable proof fix i have "density (density M f) (indicator (A i)) = density (density M g) (indicator (A i))" unfolding eq .. moreover have "(\\<^sup>+x. f x * indicator (A i) x \M) \ \" using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq) ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x" using borel cover(1) by (intro finite_density_unique[THEN iffD1]) (auto simp: density_density_eq subset_eq) then show "AE x in M. x \ A i \ f x = g x" by auto qed with AE_space show "AE x in M. f x = g x" apply eventually_elim using cover(2)[symmetric] apply auto done qed lemma (in sigma_finite_measure) sigma_finite_iff_density_finite': assumes f: "f \ borel_measurable M" shows "sigma_finite_measure (density M f) \ (AE x in M. f x \ \)" (is "sigma_finite_measure ?N \ _") proof assume "sigma_finite_measure ?N" then interpret N: sigma_finite_measure ?N . from N.Ex_finite_integrable_function obtain h where h: "h \ borel_measurable M" "integral\<^sup>N ?N h \ \" and fin: "\x\space M. 0 < h x \ h x < \" by auto have "AE x in M. f x * h x \ \" proof (rule AE_I') have "integral\<^sup>N ?N h = (\\<^sup>+x. f x * h x \M)" using f h by (auto intro!: nn_integral_density) then have "(\\<^sup>+x. f x * h x \M) \ \" using h(2) by simp then show "(\x. f x * h x) -` {\} \ space M \ null_sets M" using f h(1) by (auto intro!: nn_integral_PInf[unfolded infinity_ennreal_def] borel_measurable_vimage) qed auto then show "AE x in M. f x \ \" using fin by (auto elim!: AE_Ball_mp simp: less_top ennreal_mult_less_top) next assume AE: "AE x in M. f x \ \" - from sigma_finite guess Q . note Q = this + from sigma_finite obtain Q :: "nat \ 'a set" + where Q: "range Q \ sets M" "\ (range Q) = space M" "\i. emeasure M (Q i) \ \" + by auto define A where "A i = f -` (case i of 0 \ {\} | Suc n \ {.. ennreal(of_nat (Suc n))}) \ space M" for i { fix i j have "A i \ Q j \ sets M" unfolding A_def using f Q apply (rule_tac sets.Int) by (cases i) (auto intro: measurable_sets[OF f(1)]) } note A_in_sets = this show "sigma_finite_measure ?N" proof (standard, intro exI conjI ballI) show "countable (range (\(i, j). A i \ Q j))" by auto show "range (\(i, j). A i \ Q j) \ sets (density M f)" using A_in_sets by auto next have "\(range (\(i, j). A i \ Q j)) = (\i j. A i \ Q j)" by auto also have "\ = (\i. A i) \ space M" using Q by auto also have "(\i. A i) = space M" proof safe fix x assume x: "x \ space M" show "x \ (\i. A i)" proof (cases "f x" rule: ennreal_cases) case top with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0]) next case (real r) with ennreal_Ex_less_of_nat[of "f x"] obtain n :: nat where "f x < n" by auto also have "n < (Suc n :: ennreal)" by simp finally show ?thesis using x real by (auto simp: A_def ennreal_of_nat_eq_real_of_nat intro!: exI[of _ "Suc n"]) qed qed (auto simp: A_def) finally show "\(range (\(i, j). A i \ Q j)) = space ?N" by simp next fix X assume "X \ range (\(i, j). A i \ Q j)" then obtain i j where [simp]:"X = A i \ Q j" by auto have "(\\<^sup>+x. f x * indicator (A i \ Q j) x \M) \ \" proof (cases i) case 0 have "AE x in M. f x * indicator (A i \ Q j) x = 0" using AE by (auto simp: A_def \i = 0\) from nn_integral_cong_AE[OF this] show ?thesis by simp next case (Suc n) then have "(\\<^sup>+x. f x * indicator (A i \ Q j) x \M) \ (\\<^sup>+x. (Suc n :: ennreal) * indicator (Q j) x \M)" by (auto intro!: nn_integral_mono simp: indicator_def A_def ennreal_of_nat_eq_real_of_nat) also have "\ = Suc n * emeasure M (Q j)" using Q by (auto intro!: nn_integral_cmult_indicator) also have "\ < \" using Q by (auto simp: ennreal_mult_less_top less_top of_nat_less_top) finally show ?thesis by simp qed then show "emeasure ?N X \ \" using A_in_sets Q f by (auto simp: emeasure_density) qed qed lemma (in sigma_finite_measure) sigma_finite_iff_density_finite: "f \ borel_measurable M \ sigma_finite_measure (density M f) \ (AE x in M. f x \ \)" by (subst sigma_finite_iff_density_finite') (auto simp: max_def intro!: measurable_If) subsection \Radon-Nikodym derivative\ definition\<^marker>\tag important\ RN_deriv :: "'a measure \ 'a measure \ 'a \ ennreal" where "RN_deriv M N = (if \f. f \ borel_measurable M \ density M f = N then SOME f. f \ borel_measurable M \ density M f = N else (\_. 0))" lemma RN_derivI: assumes "f \ borel_measurable M" "density M f = N" shows "density M (RN_deriv M N) = N" proof - have *: "\f. f \ borel_measurable M \ density M f = N" using assms by auto then have "density M (SOME f. f \ borel_measurable M \ density M f = N) = N" by (rule someI2_ex) auto with * show ?thesis by (auto simp: RN_deriv_def) qed lemma borel_measurable_RN_deriv[measurable]: "RN_deriv M N \ borel_measurable M" proof - { assume ex: "\f. f \ borel_measurable M \ density M f = N" have 1: "(SOME f. f \ borel_measurable M \ density M f = N) \ borel_measurable M" using ex by (rule someI2_ex) auto } from this show ?thesis by (auto simp: RN_deriv_def) qed lemma density_RN_deriv_density: assumes f: "f \ borel_measurable M" shows "density M (RN_deriv M (density M f)) = density M f" by (rule RN_derivI[OF f]) simp lemma (in sigma_finite_measure) density_RN_deriv: "absolutely_continuous M N \ sets N = sets M \ density M (RN_deriv M N) = N" by (metis RN_derivI Radon_Nikodym) lemma (in sigma_finite_measure) RN_deriv_nn_integral: assumes N: "absolutely_continuous M N" "sets N = sets M" and f: "f \ borel_measurable M" shows "integral\<^sup>N N f = (\\<^sup>+x. RN_deriv M N x * f x \M)" proof - have "integral\<^sup>N N f = integral\<^sup>N (density M (RN_deriv M N)) f" using N by (simp add: density_RN_deriv) also have "\ = (\\<^sup>+x. RN_deriv M N x * f x \M)" using f by (simp add: nn_integral_density) finally show ?thesis by simp qed lemma (in sigma_finite_measure) RN_deriv_unique: assumes f: "f \ borel_measurable M" and eq: "density M f = N" shows "AE x in M. f x = RN_deriv M N x" unfolding eq[symmetric] by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv density_RN_deriv_density[symmetric]) lemma RN_deriv_unique_sigma_finite: assumes f: "f \ borel_measurable M" and eq: "density M f = N" and fin: "sigma_finite_measure N" shows "AE x in M. f x = RN_deriv M N x" using fin unfolding eq[symmetric] by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv density_RN_deriv_density[symmetric]) lemma (in sigma_finite_measure) RN_deriv_distr: fixes T :: "'a \ 'b" assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" and inv: "\x\space M. T' (T x) = x" and ac[simp]: "absolutely_continuous (distr M M' T) (distr N M' T)" and N: "sets N = sets M" shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x" proof (rule RN_deriv_unique) have [simp]: "sets N = sets M" by fact note sets_eq_imp_space_eq[OF N, simp] have measurable_N[simp]: "\M'. measurable N M' = measurable M M'" by (auto simp: measurable_def) { fix A assume "A \ sets M" with inv T T' sets.sets_into_space[OF this] have "T -` T' -` A \ T -` space M' \ space M = A" by (auto simp: measurable_def) } note eq = this[simp] { fix A assume "A \ sets M" with inv T T' sets.sets_into_space[OF this] have "(T' \ T) -` A \ space M = A" by (auto simp: measurable_def) } note eq2 = this[simp] let ?M' = "distr M M' T" and ?N' = "distr N M' T" interpret M': sigma_finite_measure ?M' proof - from sigma_finite_countable guess F .. note F = this + from sigma_finite_countable obtain F + where F: "countable F \ F \ sets M \ \ F = space M \ (\a\F. emeasure M a \ \)" .. show "\A. countable A \ A \ sets (distr M M' T) \ \A = space (distr M M' T) \ (\a\A. emeasure (distr M M' T) a \ \)" proof (intro exI conjI ballI) show *: "(\A. T' -` A \ space ?M') ` F \ sets ?M'" using F T' by (auto simp: measurable_def) show "\((\A. T' -` A \ space ?M')`F) = space ?M'" using F T'[THEN measurable_space] by (auto simp: set_eq_iff) next fix X assume "X \ (\A. T' -` A \ space ?M')`F" then obtain A where [simp]: "X = T' -` A \ space ?M'" and "A \ F" by auto have "X \ sets M'" using F T' \A\F\ by auto moreover have Fi: "A \ sets M" using F \A\F\ by auto ultimately show "emeasure ?M' X \ \" using F T T' \A\F\ by (simp add: emeasure_distr) - qed (insert F, auto) + qed (use F in auto) qed have "(RN_deriv ?M' ?N') \ T \ borel_measurable M" using T ac by measurable then show "(\x. RN_deriv ?M' ?N' (T x)) \ borel_measurable M" by (simp add: comp_def) have "N = distr N M (T' \ T)" by (subst measure_of_of_measure[of N, symmetric]) (auto simp add: distr_def sets.sigma_sets_eq intro!: measure_of_eq sets.space_closed) also have "\ = distr (distr N M' T) M T'" using T T' by (simp add: distr_distr) also have "\ = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'" using ac by (simp add: M'.density_RN_deriv) also have "\ = density M (RN_deriv (distr M M' T) (distr N M' T) \ T)" by (simp add: distr_density_distr[OF T T', OF inv]) finally show "density M (\x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N" by (simp add: comp_def) qed lemma (in sigma_finite_measure) RN_deriv_finite: assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M" shows "AE x in M. RN_deriv M N x \ \" proof - interpret N: sigma_finite_measure N by fact from N show ?thesis using sigma_finite_iff_density_finite[OF borel_measurable_RN_deriv, of N] density_RN_deriv[OF ac] by simp qed lemma (in sigma_finite_measure) assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M" and f: "f \ borel_measurable M" shows RN_deriv_integrable: "integrable N f \ integrable M (\x. enn2real (RN_deriv M N x) * f x)" (is ?integrable) and RN_deriv_integral: "integral\<^sup>L N f = (\x. enn2real (RN_deriv M N x) * f x \M)" (is ?integral) proof - note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp] interpret N: sigma_finite_measure N by fact have eq: "density M (RN_deriv M N) = density M (\x. enn2real (RN_deriv M N x))" proof (rule density_cong) from RN_deriv_finite[OF assms(1,2,3)] show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))" by eventually_elim (auto simp: less_top) qed (insert ac, auto) show ?integrable apply (subst density_RN_deriv[OF ac, symmetric]) unfolding eq apply (intro integrable_real_density f AE_I2 enn2real_nonneg) apply (insert ac, auto) done show ?integral apply (subst density_RN_deriv[OF ac, symmetric]) unfolding eq apply (intro integral_real_density f AE_I2 enn2real_nonneg) apply (insert ac, auto) done qed proposition (in sigma_finite_measure) real_RN_deriv: assumes "finite_measure N" assumes ac: "absolutely_continuous M N" "sets N = sets M" obtains D where "D \ borel_measurable M" and "AE x in M. RN_deriv M N x = ennreal (D x)" and "AE x in N. 0 < D x" and "\x. 0 \ D x" proof interpret N: finite_measure N by fact note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac] let ?RN = "\t. {x \ space M. RN_deriv M N x = t}" show "(\x. enn2real (RN_deriv M N x)) \ borel_measurable M" using RN by auto have "N (?RN \) = (\\<^sup>+ x. RN_deriv M N x * indicator (?RN \) x \M)" using RN(1) by (subst RN(2)[symmetric]) (auto simp: emeasure_density) also have "\ = (\\<^sup>+ x. \ * indicator (?RN \) x \M)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = \ * emeasure M (?RN \)" using RN by (intro nn_integral_cmult_indicator) auto finally have eq: "N (?RN \) = \ * emeasure M (?RN \)" . moreover have "emeasure M (?RN \) = 0" proof (rule ccontr) assume "emeasure M {x \ space M. RN_deriv M N x = \} \ 0" then have "0 < emeasure M {x \ space M. RN_deriv M N x = \}" by (auto simp: zero_less_iff_neq_zero) with eq have "N (?RN \) = \" by (simp add: ennreal_mult_eq_top_iff) with N.emeasure_finite[of "?RN \"] RN show False by auto qed ultimately have "AE x in M. RN_deriv M N x < \" using RN by (intro AE_iff_measurable[THEN iffD2]) (auto simp: less_top[symmetric]) then show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))" by auto then have eq: "AE x in N. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))" using ac absolutely_continuous_AE by auto have "N (?RN 0) = (\\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \M)" by (subst RN(2)[symmetric]) (auto simp: emeasure_density) also have "\ = (\\<^sup>+ x. 0 \M)" by (intro nn_integral_cong) (auto simp: indicator_def) finally have "AE x in N. RN_deriv M N x \ 0" using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq) with eq show "AE x in N. 0 < enn2real (RN_deriv M N x)" by (auto simp: enn2real_positive_iff less_top[symmetric] zero_less_iff_neq_zero) qed (rule enn2real_nonneg) lemma (in sigma_finite_measure) RN_deriv_singleton: assumes ac: "absolutely_continuous M N" "sets N = sets M" and x: "{x} \ sets M" shows "N {x} = RN_deriv M N x * emeasure M {x}" proof - from \{x} \ sets M\ have "density M (RN_deriv M N) {x} = (\\<^sup>+w. RN_deriv M N x * indicator {x} w \M)" by (auto simp: indicator_def emeasure_density intro!: nn_integral_cong) with x density_RN_deriv[OF ac] show ?thesis by (auto simp: max_def) qed end diff --git a/src/HOL/Analysis/Regularity.thy b/src/HOL/Analysis/Regularity.thy --- a/src/HOL/Analysis/Regularity.thy +++ b/src/HOL/Analysis/Regularity.thy @@ -1,391 +1,392 @@ (* Title: HOL/Analysis/Regularity.thy Author: Fabian Immler, TU München *) section \Regularity of Measures\ theory Regularity (* FIX suggestion to rename e.g. RegularityMeasures and/ or move as this theory consists of 1 result only *) imports Measure_Space Borel_Space begin theorem fixes M::"'a::{second_countable_topology, complete_space} measure" assumes sb: "sets M = sets borel" assumes "emeasure M (space M) \ \" assumes "B \ sets borel" shows inner_regular: "emeasure M B = (SUP K \ {K. K \ B \ compact K}. emeasure M K)" (is "?inner B") and outer_regular: "emeasure M B = (INF U \ {U. B \ U \ open U}. emeasure M U)" (is "?outer B") proof - have Us: "UNIV = space M" by (metis assms(1) sets_eq_imp_space_eq space_borel) hence sU: "space M = UNIV" by simp interpret finite_measure M by rule fact have approx_inner: "\A. A \ sets M \ (\e. e > 0 \ \K. K \ A \ compact K \ emeasure M A \ emeasure M K + ennreal e) \ ?inner A" by (rule ennreal_approx_SUP) (force intro!: emeasure_mono simp: compact_imp_closed emeasure_eq_measure)+ have approx_outer: "\A. A \ sets M \ (\e. e > 0 \ \B. A \ B \ open B \ emeasure M B \ emeasure M A + ennreal e) \ ?outer A" by (rule ennreal_approx_INF) (force intro!: emeasure_mono simp: emeasure_eq_measure sb)+ - from countable_dense_setE guess X::"'a set" . note X = this + from countable_dense_setE obtain X :: "'a set" + where X: "countable X" "\Y :: 'a set. open Y \ Y \ {} \ \d\X. d \ Y" + by auto { fix r::real assume "r > 0" hence "\y. open (ball y r)" "\y. ball y r \ {}" by auto with X(2)[OF this] have x: "space M = (\x\X. cball x r)" by (auto simp add: sU) (metis dist_commute order_less_imp_le) let ?U = "\k. (\n\{0..k}. cball (from_nat_into X n) r)" have "(\k. emeasure M (\n\{0..k}. cball (from_nat_into X n) r)) \ M ?U" by (rule Lim_emeasure_incseq) (auto intro!: borel_closed bexI simp: incseq_def Us sb) also have "?U = space M" proof safe fix x from X(2)[OF open_ball[of x r]] \r > 0\ obtain d where d: "d\X" "d \ ball x r" by auto show "x \ ?U" using X(1) d by simp (auto intro!: exI [where x = "to_nat_on X d"] simp: dist_commute Bex_def) qed (simp add: sU) finally have "(\k. M (\n\{0..k}. cball (from_nat_into X n) r)) \ M (space M)" . } note M_space = this { fix e ::real and n :: nat assume "e > 0" "n > 0" hence "1/n > 0" "e * 2 powr - n > 0" by (auto) from M_space[OF \1/n>0\] have "(\k. measure M (\i\{0..k}. cball (from_nat_into X i) (1/real n))) \ measure M (space M)" unfolding emeasure_eq_measure by (auto) from metric_LIMSEQ_D[OF this \0 < e * 2 powr -n\] obtain k where "dist (measure M (\i\{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) < e * 2 powr -n" by auto hence "measure M (\i\{0..k}. cball (from_nat_into X i) (1/real n)) \ measure M (space M) - e * 2 powr -real n" by (auto simp: dist_real_def) hence "\k. measure M (\i\{0..k}. cball (from_nat_into X i) (1/real n)) \ measure M (space M) - e * 2 powr - real n" .. } note k=this hence "\e\{0<..}. \(n::nat)\{0<..}. \k. measure M (\i\{0..k}. cball (from_nat_into X i) (1/real n)) \ measure M (space M) - e * 2 powr - real n" by blast then obtain k where k: "\e\{0<..}. \n\{0<..}. measure M (space M) - e * 2 powr - real (n::nat) \ measure M (\i\{0..k e n}. cball (from_nat_into X i) (1 / n))" by metis hence k: "\e n. e > 0 \ n > 0 \ measure M (space M) - e * 2 powr - n \ measure M (\i\{0..k e n}. cball (from_nat_into X i) (1 / n))" unfolding Ball_def by blast have approx_space: "\K \ {K. K \ space M \ compact K}. emeasure M (space M) \ emeasure M K + ennreal e" (is "?thesis e") if "0 < e" for e :: real proof - define B where [abs_def]: "B n = (\i\{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n))" for n have "\n. closed (B n)" by (auto simp: B_def) hence [simp]: "\n. B n \ sets M" by (simp add: sb) from k[OF \e > 0\ zero_less_Suc] have "\n. measure M (space M) - measure M (B n) \ e * 2 powr - real (Suc n)" by (simp add: algebra_simps B_def finite_measure_compl) hence B_compl_le: "\n::nat. measure M (space M - B n) \ e * 2 powr - real (Suc n)" by (simp add: finite_measure_compl) define K where "K = (\n. B n)" from \closed (B _)\ have "closed K" by (auto simp: K_def) hence [simp]: "K \ sets M" by (simp add: sb) have "measure M (space M) - measure M K = measure M (space M - K)" by (simp add: finite_measure_compl) also have "\ = emeasure M (\n. space M - B n)" by (auto simp: K_def emeasure_eq_measure) also have "\ \ (\n. emeasure M (space M - B n))" by (rule emeasure_subadditive_countably) (auto simp: summable_def) also have "\ \ (\n. ennreal (e*2 powr - real (Suc n)))" using B_compl_le by (intro suminf_le) (simp_all add: emeasure_eq_measure ennreal_leI) also have "\ \ (\n. ennreal (e * (1 / 2) ^ Suc n))" by (simp add: powr_minus powr_realpow field_simps del: of_nat_Suc) also have "\ = ennreal e * (\n. ennreal ((1 / 2) ^ Suc n))" unfolding ennreal_power[symmetric] using \0 < e\ by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def ennreal_power[symmetric]) also have "\ = e" by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto finally have "measure M (space M) \ measure M K + e" using \0 < e\ by simp hence "emeasure M (space M) \ emeasure M K + e" using \0 < e\ by (simp add: emeasure_eq_measure flip: ennreal_plus) moreover have "compact K" unfolding compact_eq_totally_bounded proof safe show "complete K" using \closed K\ by (simp add: complete_eq_closed) fix e'::real assume "0 < e'" - from nat_approx_posE[OF this] guess n . note n = this + then obtain n where n: "1 / real (Suc n) < e'" by (rule nat_approx_posE) let ?k = "from_nat_into X ` {0..k e (Suc n)}" have "finite ?k" by simp moreover have "K \ (\x\?k. ball x e')" unfolding K_def B_def using n by force ultimately show "\k. finite k \ K \ (\x\k. ball x e')" by blast qed ultimately show ?thesis by (auto simp: sU) qed { fix A::"'a set" assume "closed A" hence "A \ sets borel" by (simp add: compact_imp_closed) hence [simp]: "A \ sets M" by (simp add: sb) have "?inner A" proof (rule approx_inner) fix e::real assume "e > 0" from approx_space[OF this] obtain K where K: "K \ space M" "compact K" "emeasure M (space M) \ emeasure M K + e" by (auto simp: emeasure_eq_measure) hence [simp]: "K \ sets M" by (simp add: sb compact_imp_closed) have "measure M A - measure M (A \ K) = measure M (A - A \ K)" by (subst finite_measure_Diff) auto also have "A - A \ K = A \ K - K" by auto also have "measure M \ = measure M (A \ K) - measure M K" by (subst finite_measure_Diff) auto also have "\ \ measure M (space M) - measure M K" by (simp add: emeasure_eq_measure sU sb finite_measure_mono) also have "\ \ e" using K \0 < e\ by (simp add: emeasure_eq_measure flip: ennreal_plus) finally have "emeasure M A \ emeasure M (A \ K) + ennreal e" using \0 by (simp add: emeasure_eq_measure algebra_simps flip: ennreal_plus) moreover have "A \ K \ A" "compact (A \ K)" using \closed A\ \compact K\ by auto ultimately show "\K \ A. compact K \ emeasure M A \ emeasure M K + ennreal e" by blast qed simp have "?outer A" proof cases assume "A \ {}" let ?G = "\d. {x. infdist x A < d}" { fix d have "?G d = (\x. infdist x A) -` {.." by (intro continuous_open_vimage) (auto intro!: continuous_infdist continuous_ident) finally have "open (?G d)" . } note open_G = this from in_closed_iff_infdist_zero[OF \closed A\ \A \ {}\] have "A = {x. infdist x A = 0}" by auto also have "\ = (\i. ?G (1/real (Suc i)))" proof (auto simp del: of_nat_Suc, rule ccontr) fix x assume "infdist x A \ 0" - hence pos: "infdist x A > 0" using infdist_nonneg[of x A] by simp - from nat_approx_posE[OF this] guess n . - moreover + then have pos: "infdist x A > 0" using infdist_nonneg[of x A] by simp + then obtain n where n: "1 / real (Suc n) < infdist x A" by (rule nat_approx_posE) assume "\i. infdist x A < 1 / real (Suc i)" - hence "infdist x A < 1 / real (Suc n)" by auto - ultimately show False by simp + then have "infdist x A < 1 / real (Suc n)" by auto + with n show False by simp qed also have "M \ = (INF n. emeasure M (?G (1 / real (Suc n))))" proof (rule INF_emeasure_decseq[symmetric], safe) fix i::nat from open_G[of "1 / real (Suc i)"] show "?G (1 / real (Suc i)) \ sets M" by (simp add: sb borel_open) next show "decseq (\i. {x. infdist x A < 1 / real (Suc i)})" by (auto intro: less_trans intro!: divide_strict_left_mono simp: decseq_def le_eq_less_or_eq) qed simp finally have "emeasure M A = (INF n. emeasure M {x. infdist x A < 1 / real (Suc n)})" . moreover have "\ \ (INF U\{U. A \ U \ open U}. emeasure M U)" proof (intro INF_mono) fix m have "?G (1 / real (Suc m)) \ {U. A \ U \ open U}" using open_G by auto moreover have "M (?G (1 / real (Suc m))) \ M (?G (1 / real (Suc m)))" by simp ultimately show "\U\{U. A \ U \ open U}. emeasure M U \ emeasure M {x. infdist x A < 1 / real (Suc m)}" by blast qed moreover have "emeasure M A \ (INF U\{U. A \ U \ open U}. emeasure M U)" by (rule INF_greatest) (auto intro!: emeasure_mono simp: sb) ultimately show ?thesis by simp qed (auto intro!: INF_eqI) note \?inner A\ \?outer A\ } note closed_in_D = this from \B \ sets borel\ have "Int_stable (Collect closed)" "Collect closed \ Pow UNIV" "B \ sigma_sets UNIV (Collect closed)" by (auto simp: Int_stable_def borel_eq_closed) then show "?inner B" "?outer B" proof (induct B rule: sigma_sets_induct_disjoint) case empty { case 1 show ?case by (intro SUP_eqI[symmetric]) auto } { case 2 show ?case by (intro INF_eqI[symmetric]) (auto elim!: meta_allE[of _ "{}"]) } next case (basic B) { case 1 from basic closed_in_D show ?case by auto } { case 2 from basic closed_in_D show ?case by auto } next case (compl B) note inner = compl(2) and outer = compl(3) from compl have [simp]: "B \ sets M" by (auto simp: sb borel_eq_closed) case 2 have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl) also have "\ = (INF K\{K. K \ B \ compact K}. M (space M) - M K)" by (subst ennreal_SUP_const_minus) (auto simp: less_top[symmetric] inner) also have "\ = (INF U\{U. U \ B \ compact U}. M (space M - U))" by (auto simp add: emeasure_compl sb compact_imp_closed) also have "\ \ (INF U\{U. U \ B \ closed U}. M (space M - U))" by (rule INF_superset_mono) (auto simp add: compact_imp_closed) also have "(INF U\{U. U \ B \ closed U}. M (space M - U)) = (INF U\{U. space M - B \ U \ open U}. emeasure M U)" apply (rule arg_cong [of _ _ Inf]) using sU apply (auto simp add: image_iff) apply (rule exI [of _ "UNIV - y" for y]) apply safe apply (auto simp add: double_diff) done finally have "(INF U\{U. space M - B \ U \ open U}. emeasure M U) \ emeasure M (space M - B)" . moreover have "(INF U\{U. space M - B \ U \ open U}. emeasure M U) \ emeasure M (space M - B)" by (auto simp: sb sU intro!: INF_greatest emeasure_mono) ultimately show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb]) case 1 have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl) also have "\ = (SUP U\ {U. B \ U \ open U}. M (space M) - M U)" unfolding outer by (subst ennreal_INF_const_minus) auto also have "\ = (SUP U\{U. B \ U \ open U}. M (space M - U))" by (auto simp add: emeasure_compl sb compact_imp_closed) also have "\ = (SUP K\{K. K \ space M - B \ closed K}. emeasure M K)" unfolding SUP_image [of _ "\u. space M - u" _, symmetric, unfolded comp_def] apply (rule arg_cong [of _ _ Sup]) using sU apply (auto intro!: imageI) done also have "\ = (SUP K\{K. K \ space M - B \ compact K}. emeasure M K)" proof (safe intro!: antisym SUP_least) fix K assume "closed K" "K \ space M - B" from closed_in_D[OF \closed K\] have K_inner: "emeasure M K = (SUP K\{Ka. Ka \ K \ compact Ka}. emeasure M K)" by simp show "emeasure M K \ (SUP K\{K. K \ space M - B \ compact K}. emeasure M K)" unfolding K_inner using \K \ space M - B\ by (auto intro!: SUP_upper SUP_least) qed (fastforce intro!: SUP_least SUP_upper simp: compact_imp_closed) finally show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb]) next case (union D) then have "range D \ sets M" by (auto simp: sb borel_eq_closed) with union have M[symmetric]: "(\i. M (D i)) = M (\i. D i)" by (intro suminf_emeasure) also have "(\n. \i (\i. M (D i))" by (intro summable_LIMSEQ) auto finally have measure_LIMSEQ: "(\n. \i measure M (\i. D i)" by (simp add: emeasure_eq_measure sum_nonneg) have "(\i. D i) \ sets M" using \range D \ sets M\ by auto case 1 show ?case proof (rule approx_inner) fix e::real assume "e > 0" with measure_LIMSEQ have "\no. \n\no. \(\ix. D x)\ < e/2" by (auto simp: lim_sequentially dist_real_def simp del: less_divide_eq_numeral1) hence "\n0. \(\ix. D x)\ < e/2" by auto then obtain n0 where n0: "\(\ii. D i)\ < e/2" unfolding choice_iff by blast have "ennreal (\ii \ (\i. M (D i))" by (rule sum_le_suminf) auto also have "\ = M (\i. D i)" by (simp add: M) also have "\ = measure M (\i. D i)" by (simp add: emeasure_eq_measure) finally have n0: "measure M (\i. D i) - (\ii. \K. K \ D i \ compact K \ emeasure M (D i) \ emeasure M K + e/(2*Suc n0)" proof fix i from \0 < e\ have "0 < e/(2*Suc n0)" by simp have "emeasure M (D i) = (SUP K\{K. K \ (D i) \ compact K}. emeasure M K)" using union by blast from SUP_approx_ennreal[OF \0 < e/(2*Suc n0)\ _ this] show "\K. K \ D i \ compact K \ emeasure M (D i) \ emeasure M K + e/(2*Suc n0)" by (auto simp: emeasure_eq_measure intro: less_imp_le compact_empty) qed then obtain K where K: "\i. K i \ D i" "\i. compact (K i)" "\i. emeasure M (D i) \ emeasure M (K i) + e/(2*Suc n0)" unfolding choice_iff by blast let ?K = "\i\{..disjoint_family D\ unfolding disjoint_family_on_def by blast hence mK: "measure M ?K = (\ii. D i) < (\ii (\i0 < e\ by (auto intro: sum_mono simp: emeasure_eq_measure simp flip: ennreal_plus) also have "\ = (\ii \ (\i0 < e\ by (auto simp: field_simps intro!: mult_left_mono) finally have "measure M (\i. D i) < (\ii. D i) < M ?K + e" using \0 by (auto simp: mK emeasure_eq_measure sum_nonneg ennreal_less_iff simp flip: ennreal_plus) moreover have "?K \ (\i. D i)" using K by auto moreover have "compact ?K" using K by auto ultimately have "?K\(\i. D i) \ compact ?K \ emeasure M (\i. D i) \ emeasure M ?K + ennreal e" by simp thus "\K\\i. D i. compact K \ emeasure M (\i. D i) \ emeasure M K + ennreal e" .. qed fact case 2 show ?case proof (rule approx_outer[OF \(\i. D i) \ sets M\]) fix e::real assume "e > 0" have "\i::nat. \U. D i \ U \ open U \ e/(2 powr Suc i) > emeasure M U - emeasure M (D i)" proof fix i::nat from \0 < e\ have "0 < e/(2 powr Suc i)" by simp have "emeasure M (D i) = (INF U\{U. (D i) \ U \ open U}. emeasure M U)" using union by blast from INF_approx_ennreal[OF \0 < e/(2 powr Suc i)\ this] show "\U. D i \ U \ open U \ e/(2 powr Suc i) > emeasure M U - emeasure M (D i)" using \0 by (auto simp: emeasure_eq_measure sum_nonneg ennreal_less_iff ennreal_minus finite_measure_mono sb simp flip: ennreal_plus) qed then obtain U where U: "\i. D i \ U i" "\i. open (U i)" "\i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)" unfolding choice_iff by blast let ?U = "\i. U i" have "ennreal (measure M ?U - measure M (\i. D i)) = M ?U - M (\i. D i)" using U(1,2) by (subst ennreal_minus[symmetric]) (auto intro!: finite_measure_mono simp: sb emeasure_eq_measure) also have "\ = M (?U - (\i. D i))" using U \(\i. D i) \ sets M\ by (subst emeasure_Diff) (auto simp: sb) also have "\ \ M (\i. U i - D i)" using U \range D \ sets M\ by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff) also have "\ \ (\i. M (U i - D i))" using U \range D \ sets M\ by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb) also have "\ \ (\i. ennreal e/(2 powr Suc i))" using U \range D \ sets M\ using \0 by (intro suminf_le, subst emeasure_Diff) (auto simp: emeasure_Diff emeasure_eq_measure sb ennreal_minus finite_measure_mono divide_ennreal ennreal_less_iff intro: less_imp_le) also have "\ \ (\n. ennreal (e * (1 / 2) ^ Suc n))" using \0 by (simp add: powr_minus powr_realpow field_simps divide_ennreal del: of_nat_Suc) also have "\ = ennreal e * (\n. ennreal ((1 / 2) ^ Suc n))" unfolding ennreal_power[symmetric] using \0 < e\ by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def ennreal_power[symmetric]) also have "\ = ennreal e" by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto finally have "emeasure M ?U \ emeasure M (\i. D i) + ennreal e" using \0 by (simp add: emeasure_eq_measure flip: ennreal_plus) moreover have "(\i. D i) \ ?U" using U by auto moreover have "open ?U" using U by auto ultimately have "(\i. D i) \ ?U \ open ?U \ emeasure M ?U \ emeasure M (\i. D i) + ennreal e" by simp thus "\B. (\i. D i) \ B \ open B \ emeasure M B \ emeasure M (\i. D i) + ennreal e" .. qed qed qed end diff --git a/src/HOL/Analysis/Set_Integral.thy b/src/HOL/Analysis/Set_Integral.thy --- a/src/HOL/Analysis/Set_Integral.thy +++ b/src/HOL/Analysis/Set_Integral.thy @@ -1,1163 +1,1163 @@ (* Title: HOL/Analysis/Set_Integral.thy Author: Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU) Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr Notation and useful facts for working with integrals over a set. TODO: keep all these? Need unicode translations as well. *) theory Set_Integral imports Radon_Nikodym begin (* Notation *) definition\<^marker>\tag important\ "set_borel_measurable M A f \ (\x. indicator A x *\<^sub>R f x) \ borel_measurable M" definition\<^marker>\tag important\ "set_integrable M A f \ integrable M (\x. indicator A x *\<^sub>R f x)" definition\<^marker>\tag important\ "set_lebesgue_integral M A f \ lebesgue_integral M (\x. indicator A x *\<^sub>R f x)" syntax "_ascii_set_lebesgue_integral" :: "pttrn \ 'a set \ 'a measure \ real \ real" ("(4LINT (_):(_)/|(_)./ _)" [0,60,110,61] 60) translations "LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\x. f)" (* Notation for integration wrt lebesgue measure on the reals: LBINT x. f LBINT x : A. f TODO: keep all these? Need unicode. *) syntax "_lebesgue_borel_integral" :: "pttrn \ real \ real" ("(2LBINT _./ _)" [0,60] 60) syntax "_set_lebesgue_borel_integral" :: "pttrn \ real set \ real \ real" ("(3LBINT _:_./ _)" [0,60,61] 60) (* Basic properties *) (* lemma indicator_abs_eq: "\A x. \indicator A x\ = ((indicator A x) :: real)" by (auto simp add: indicator_def) *) lemma set_integrable_cong: assumes "M = M'" "A = A'" "\x. x \ A \ f x = f' x" shows "set_integrable M A f = set_integrable M' A' f'" proof - have "(\x. indicator A x *\<^sub>R f x) = (\x. indicator A' x *\<^sub>R f' x)" using assms by (auto simp: indicator_def of_bool_def) thus ?thesis by (simp add: set_integrable_def assms) qed lemma set_borel_measurable_sets: fixes f :: "_ \ _::real_normed_vector" assumes "set_borel_measurable M X f" "B \ sets borel" "X \ sets M" shows "f -` B \ X \ sets M" proof - have "f \ borel_measurable (restrict_space M X)" using assms unfolding set_borel_measurable_def by (subst borel_measurable_restrict_space_iff) auto then have "f -` B \ space (restrict_space M X) \ sets (restrict_space M X)" by (rule measurable_sets) fact with \X \ sets M\ show ?thesis by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space) qed lemma set_lebesgue_integral_zero [simp]: "set_lebesgue_integral M A (\x. 0) = 0" by (auto simp: set_lebesgue_integral_def) lemma set_lebesgue_integral_cong: assumes "A \ sets M" and "\x. x \ A \ f x = g x" shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)" unfolding set_lebesgue_integral_def using assms by (metis indicator_simps(2) real_vector.scale_zero_left) lemma set_lebesgue_integral_cong_AE: assumes [measurable]: "A \ sets M" "f \ borel_measurable M" "g \ borel_measurable M" assumes "AE x \ A in M. f x = g x" shows "LINT x:A|M. f x = LINT x:A|M. g x" proof- have "AE x in M. indicator A x *\<^sub>R f x = indicator A x *\<^sub>R g x" using assms by auto thus ?thesis unfolding set_lebesgue_integral_def by (intro integral_cong_AE) auto qed lemma set_integrable_cong_AE: "f \ borel_measurable M \ g \ borel_measurable M \ AE x \ A in M. f x = g x \ A \ sets M \ set_integrable M A f = set_integrable M A g" unfolding set_integrable_def by (rule integrable_cong_AE) auto lemma set_integrable_subset: fixes M A B and f :: "_ \ _ :: {banach, second_countable_topology}" assumes "set_integrable M A f" "B \ sets M" "B \ A" shows "set_integrable M B f" proof - have "set_integrable M B (\x. indicator A x *\<^sub>R f x)" using assms integrable_mult_indicator set_integrable_def by blast with \B \ A\ show ?thesis unfolding set_integrable_def by (simp add: indicator_inter_arith[symmetric] Int_absorb2) qed lemma set_integrable_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f: "set_integrable M S f" and T: "T \ sets (restrict_space M S)" shows "set_integrable M T f" proof - obtain T' where T_eq: "T = S \ T'" and "T' \ sets M" using T by (auto simp: sets_restrict_space) have \integrable M (\x. indicator T' x *\<^sub>R (indicator S x *\<^sub>R f x))\ using \T' \ sets M\ f integrable_mult_indicator set_integrable_def by blast then show ?thesis unfolding set_integrable_def unfolding T_eq indicator_inter_arith by (simp add: ac_simps) qed (* TODO: integral_cmul_indicator should be named set_integral_const *) (* TODO: borel_integrable_atLeastAtMost should be named something like set_integrable_Icc_isCont *) lemma set_integral_scaleR_right [simp]: "LINT t:A|M. a *\<^sub>R f t = a *\<^sub>R (LINT t:A|M. f t)" unfolding set_lebesgue_integral_def by (subst integral_scaleR_right[symmetric]) (auto intro!: Bochner_Integration.integral_cong) lemma set_integral_mult_right [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "LINT t:A|M. a * f t = a * (LINT t:A|M. f t)" unfolding set_lebesgue_integral_def by (subst integral_mult_right_zero[symmetric]) auto lemma set_integral_mult_left [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "LINT t:A|M. f t * a = (LINT t:A|M. f t) * a" unfolding set_lebesgue_integral_def by (subst integral_mult_left_zero[symmetric]) auto lemma set_integral_divide_zero [simp]: fixes a :: "'a::{real_normed_field, field, second_countable_topology}" shows "LINT t:A|M. f t / a = (LINT t:A|M. f t) / a" unfolding set_lebesgue_integral_def by (subst integral_divide_zero[symmetric], intro Bochner_Integration.integral_cong) (auto split: split_indicator) lemma set_integrable_scaleR_right [simp, intro]: shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. a *\<^sub>R f t)" unfolding set_integrable_def unfolding scaleR_left_commute by (rule integrable_scaleR_right) lemma set_integrable_scaleR_left [simp, intro]: fixes a :: "_ :: {banach, second_countable_topology}" shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. f t *\<^sub>R a)" unfolding set_integrable_def using integrable_scaleR_left[of a M "\x. indicator A x *\<^sub>R f x"] by simp lemma set_integrable_mult_right [simp, intro]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. a * f t)" unfolding set_integrable_def using integrable_mult_right[of a M "\x. indicator A x *\<^sub>R f x"] by simp lemma set_integrable_mult_right_iff [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" assumes "a \ 0" shows "set_integrable M A (\t. a * f t) \ set_integrable M A f" proof assume "set_integrable M A (\t. a * f t)" then have "set_integrable M A (\t. 1/a * (a * f t))" using set_integrable_mult_right by blast then show "set_integrable M A f" using assms by auto qed auto lemma set_integrable_mult_left [simp, intro]: fixes a :: "'a::{real_normed_field, second_countable_topology}" shows "(a \ 0 \ set_integrable M A f) \ set_integrable M A (\t. f t * a)" unfolding set_integrable_def using integrable_mult_left[of a M "\x. indicator A x *\<^sub>R f x"] by simp lemma set_integrable_mult_left_iff [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" assumes "a \ 0" shows "set_integrable M A (\t. f t * a) \ set_integrable M A f" using assms by (subst set_integrable_mult_right_iff [symmetric]) (auto simp: mult.commute) lemma set_integrable_divide [simp, intro]: fixes a :: "'a::{real_normed_field, field, second_countable_topology}" assumes "a \ 0 \ set_integrable M A f" shows "set_integrable M A (\t. f t / a)" proof - have "integrable M (\x. indicator A x *\<^sub>R f x / a)" using assms unfolding set_integrable_def by (rule integrable_divide_zero) also have "(\x. indicator A x *\<^sub>R f x / a) = (\x. indicator A x *\<^sub>R (f x / a))" by (auto split: split_indicator) finally show ?thesis unfolding set_integrable_def . qed lemma set_integrable_mult_divide_iff [simp]: fixes a :: "'a::{real_normed_field, second_countable_topology}" assumes "a \ 0" shows "set_integrable M A (\t. f t / a) \ set_integrable M A f" by (simp add: divide_inverse assms) lemma set_integral_add [simp, intro]: fixes f g :: "_ \ _ :: {banach, second_countable_topology}" assumes "set_integrable M A f" "set_integrable M A g" shows "set_integrable M A (\x. f x + g x)" and "LINT x:A|M. f x + g x = (LINT x:A|M. f x) + (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_add_right) lemma set_integral_diff [simp, intro]: assumes "set_integrable M A f" "set_integrable M A g" shows "set_integrable M A (\x. f x - g x)" and "LINT x:A|M. f x - g x = (LINT x:A|M. f x) - (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_diff_right) lemma set_integral_uminus: "set_integrable M A f \ LINT x:A|M. - f x = - (LINT x:A|M. f x)" unfolding set_integrable_def set_lebesgue_integral_def by (subst integral_minus[symmetric]) simp_all lemma set_integral_complex_of_real: "LINT x:A|M. complex_of_real (f x) = of_real (LINT x:A|M. f x)" unfolding set_lebesgue_integral_def by (subst integral_complex_of_real[symmetric]) (auto intro!: Bochner_Integration.integral_cong split: split_indicator) lemma set_integral_mono: fixes f g :: "_ \ real" assumes "set_integrable M A f" "set_integrable M A g" "\x. x \ A \ f x \ g x" shows "(LINT x:A|M. f x) \ (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (auto intro: integral_mono split: split_indicator) lemma set_integral_mono_AE: fixes f g :: "_ \ real" assumes "set_integrable M A f" "set_integrable M A g" "AE x \ A in M. f x \ g x" shows "(LINT x:A|M. f x) \ (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (auto intro: integral_mono_AE split: split_indicator) lemma set_integrable_abs: "set_integrable M A f \ set_integrable M A (\x. \f x\ :: real)" using integrable_abs[of M "\x. f x * indicator A x"]unfolding set_integrable_def by (simp add: abs_mult ac_simps) lemma set_integrable_abs_iff: fixes f :: "_ \ real" shows "set_borel_measurable M A f \ set_integrable M A (\x. \f x\) = set_integrable M A f" unfolding set_integrable_def set_borel_measurable_def by (subst (2) integrable_abs_iff[symmetric]) (simp_all add: abs_mult ac_simps) lemma set_integrable_abs_iff': fixes f :: "_ \ real" shows "f \ borel_measurable M \ A \ sets M \ set_integrable M A (\x. \f x\) = set_integrable M A f" by (simp add: set_borel_measurable_def set_integrable_abs_iff) lemma set_integrable_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "countable X" assumes diff: "(A - B) \ (B - A) \ X" assumes "\x. x \ X \ emeasure M {x} = 0" "\x. x \ X \ {x} \ sets M" shows "set_integrable M A f \ set_integrable M B f" unfolding set_integrable_def proof (rule integrable_discrete_difference[where X=X]) show "\x. x \ space M \ x \ X \ indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x" using diff by (auto split: split_indicator) qed fact+ lemma set_integral_discrete_difference: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "countable X" assumes diff: "(A - B) \ (B - A) \ X" assumes "\x. x \ X \ emeasure M {x} = 0" "\x. x \ X \ {x} \ sets M" shows "set_lebesgue_integral M A f = set_lebesgue_integral M B f" unfolding set_lebesgue_integral_def proof (rule integral_discrete_difference[where X=X]) show "\x. x \ space M \ x \ X \ indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x" using diff by (auto split: split_indicator) qed fact+ lemma set_integrable_Un: fixes f g :: "_ \ _ :: {banach, second_countable_topology}" assumes f_A: "set_integrable M A f" and f_B: "set_integrable M B f" and [measurable]: "A \ sets M" "B \ sets M" shows "set_integrable M (A \ B) f" proof - have "set_integrable M (A - B) f" using f_A by (rule set_integrable_subset) auto with f_B have "integrable M (\x. indicator (A - B) x *\<^sub>R f x + indicator B x *\<^sub>R f x)" unfolding set_integrable_def using integrable_add by blast then show ?thesis unfolding set_integrable_def by (rule integrable_cong[THEN iffD1, rotated 2]) (auto split: split_indicator) qed lemma set_integrable_empty [simp]: "set_integrable M {} f" by (auto simp: set_integrable_def) lemma set_integrable_UN: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" "\i. i\I \ set_integrable M (A i) f" "\i. i\I \ A i \ sets M" shows "set_integrable M (\i\I. A i) f" using assms by (induct I) (auto simp: set_integrable_Un sets.finite_UN) lemma set_integral_Un: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "A \ B = {}" and "set_integrable M A f" and "set_integrable M B f" shows "LINT x:A\B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (auto simp add: indicator_union_arith indicator_inter_arith[symmetric] scaleR_add_left) lemma set_integral_cong_set: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "set_borel_measurable M A f" "set_borel_measurable M B f" and ae: "AE x in M. x \ A \ x \ B" shows "LINT x:B|M. f x = LINT x:A|M. f x" unfolding set_lebesgue_integral_def proof (rule integral_cong_AE) show "AE x in M. indicator B x *\<^sub>R f x = indicator A x *\<^sub>R f x" using ae by (auto simp: subset_eq split: split_indicator) qed (use assms in \auto simp: set_borel_measurable_def\) proposition set_borel_measurable_subset: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes [measurable]: "set_borel_measurable M A f" "B \ sets M" and "B \ A" shows "set_borel_measurable M B f" proof- have "set_borel_measurable M B (\x. indicator A x *\<^sub>R f x)" using assms unfolding set_borel_measurable_def using borel_measurable_indicator borel_measurable_scaleR by blast moreover have "(\x. indicator B x *\<^sub>R indicator A x *\<^sub>R f x) = (\x. indicator B x *\<^sub>R f x)" using \B \ A\ by (auto simp: fun_eq_iff split: split_indicator) ultimately show ?thesis unfolding set_borel_measurable_def by simp qed lemma set_integral_Un_AE: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes ae: "AE x in M. \ (x \ A \ x \ B)" and [measurable]: "A \ sets M" "B \ sets M" and "set_integrable M A f" and "set_integrable M B f" shows "LINT x:A\B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)" proof - have f: "set_integrable M (A \ B) f" by (intro set_integrable_Un assms) then have f': "set_borel_measurable M (A \ B) f" using integrable_iff_bounded set_borel_measurable_def set_integrable_def by blast have "LINT x:A\B|M. f x = LINT x:(A - A \ B) \ (B - A \ B)|M. f x" proof (rule set_integral_cong_set) show "AE x in M. (x \ A - A \ B \ (B - A \ B)) = (x \ A \ B)" using ae by auto show "set_borel_measurable M (A - A \ B \ (B - A \ B)) f" using f' by (rule set_borel_measurable_subset) auto qed fact also have "\ = (LINT x:(A - A \ B)|M. f x) + (LINT x:(B - A \ B)|M. f x)" by (auto intro!: set_integral_Un set_integrable_subset[OF f]) also have "\ = (LINT x:A|M. f x) + (LINT x:B|M. f x)" using ae by (intro arg_cong2[where f="(+)"] set_integral_cong_set) (auto intro!: set_borel_measurable_subset[OF f']) finally show ?thesis . qed lemma set_integral_finite_Union: fixes f :: "_ \ _ :: {banach, second_countable_topology}" assumes "finite I" "disjoint_family_on A I" and "\i. i \ I \ set_integrable M (A i) f" "\i. i \ I \ A i \ sets M" shows "(LINT x:(\i\I. A i)|M. f x) = (\i\I. LINT x:A i|M. f x)" using assms proof induction case (insert x F) then have "A x \ \(A ` F) = {}" by (meson disjoint_family_on_insert) with insert show ?case by (simp add: set_integral_Un set_integrable_Un set_integrable_UN disjoint_family_on_insert) qed (simp add: set_lebesgue_integral_def) (* TODO: find a better name? *) lemma pos_integrable_to_top: fixes l::real assumes "\i. A i \ sets M" "mono A" assumes nneg: "\x i. x \ A i \ 0 \ f x" and intgbl: "\i::nat. set_integrable M (A i) f" and lim: "(\i::nat. LINT x:A i|M. f x) \ l" shows "set_integrable M (\i. A i) f" unfolding set_integrable_def apply (rule integrable_monotone_convergence[where f = "\i::nat. \x. indicator (A i) x *\<^sub>R f x" and x = l]) apply (rule intgbl [unfolded set_integrable_def]) prefer 3 apply (rule lim [unfolded set_lebesgue_integral_def]) apply (rule AE_I2) using \mono A\ apply (auto simp: mono_def nneg split: split_indicator) [] proof (rule AE_I2) { fix x assume "x \ space M" show "(\i. indicator (A i) x *\<^sub>R f x) \ indicator (\i. A i) x *\<^sub>R f x" proof cases assume "\i. x \ A i" - then guess i .. + then obtain i where "x \ A i" .. then have *: "eventually (\i. x \ A i) sequentially" using \x \ A i\ \mono A\ by (auto simp: eventually_sequentially mono_def) show ?thesis apply (intro tendsto_eventually) using * apply eventually_elim apply (auto split: split_indicator) done qed auto } then show "(\x. indicator (\i. A i) x *\<^sub>R f x) \ borel_measurable M" apply (rule borel_measurable_LIMSEQ_real) apply assumption using intgbl set_integrable_def by blast qed (* Proof from Royden Real Analysis, p. 91. *) lemma lebesgue_integral_countable_add: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" assumes meas[intro]: "\i::nat. A i \ sets M" and disj: "\i j. i \ j \ A i \ A j = {}" and intgbl: "set_integrable M (\i. A i) f" shows "LINT x:(\i. A i)|M. f x = (\i. (LINT x:(A i)|M. f x))" unfolding set_lebesgue_integral_def proof (subst integral_suminf[symmetric]) show int_A: "integrable M (\x. indicat_real (A i) x *\<^sub>R f x)" for i using intgbl unfolding set_integrable_def [symmetric] by (rule set_integrable_subset) auto { fix x assume "x \ space M" have "(\i. indicator (A i) x *\<^sub>R f x) sums (indicator (\i. A i) x *\<^sub>R f x)" by (intro sums_scaleR_left indicator_sums) fact } note sums = this have norm_f: "\i. set_integrable M (A i) (\x. norm (f x))" using int_A[THEN integrable_norm] unfolding set_integrable_def by auto show "AE x in M. summable (\i. norm (indicator (A i) x *\<^sub>R f x))" using disj by (intro AE_I2) (auto intro!: summable_mult2 sums_summable[OF indicator_sums]) show "summable (\i. LINT x|M. norm (indicator (A i) x *\<^sub>R f x))" proof (rule summableI_nonneg_bounded) fix n show "0 \ LINT x|M. norm (indicator (A n) x *\<^sub>R f x)" using norm_f by (auto intro!: integral_nonneg_AE) have "(\iR f x)) = (\i = set_lebesgue_integral M (\ix. norm (f x))" using norm_f by (subst set_integral_finite_Union) (auto simp: disjoint_family_on_def disj) also have "\ \ set_lebesgue_integral M (\i. A i) (\x. norm (f x))" using intgbl[unfolded set_integrable_def, THEN integrable_norm] norm_f unfolding set_lebesgue_integral_def set_integrable_def apply (intro integral_mono set_integrable_UN[of "{..iR f x)) \ set_lebesgue_integral M (\i. A i) (\x. norm (f x))" by simp qed show "LINT x|M. indicator (\(A ` UNIV)) x *\<^sub>R f x = LINT x|M. (\i. indicator (A i) x *\<^sub>R f x)" apply (rule Bochner_Integration.integral_cong[OF refl]) apply (subst suminf_scaleR_left[OF sums_summable[OF indicator_sums, OF disj], symmetric]) using sums_unique[OF indicator_sums[OF disj]] apply auto done qed lemma set_integral_cont_up: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" assumes [measurable]: "\i. A i \ sets M" and A: "incseq A" and intgbl: "set_integrable M (\i. A i) f" shows "(\i. LINT x:(A i)|M. f x) \ LINT x:(\i. A i)|M. f x" unfolding set_lebesgue_integral_def proof (intro integral_dominated_convergence[where w="\x. indicator (\i. A i) x *\<^sub>R norm (f x)"]) have int_A: "\i. set_integrable M (A i) f" using intgbl by (rule set_integrable_subset) auto show "\i. (\x. indicator (A i) x *\<^sub>R f x) \ borel_measurable M" using int_A integrable_iff_bounded set_integrable_def by blast show "(\x. indicator (\(A ` UNIV)) x *\<^sub>R f x) \ borel_measurable M" using integrable_iff_bounded intgbl set_integrable_def by blast show "integrable M (\x. indicator (\i. A i) x *\<^sub>R norm (f x))" using int_A intgbl integrable_norm unfolding set_integrable_def by fastforce { fix x i assume "x \ A i" with A have "(\xa. indicator (A xa) x::real) \ 1 \ (\xa. 1::real) \ 1" by (intro filterlim_cong refl) (fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) } then show "AE x in M. (\i. indicator (A i) x *\<^sub>R f x) \ indicator (\i. A i) x *\<^sub>R f x" by (intro AE_I2 tendsto_intros) (auto split: split_indicator) qed (auto split: split_indicator) (* Can the int0 hypothesis be dropped? *) lemma set_integral_cont_down: fixes f :: "_ \ 'a :: {banach, second_countable_topology}" assumes [measurable]: "\i. A i \ sets M" and A: "decseq A" and int0: "set_integrable M (A 0) f" shows "(\i::nat. LINT x:(A i)|M. f x) \ LINT x:(\i. A i)|M. f x" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) have int_A: "\i. set_integrable M (A i) f" using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def) have "integrable M (\c. norm (indicat_real (A 0) c *\<^sub>R f c))" by (metis (no_types) int0 integrable_norm set_integrable_def) then show "integrable M (\x. indicator (A 0) x *\<^sub>R norm (f x))" by force have "set_integrable M (\i. A i) f" using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def) with int_A show "(\x. indicat_real (\(A ` UNIV)) x *\<^sub>R f x) \ borel_measurable M" "\i. (\x. indicat_real (A i) x *\<^sub>R f x) \ borel_measurable M" by (auto simp: set_integrable_def) show "\i. AE x in M. norm (indicator (A i) x *\<^sub>R f x) \ indicator (A 0) x *\<^sub>R norm (f x)" using A by (auto split: split_indicator simp: decseq_def) { fix x i assume "x \ space M" "x \ A i" with A have "(\i. indicator (A i) x::real) \ 0 \ (\i. 0::real) \ 0" by (intro filterlim_cong refl) (auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) } then show "AE x in M. (\i. indicator (A i) x *\<^sub>R f x) \ indicator (\i. A i) x *\<^sub>R f x" by (intro AE_I2 tendsto_intros) (auto split: split_indicator) qed lemma set_integral_at_point: fixes a :: real assumes "set_integrable M {a} f" and [simp]: "{a} \ sets M" and "(emeasure M) {a} \ \" shows "(LINT x:{a} | M. f x) = f a * measure M {a}" proof- have "set_lebesgue_integral M {a} f = set_lebesgue_integral M {a} (%x. f a)" by (intro set_lebesgue_integral_cong) simp_all then show ?thesis using assms unfolding set_lebesgue_integral_def by simp qed abbreviation complex_integrable :: "'a measure \ ('a \ complex) \ bool" where "complex_integrable M f \ integrable M f" abbreviation complex_lebesgue_integral :: "'a measure \ ('a \ complex) \ complex" ("integral\<^sup>C") where "integral\<^sup>C M f == integral\<^sup>L M f" syntax "_complex_lebesgue_integral" :: "pttrn \ complex \ 'a measure \ complex" ("\\<^sup>C _. _ \_" [60,61] 110) translations "\\<^sup>Cx. f \M" == "CONST complex_lebesgue_integral M (\x. f)" syntax "_ascii_complex_lebesgue_integral" :: "pttrn \ 'a measure \ real \ real" ("(3CLINT _|_. _)" [0,110,60] 60) translations "CLINT x|M. f" == "CONST complex_lebesgue_integral M (\x. f)" lemma complex_integrable_cnj [simp]: "complex_integrable M (\x. cnj (f x)) \ complex_integrable M f" proof assume "complex_integrable M (\x. cnj (f x))" then have "complex_integrable M (\x. cnj (cnj (f x)))" by (rule integrable_cnj) then show "complex_integrable M f" by simp qed simp lemma complex_of_real_integrable_eq: "complex_integrable M (\x. complex_of_real (f x)) \ integrable M f" proof assume "complex_integrable M (\x. complex_of_real (f x))" then have "integrable M (\x. Re (complex_of_real (f x)))" by (rule integrable_Re) then show "integrable M f" by simp qed simp abbreviation complex_set_integrable :: "'a measure \ 'a set \ ('a \ complex) \ bool" where "complex_set_integrable M A f \ set_integrable M A f" abbreviation complex_set_lebesgue_integral :: "'a measure \ 'a set \ ('a \ complex) \ complex" where "complex_set_lebesgue_integral M A f \ set_lebesgue_integral M A f" syntax "_ascii_complex_set_lebesgue_integral" :: "pttrn \ 'a set \ 'a measure \ real \ real" ("(4CLINT _:_|_. _)" [0,60,110,61] 60) translations "CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (\x. f)" lemma set_measurable_continuous_on_ivl: assumes "continuous_on {a..b} (f :: real \ real)" shows "set_borel_measurable borel {a..b} f" unfolding set_borel_measurable_def by (rule borel_measurable_continuous_on_indicator[OF _ assms]) simp text\This notation is from Sébastien Gouëzel: His use is not directly in line with the notations in this file, they are more in line with sum, and more readable he thinks.\ abbreviation "set_nn_integral M A f \ nn_integral M (\x. f x * indicator A x)" syntax "_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" ("(\\<^sup>+((_)\(_)./ _)/\_)" [0,60,110,61] 60) "_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" ("(\((_)\(_)./ _)/\_)" [0,60,110,61] 60) translations "\\<^sup>+x \ A. f \M" == "CONST set_nn_integral M A (\x. f)" "\x \ A. f \M" == "CONST set_lebesgue_integral M A (\x. f)" lemma nn_integral_disjoint_pair: assumes [measurable]: "f \ borel_measurable M" "B \ sets M" "C \ sets M" "B \ C = {}" shows "(\\<^sup>+x \ B \ C. f x \M) = (\\<^sup>+x \ B. f x \M) + (\\<^sup>+x \ C. f x \M)" proof - have mes: "\D. D \ sets M \ (\x. f x * indicator D x) \ borel_measurable M" by simp have pos: "\D. AE x in M. f x * indicator D x \ 0" using assms(2) by auto have "\x. f x * indicator (B \ C) x = f x * indicator B x + f x * indicator C x" using assms(4) by (auto split: split_indicator) then have "(\\<^sup>+x. f x * indicator (B \ C) x \M) = (\\<^sup>+x. f x * indicator B x + f x * indicator C x \M)" by simp also have "... = (\\<^sup>+x. f x * indicator B x \M) + (\\<^sup>+x. f x * indicator C x \M)" by (rule nn_integral_add) (auto simp add: assms mes pos) finally show ?thesis by simp qed lemma nn_integral_disjoint_pair_countspace: assumes "B \ C = {}" shows "(\\<^sup>+x \ B \ C. f x \count_space UNIV) = (\\<^sup>+x \ B. f x \count_space UNIV) + (\\<^sup>+x \ C. f x \count_space UNIV)" by (rule nn_integral_disjoint_pair) (simp_all add: assms) lemma nn_integral_null_delta: assumes "A \ sets M" "B \ sets M" "(A - B) \ (B - A) \ null_sets M" shows "(\\<^sup>+x \ A. f x \M) = (\\<^sup>+x \ B. f x \M)" proof (rule nn_integral_cong_AE) have *: "AE a in M. a \ (A - B) \ (B - A)" using assms(3) AE_not_in by blast then show \AE x in M. f x * indicator A x = f x * indicator B x\ by auto qed proposition nn_integral_disjoint_family: assumes [measurable]: "f \ borel_measurable M" "\(n::nat). B n \ sets M" and "disjoint_family B" shows "(\\<^sup>+x \ (\n. B n). f x \M) = (\n. (\\<^sup>+x \ B n. f x \M))" proof - have "(\\<^sup>+ x. (\n. f x * indicator (B n) x) \M) = (\n. (\\<^sup>+ x. f x * indicator (B n) x \M))" by (rule nn_integral_suminf) simp moreover have "(\n. f x * indicator (B n) x) = f x * indicator (\n. B n) x" for x proof (cases) assume "x \ (\n. B n)" then obtain n where "x \ B n" by blast have a: "finite {n}" by simp have "\i. i \ n \ x \ B i" using \x \ B n\ assms(3) disjoint_family_on_def by (metis IntI UNIV_I empty_iff) then have "\i. i \ {n} \ indicator (B i) x = (0::ennreal)" using indicator_def by simp then have b: "\i. i \ {n} \ f x * indicator (B i) x = (0::ennreal)" by simp define h where "h = (\i. f x * indicator (B i) x)" then have "\i. i \ {n} \ h i = 0" using b by simp then have "(\i. h i) = (\i\{n}. h i)" by (metis sums_unique[OF sums_finite[OF a]]) then have "(\i. h i) = h n" by simp then have "(\n. f x * indicator (B n) x) = f x * indicator (B n) x" using h_def by simp then have "(\n. f x * indicator (B n) x) = f x" using \x \ B n\ indicator_def by simp then show ?thesis using \x \ (\n. B n)\ by auto next assume "x \ (\n. B n)" then have "\n. f x * indicator (B n) x = 0" by simp have "(\n. f x * indicator (B n) x) = 0" by (simp add: \\n. f x * indicator (B n) x = 0\) then show ?thesis using \x \ (\n. B n)\ by auto qed ultimately show ?thesis by simp qed lemma nn_set_integral_add: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" "A \ sets M" shows "(\\<^sup>+x \ A. (f x + g x) \M) = (\\<^sup>+x \ A. f x \M) + (\\<^sup>+x \ A. g x \M)" proof - have "(\\<^sup>+x \ A. (f x + g x) \M) = (\\<^sup>+x. (f x * indicator A x + g x * indicator A x) \M)" by (auto simp add: indicator_def intro!: nn_integral_cong) also have "... = (\\<^sup>+x. f x * indicator A x \M) + (\\<^sup>+x. g x * indicator A x \M)" apply (rule nn_integral_add) using assms(1) assms(2) by auto finally show ?thesis by simp qed lemma nn_set_integral_cong: assumes "AE x in M. f x = g x" shows "(\\<^sup>+x \ A. f x \M) = (\\<^sup>+x \ A. g x \M)" apply (rule nn_integral_cong_AE) using assms(1) by auto lemma nn_set_integral_set_mono: "A \ B \ (\\<^sup>+ x \ A. f x \M) \ (\\<^sup>+ x \ B. f x \M)" by (auto intro!: nn_integral_mono split: split_indicator) lemma nn_set_integral_mono: assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" "A \ sets M" and "AE x\A in M. f x \ g x" shows "(\\<^sup>+x \ A. f x \M) \ (\\<^sup>+x \ A. g x \M)" by (auto intro!: nn_integral_mono_AE split: split_indicator simp: assms) lemma nn_set_integral_space [simp]: shows "(\\<^sup>+ x \ space M. f x \M) = (\\<^sup>+x. f x \M)" by (metis (mono_tags, lifting) indicator_simps(1) mult.right_neutral nn_integral_cong) lemma nn_integral_count_compose_inj: assumes "inj_on g A" shows "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ g`A. f y \count_space UNIV)" proof - have "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+x. f (g x) \count_space A)" by (auto simp add: nn_integral_count_space_indicator[symmetric]) also have "... = (\\<^sup>+y. f y \count_space (g`A))" by (simp add: assms nn_integral_bij_count_space inj_on_imp_bij_betw) also have "... = (\\<^sup>+y \ g`A. f y \count_space UNIV)" by (auto simp add: nn_integral_count_space_indicator[symmetric]) finally show ?thesis by simp qed lemma nn_integral_count_compose_bij: assumes "bij_betw g A B" shows "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ B. f y \count_space UNIV)" proof - have "inj_on g A" using assms bij_betw_def by auto then have "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ g`A. f y \count_space UNIV)" by (rule nn_integral_count_compose_inj) then show ?thesis using assms by (simp add: bij_betw_def) qed lemma set_integral_null_delta: fixes f::"_ \ _ :: {banach, second_countable_topology}" assumes [measurable]: "integrable M f" "A \ sets M" "B \ sets M" and null: "(A - B) \ (B - A) \ null_sets M" shows "(\x \ A. f x \M) = (\x \ B. f x \M)" proof (rule set_integral_cong_set) have *: "AE a in M. a \ (A - B) \ (B - A)" using null AE_not_in by blast then show "AE x in M. (x \ B) = (x \ A)" by auto qed (simp_all add: set_borel_measurable_def) lemma set_integral_space: assumes "integrable M f" shows "(\x \ space M. f x \M) = (\x. f x \M)" by (metis (no_types, lifting) indicator_simps(1) integral_cong scaleR_one set_lebesgue_integral_def) lemma null_if_pos_func_has_zero_nn_int: fixes f::"'a \ ennreal" assumes [measurable]: "f \ borel_measurable M" "A \ sets M" and "AE x\A in M. f x > 0" "(\\<^sup>+x\A. f x \M) = 0" shows "A \ null_sets M" proof - have "AE x in M. f x * indicator A x = 0" by (subst nn_integral_0_iff_AE[symmetric], auto simp add: assms(4)) then have "AE x\A in M. False" using assms(3) by auto then show "A \ null_sets M" using assms(2) by (simp add: AE_iff_null_sets) qed lemma null_if_pos_func_has_zero_int: assumes [measurable]: "integrable M f" "A \ sets M" and "AE x\A in M. f x > 0" "(\x\A. f x \M) = (0::real)" shows "A \ null_sets M" proof - have "AE x in M. indicator A x * f x = 0" apply (subst integral_nonneg_eq_0_iff_AE[symmetric]) using assms integrable_mult_indicator[OF \A \ sets M\ assms(1)] by (auto simp: set_lebesgue_integral_def) then have "AE x\A in M. f x = 0" by auto then have "AE x\A in M. False" using assms(3) by auto then show "A \ null_sets M" using assms(2) by (simp add: AE_iff_null_sets) qed text\The next lemma is a variant of \density_unique\. Note that it uses the notation for nonnegative set integrals introduced earlier.\ lemma (in sigma_finite_measure) density_unique2: assumes [measurable]: "f \ borel_measurable M" "f' \ borel_measurable M" assumes density_eq: "\A. A \ sets M \ (\\<^sup>+ x \ A. f x \M) = (\\<^sup>+ x \ A. f' x \M)" shows "AE x in M. f x = f' x" proof (rule density_unique) show "density M f = density M f'" by (intro measure_eqI) (auto simp: emeasure_density intro!: density_eq) qed (auto simp add: assms) text \The next lemma implies the same statement for Banach-space valued functions using Hahn-Banach theorem and linear forms. Since they are not yet easily available, I only formulate it for real-valued functions.\ lemma density_unique_real: fixes f f'::"_ \ real" assumes M[measurable]: "integrable M f" "integrable M f'" assumes density_eq: "\A. A \ sets M \ (\x \ A. f x \M) = (\x \ A. f' x \M)" shows "AE x in M. f x = f' x" proof - define A where "A = {x \ space M. f x < f' x}" then have [measurable]: "A \ sets M" by simp have "(\x \ A. (f' x - f x) \M) = (\x \ A. f' x \M) - (\x \ A. f x \M)" using \A \ sets M\ M integrable_mult_indicator set_integrable_def by blast then have "(\x \ A. (f' x - f x) \M) = 0" using assms(3) by simp then have "A \ null_sets M" using A_def null_if_pos_func_has_zero_int[where ?f = "\x. f' x - f x" and ?A = A] assms by auto then have "AE x in M. x \ A" by (simp add: AE_not_in) then have *: "AE x in M. f' x \ f x" unfolding A_def by auto define B where "B = {x \ space M. f' x < f x}" then have [measurable]: "B \ sets M" by simp have "(\x \ B. (f x - f' x) \M) = (\x \ B. f x \M) - (\x \ B. f' x \M)" using \B \ sets M\ M integrable_mult_indicator set_integrable_def by blast then have "(\x \ B. (f x - f' x) \M) = 0" using assms(3) by simp then have "B \ null_sets M" using B_def null_if_pos_func_has_zero_int[where ?f = "\x. f x - f' x" and ?A = B] assms by auto then have "AE x in M. x \ B" by (simp add: AE_not_in) then have "AE x in M. f' x \ f x" unfolding B_def by auto then show ?thesis using * by auto qed text \The next lemma shows that \L\<^sup>1\ convergence of a sequence of functions follows from almost everywhere convergence and the weaker condition of the convergence of the integrated norms (or even just the nontrivial inequality about them). Useful in a lot of contexts! This statement (or its variations) are known as Scheffe lemma. The formalization is more painful as one should jump back and forth between reals and ereals and justify all the time positivity or integrability (thankfully, measurability is handled more or less automatically).\ proposition Scheffe_lemma1: assumes "\n. integrable M (F n)" "integrable M f" "AE x in M. (\n. F n x) \ f x" "limsup (\n. \\<^sup>+ x. norm(F n x) \M) \ (\\<^sup>+ x. norm(f x) \M)" shows "(\n. \\<^sup>+ x. norm(F n x - f x) \M) \ 0" proof - have [measurable]: "\n. F n \ borel_measurable M" "f \ borel_measurable M" using assms(1) assms(2) by simp_all define G where "G = (\n x. norm(f x) + norm(F n x) - norm(F n x - f x))" have [measurable]: "\n. G n \ borel_measurable M" unfolding G_def by simp have G_pos[simp]: "\n x. G n x \ 0" unfolding G_def by (metis ge_iff_diff_ge_0 norm_minus_commute norm_triangle_ineq4) have finint: "(\\<^sup>+ x. norm(f x) \M) \ \" using has_bochner_integral_implies_finite_norm[OF has_bochner_integral_integrable[OF \integrable M f\]] by simp then have fin2: "2 * (\\<^sup>+ x. norm(f x) \M) \ \" by (auto simp: ennreal_mult_eq_top_iff) { fix x assume *: "(\n. F n x) \ f x" then have "(\n. norm(F n x)) \ norm(f x)" using tendsto_norm by blast moreover have "(\n. norm(F n x - f x)) \ 0" using * Lim_null tendsto_norm_zero_iff by fastforce ultimately have a: "(\n. norm(F n x) - norm(F n x - f x)) \ norm(f x)" using tendsto_diff by fastforce have "(\n. norm(f x) + (norm(F n x) - norm(F n x - f x))) \ norm(f x) + norm(f x)" by (rule tendsto_add) (auto simp add: a) moreover have "\n. G n x = norm(f x) + (norm(F n x) - norm(F n x - f x))" unfolding G_def by simp ultimately have "(\n. G n x) \ 2 * norm(f x)" by simp then have "(\n. ennreal(G n x)) \ ennreal(2 * norm(f x))" by simp then have "liminf (\n. ennreal(G n x)) = ennreal(2 * norm(f x))" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast } then have "AE x in M. liminf (\n. ennreal(G n x)) = ennreal(2 * norm(f x))" using assms(3) by auto then have "(\\<^sup>+ x. liminf (\n. ennreal (G n x)) \M) = (\\<^sup>+ x. 2 * ennreal(norm(f x)) \M)" by (simp add: nn_integral_cong_AE ennreal_mult) also have "... = 2 * (\\<^sup>+ x. norm(f x) \M)" by (rule nn_integral_cmult) auto finally have int_liminf: "(\\<^sup>+ x. liminf (\n. ennreal (G n x)) \M) = 2 * (\\<^sup>+ x. norm(f x) \M)" by simp have "(\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) = (\\<^sup>+x. norm(f x) \M) + (\\<^sup>+x. norm(F n x) \M)" for n by (rule nn_integral_add) (auto simp add: assms) then have "limsup (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M)) = limsup (\n. (\\<^sup>+x. norm(f x) \M) + (\\<^sup>+x. norm(F n x) \M))" by simp also have "... = (\\<^sup>+x. norm(f x) \M) + limsup (\n. (\\<^sup>+x. norm(F n x) \M))" by (rule Limsup_const_add, auto simp add: finint) also have "... \ (\\<^sup>+x. norm(f x) \M) + (\\<^sup>+x. norm(f x) \M)" using assms(4) by (simp add: add_left_mono) also have "... = 2 * (\\<^sup>+x. norm(f x) \M)" unfolding one_add_one[symmetric] distrib_right by simp ultimately have a: "limsup (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M)) \ 2 * (\\<^sup>+x. norm(f x) \M)" by simp have le: "ennreal (norm (F n x - f x)) \ ennreal (norm (f x)) + ennreal (norm (F n x))" for n x by (simp add: norm_minus_commute norm_triangle_ineq4 ennreal_minus flip: ennreal_plus) then have le2: "(\\<^sup>+ x. ennreal (norm (F n x - f x)) \M) \ (\\<^sup>+ x. ennreal (norm (f x)) + ennreal (norm (F n x)) \M)" for n by (rule nn_integral_mono) have "2 * (\\<^sup>+ x. norm(f x) \M) = (\\<^sup>+ x. liminf (\n. ennreal (G n x)) \M)" by (simp add: int_liminf) also have "\ \ liminf (\n. (\\<^sup>+x. G n x \M))" by (rule nn_integral_liminf) auto also have "liminf (\n. (\\<^sup>+x. G n x \M)) = liminf (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M))" proof (intro arg_cong[where f=liminf] ext) fix n have "\x. ennreal(G n x) = ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x))" unfolding G_def by (simp add: ennreal_minus flip: ennreal_plus) moreover have "(\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x)) \M) = (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M)" proof (rule nn_integral_diff) from le show "AE x in M. ennreal (norm (F n x - f x)) \ ennreal (norm (f x)) + ennreal (norm (F n x))" by simp from le2 have "(\\<^sup>+x. ennreal (norm (F n x - f x)) \M) < \" using assms(1) assms(2) by (metis has_bochner_integral_implies_finite_norm integrable.simps Bochner_Integration.integrable_diff) then show "(\\<^sup>+x. ennreal (norm (F n x - f x)) \M) \ \" by simp qed (auto simp add: assms) ultimately show "(\\<^sup>+x. G n x \M) = (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M)" by simp qed finally have "2 * (\\<^sup>+ x. norm(f x) \M) + limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M)) \ liminf (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - (\\<^sup>+x. norm(F n x - f x) \M)) + limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M))" by (intro add_mono) auto also have "\ \ (limsup (\n. \\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M) - limsup (\n. \\<^sup>+x. norm (F n x - f x) \M)) + limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M))" by (intro add_mono liminf_minus_ennreal le2) auto also have "\ = limsup (\n. (\\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \M))" by (intro diff_add_cancel_ennreal Limsup_mono always_eventually allI le2) also have "\ \ 2 * (\\<^sup>+x. norm(f x) \M)" by fact finally have "limsup (\n. (\\<^sup>+x. norm(F n x - f x) \M)) = 0" using fin2 by simp then show ?thesis by (rule tendsto_0_if_Limsup_eq_0_ennreal) qed proposition Scheffe_lemma2: fixes F::"nat \ 'a \ 'b::{banach, second_countable_topology}" assumes "\ n::nat. F n \ borel_measurable M" "integrable M f" "AE x in M. (\n. F n x) \ f x" "\n. (\\<^sup>+ x. norm(F n x) \M) \ (\\<^sup>+ x. norm(f x) \M)" shows "(\n. \\<^sup>+ x. norm(F n x - f x) \M) \ 0" proof (rule Scheffe_lemma1) fix n::nat have "(\\<^sup>+ x. norm(f x) \M) < \" using assms(2) by (metis has_bochner_integral_implies_finite_norm integrable.cases) then have "(\\<^sup>+ x. norm(F n x) \M) < \" using assms(4)[of n] by auto then show "integrable M (F n)" by (subst integrable_iff_bounded, simp add: assms(1)[of n]) qed (auto simp add: assms Limsup_bounded) lemma tendsto_set_lebesgue_integral_at_right: fixes a b :: real and f :: "real \ 'a :: {banach,second_countable_topology}" assumes "a < b" and sets: "\a'. a' \ {a<..b} \ {a'..b} \ sets M" and "set_integrable M {a<..b} f" shows "((\a'. set_lebesgue_integral M {a'..b} f) \ set_lebesgue_integral M {a<..b} f) (at_right a)" proof (rule tendsto_at_right_sequentially[OF assms(1)], goal_cases) case (1 S) have eq: "(\n. {S n..b}) = {a<..b}" proof safe fix x n assume "x \ {S n..b}" with 1(1,2)[of n] show "x \ {a<..b}" by auto next fix x assume "x \ {a<..b}" with order_tendstoD[OF \S \ a\, of x] show "x \ (\n. {S n..b})" by (force simp: eventually_at_top_linorder dest: less_imp_le) qed have "(\n. set_lebesgue_integral M {S n..b} f) \ set_lebesgue_integral M (\n. {S n..b}) f" by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) with eq show ?case by simp qed text \ The next lemmas relate convergence of integrals over an interval to improper integrals. \ lemma tendsto_set_lebesgue_integral_at_left: fixes a b :: real and f :: "real \ 'a :: {banach,second_countable_topology}" assumes "a < b" and sets: "\b'. b' \ {a.. {a..b'} \ sets M" and "set_integrable M {a..b'. set_lebesgue_integral M {a..b'} f) \ set_lebesgue_integral M {a..n. {a..S n}) = {a.. {a..S n}" with 1(1,2)[of n] show "x \ {a.. {a..S \ b\, of x] show "x \ (\n. {a..S n})" by (force simp: eventually_at_top_linorder dest: less_imp_le) qed have "(\n. set_lebesgue_integral M {a..S n} f) \ set_lebesgue_integral M (\n. {a..S n}) f" by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) with eq show ?case by simp qed proposition tendsto_set_lebesgue_integral_at_top: fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes sets: "\b. b \ a \ {a..b} \ sets M" and int: "set_integrable M {a..} f" shows "((\b. set_lebesgue_integral M {a..b} f) \ set_lebesgue_integral M {a..} f) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" show "(\n. set_lebesgue_integral M {a..X n} f) \ set_lebesgue_integral M {a..} f" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) show "integrable M (\x. indicat_real {a..} x *\<^sub>R norm (f x))" using integrable_norm[OF int[unfolded set_integrable_def]] by simp show "AE x in M. (\n. indicator {a..X n} x *\<^sub>R f x) \ indicat_real {a..} x *\<^sub>R f x" proof fix x from \filterlim X at_top sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_top_ge[where c=x] by auto then show "(\n. indicator {a..X n} x *\<^sub>R f x) \ indicat_real {a..} x *\<^sub>R f x" by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {a..X n} x *\<^sub>R f x) \ indicator {a..} x *\<^sub>R norm (f x)" by (auto split: split_indicator) next from int show "(\x. indicat_real {a..} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) next fix n :: nat from sets have "{a..X n} \ sets M" by (cases "X n \ a") auto with int have "set_integrable M {a..X n} f" by (rule set_integrable_subset) auto thus "(\x. indicat_real {a..X n} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) qed qed proposition tendsto_set_lebesgue_integral_at_bot: fixes f :: "real \ 'a::{banach, second_countable_topology}" assumes sets: "\a. a \ b \ {a..b} \ sets M" and int: "set_integrable M {..b} f" shows "((\a. set_lebesgue_integral M {a..b} f) \ set_lebesgue_integral M {..b} f) at_bot" proof (rule tendsto_at_botI_sequentially) fix X :: "nat \ real" assume "filterlim X at_bot sequentially" show "(\n. set_lebesgue_integral M {X n..b} f) \ set_lebesgue_integral M {..b} f" unfolding set_lebesgue_integral_def proof (rule integral_dominated_convergence) show "integrable M (\x. indicat_real {..b} x *\<^sub>R norm (f x))" using integrable_norm[OF int[unfolded set_integrable_def]] by simp show "AE x in M. (\n. indicator {X n..b} x *\<^sub>R f x) \ indicat_real {..b} x *\<^sub>R f x" proof fix x from \filterlim X at_bot sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_bot_le[where c=x] by auto then show "(\n. indicator {X n..b} x *\<^sub>R f x) \ indicat_real {..b} x *\<^sub>R f x" by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {X n..b} x *\<^sub>R f x) \ indicator {..b} x *\<^sub>R norm (f x)" by (auto split: split_indicator) next from int show "(\x. indicat_real {..b} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) next fix n :: nat from sets have "{X n..b} \ sets M" by (cases "X n \ b") auto with int have "set_integrable M {X n..b} f" by (rule set_integrable_subset) auto thus "(\x. indicat_real {X n..b} x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_integrable_def) qed qed theorem integral_Markov_inequality': fixes u :: "'a \ real" assumes [measurable]: "set_integrable M A u" and "A \ sets M" assumes "AE x in M. x \ A \ u x \ 0" and "0 < (c::real)" shows "emeasure M {x\A. u x \ c} \ (1/c::real) * (\x\A. u x \M)" proof - have "(\x. u x * indicator A x) \ borel_measurable M" using assms by (auto simp: set_integrable_def mult_ac) hence "(\x. ennreal (u x * indicator A x)) \ borel_measurable M" by measurable also have "(\x. ennreal (u x * indicator A x)) = (\x. ennreal (u x) * indicator A x)" by (intro ext) (auto simp: indicator_def) finally have meas: "\ \ borel_measurable M" . from assms(3) have AE: "AE x in M. 0 \ u x * indicator A x" by eventually_elim (auto simp: indicator_def) have nonneg: "set_lebesgue_integral M A u \ 0" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF AE]) (auto simp: mult_ac) have A: "A \ space M" using \A \ sets M\ by (simp add: sets.sets_into_space) have "{x \ A. u x \ c} = {x \ A. ennreal(1/c) * u x \ 1}" using \c>0\ A by (auto simp: ennreal_mult'[symmetric]) then have "emeasure M {x \ A. u x \ c} = emeasure M ({x \ A. ennreal(1/c) * u x \ 1})" by simp also have "... \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator A x \M)" by (intro nn_integral_Markov_inequality meas assms) also have "(\\<^sup>+ x. ennreal(u x) * indicator A x \M) = ennreal (set_lebesgue_integral M A u)" unfolding set_lebesgue_integral_def nn_integral_set_ennreal using assms AE by (subst nn_integral_eq_integral) (simp_all add: mult_ac set_integrable_def) finally show ?thesis using \c > 0\ nonneg by (subst ennreal_mult) auto qed theorem integral_Markov_inequality'_measure: assumes [measurable]: "set_integrable M A u" and "A \ sets M" and "AE x in M. x \ A \ 0 \ u x" "0 < (c::real)" shows "measure M {x\A. u x \ c} \ (\x\A. u x \M) / c" proof - have nonneg: "set_lebesgue_integral M A u \ 0" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF assms(3)]) (auto simp: mult_ac) have le: "emeasure M {x\A. u x \ c} \ ennreal ((1/c) * (\x\A. u x \M))" by (rule integral_Markov_inequality') (use assms in auto) also have "\ < top" by auto finally have "ennreal (measure M {x\A. u x \ c}) = emeasure M {x\A. u x \ c}" by (intro emeasure_eq_ennreal_measure [symmetric]) auto also note le finally show ?thesis using nonneg by (subst (asm) ennreal_le_iff) (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms) qed theorem%important (in finite_measure) Chernoff_ineq_ge: assumes s: "s > 0" assumes integrable: "set_integrable M A (\x. exp (s * f x))" and "A \ sets M" shows "measure M {x\A. f x \ a} \ exp (-s * a) * (\x\A. exp (s * f x) \M)" proof - have "{x\A. f x \ a} = {x\A. exp (s * f x) \ exp (s * a)}" using s by auto also have "measure M \ \ set_lebesgue_integral M A (\x. exp (s * f x)) / exp (s * a)" by (intro integral_Markov_inequality'_measure assms) auto finally show ?thesis by (simp add: exp_minus field_simps) qed theorem%important (in finite_measure) Chernoff_ineq_le: assumes s: "s > 0" assumes integrable: "set_integrable M A (\x. exp (-s * f x))" and "A \ sets M" shows "measure M {x\A. f x \ a} \ exp (s * a) * (\x\A. exp (-s * f x) \M)" proof - have "{x\A. f x \ a} = {x\A. exp (-s * f x) \ exp (-s * a)}" using s by auto also have "measure M \ \ set_lebesgue_integral M A (\x. exp (-s * f x)) / exp (-s * a)" by (intro integral_Markov_inequality'_measure assms) auto finally show ?thesis by (simp add: exp_minus field_simps) qed end diff --git a/src/HOL/Analysis/Sigma_Algebra.thy b/src/HOL/Analysis/Sigma_Algebra.thy --- a/src/HOL/Analysis/Sigma_Algebra.thy +++ b/src/HOL/Analysis/Sigma_Algebra.thy @@ -1,2304 +1,2314 @@ (* 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\ locale\<^marker>\tag 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\ locale\<^marker>\tag 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 \ \C\M. 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 "(\i\I. 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\\. \i\S. P i x} \ M" proof - have "{x\\. \i\S. P i x} = (\i\S. {x\\. P i x})" using \S \ {}\ by auto with assms show ?thesis by auto qed subsubsection \Ring of sets\ locale\<^marker>\tag 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 "(\i\I. 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 "\C\M. 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 \ (\a\M. \b\M. a \ b \ M) \ (\a\M. \b\M. a - b \ M)" proof assume "ring_of_sets \ M" then interpret ring_of_sets \ M . show "M \ Pow \ \ {} \ M \ (\a\M. \b\M. a \ b \ M) \ (\a\M. \b\M. 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\\. \i\S. P i x} \ M" proof - have "{x\\. \i\S. P i x} = (\i\S. {x\\. P i x})" by auto with assms show ?thesis by auto qed subsubsection \Algebra of sets\ locale\<^marker>\tag 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) subsubsection\<^marker>\tag 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\ locale\<^marker>\tag 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) = (\s\M \ range A. s)" by auto also have "(\s\M \ 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 "(\x\X. 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) = (\x\X. 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 "(\x\X. A x) \ M" proof - have "(\x\X. A x) = (\i\to_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 \ \ (\x\X. 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 "(\i\X. A i) \ M" proof - from A have "\i\X. A i \ M" by fast hence "\ - (\i\X. \ - A i) \ M" by blast moreover have "(\i\X. A i) = \ - (\i\X. \ - 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 "(\x\X. A x) \ M" proof - have "(\x\X. A x) = (\i\to_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) \ (\i\I. 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 "(\a\A. {a}) \ M" using assms by (intro countable_UN') auto also have "(\a\A. {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\\. \i\I. P i x} \ M" proof - have "{x\\. \i\I. P i x} = (\i\I. {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\\. \i\I. P i x} \ M" proof - have "{x\\. \i\I. P i x} = (\i\I. {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\\. \!i\I. P i x} \ M" proof - have "{x\\. \!i\I. P i x} = {x\\. \i\I. P i x \ (\j\I. 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::countable\X. 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::countable\X. 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 subsubsection\<^marker>\tag 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\ text\<^marker>\tag important\ \Sigma algebras can naturally be created as the closure of any set of M with regard to the properties just postulated.\ inductive_set\<^marker>\tag 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 "(\i\S. a i) \ sigma_sets sp A" proof - from ai have "\i. (if i\S then a i else sp) \ sigma_sets sp A" by (simp add: sigma_sets.intros(2-) sigma_sets_top) hence "(\i. (if i\S then a i else sp)) \ sigma_sets sp A" by (rule sigma_sets_Inter [OF Asb]) also have "(\i. (if i\S then a i else sp)) = (\i\S. 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) - from choice[OF this] guess f .. + 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 - from choice[OF this] guess A .. note A = this - with A show ?case + 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.. 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.. 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 subsubsection\<^marker>\tag 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 guess Ca .. note Ca = this - from b guess Cb .. note Cb = this + 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 - show "disjoint (Ca \ Cb)" - using \a \ b = {}\ Ca Cb by (auto intro!: disjoint_union) - qed (insert Ca Cb, auto) + 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 guess Ca .. note Ca = this - from b guess Cb .. note Cb = this + 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 (insert Ca Cb, auto simp: C_def) + 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 assume a: "a \ ?R" then guess Ca .. note Ca = this - fix b assume b: "b \ ?R" then guess Cb .. note Cb = this - + { + 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 } + 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) subsubsection\<^marker>\tag 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\ definition\<^marker>\tag 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\ locale\<^marker>\tag 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" definition\<^marker>\tag 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 \ \k\I. 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" definition\<^marker>\tag 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\ text\<^marker>\tag 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\ definition\<^marker>\tag important\ positive :: "'a set set \ ('a set \ ennreal) \ bool" where "positive M \ \ \ {} = 0" definition\<^marker>\tag 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))" definition\<^marker>\tag important\ measure_space :: "'a set \ 'a set set \ ('a set \ ennreal) \ bool" where "measure_space \ A \ \ sigma_algebra \ A \ positive A \ \ countably_additive A \" typedef\<^marker>\tag 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 definition\<^marker>\tag important\ space :: "'a measure \ 'a set" where "space M = fst (Rep_measure M)" definition\<^marker>\tag important\ sets :: "'a measure \ 'a set set" where "sets M = fst (snd (Rep_measure M))" definition\<^marker>\tag important\ emeasure :: "'a measure \ 'a set \ ennreal" where "emeasure M = snd (snd (Rep_measure M))" definition\<^marker>\tag 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) definition\<^marker>\tag 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 \ (sigma_sets \ A) \'" using assms by (rule measure_space_eq) with eq show ?thesis by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure]) qed lemma shows space_measure_of_conv: "space (measure_of \ A \) = \" (is ?space) and sets_measure_of_conv: "sets (measure_of \ A \) = (if A \ Pow \ then sigma_sets \ A else {{}, \})" (is ?sets) and emeasure_measure_of_conv: "emeasure (measure_of \ A \) = (\B. if B \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ B else 0)" (is ?emeasure) proof - have "?space \ ?sets \ ?emeasure" proof(cases "measure_space \ (sigma_sets \ A) \") case True from measure_space_closed[OF this] sigma_sets_superset_generator[of A \] have "A \ Pow \" by simp hence "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) (\a. if a \ sigma_sets \ A then \ a else 0)" by(rule measure_space_eq) auto with True \A \ Pow \\ show ?thesis by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse) next case False thus ?thesis by(cases "A \ Pow \")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0') qed thus ?space ?sets ?emeasure by simp_all qed lemma [simp]: assumes A: "A \ Pow \" shows sets_measure_of: "sets (measure_of \ A \) = sigma_sets \ A" and space_measure_of: "space (measure_of \ A \) = \" using assms by(simp_all add: sets_measure_of_conv space_measure_of_conv) lemma space_in_measure_of[simp]: "\ \ sets (measure_of \ M \)" by (subst sets_measure_of_conv) (auto simp: sigma_sets_top) lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \ M \) = M" using space_closed by (auto intro!: sigma_sets_eq) lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \ M \) = \" by (rule space_measure_of_conv) lemma measure_of_subset: "M \ Pow \ \ M' \ M \ sets (measure_of \ M' \) \ sets (measure_of \ M \')" by (auto intro!: sigma_sets_subseteq) lemma emeasure_sigma: "emeasure (sigma \ A) = (\x. 0)" unfolding measure_of_def emeasure_def by (subst Abs_measure_inverse) (auto simp: measure_space_def positive_def countably_additive_def intro!: sigma_algebra_sigma_sets sigma_algebra_trivial) lemma sigma_sets_mono'': assumes "A \ sigma_sets C D" assumes "B \ D" assumes "D \ Pow C" shows "sigma_sets A B \ sigma_sets C D" proof fix x assume "x \ sigma_sets A B" thus "x \ sigma_sets C D" proof induct case (Basic a) with assms have "a \ D" by auto thus ?case .. next case Empty show ?case by (rule sigma_sets.Empty) next from assms have "A \ sets (sigma C D)" by (subst sets_measure_of[OF \D \ Pow C\]) moreover case (Compl a) hence "a \ sets (sigma C D)" by (subst sets_measure_of[OF \D \ Pow C\]) ultimately have "A - a \ sets (sigma C D)" .. thus ?case by (subst (asm) sets_measure_of[OF \D \ Pow C\]) next case (Union a) thus ?case by (intro sigma_sets.Union) qed qed lemma in_measure_of[intro, simp]: "M \ Pow \ \ A \ M \ A \ sets (measure_of \ M \)" by auto lemma space_empty_iff: "space N = {} \ sets N = {{}}" by (metis Pow_empty Sup_bot_conv(1) cSup_singleton empty_iff sets.sigma_sets_eq sets.space_closed sigma_sets_top subset_singletonD) subsubsection \Constructing simple \<^typ>\'a measure\\ proposition emeasure_measure_of: assumes M: "M = measure_of \ A \" assumes ms: "A \ Pow \" "positive (sets M) \" "countably_additive (sets M) \" assumes X: "X \ sets M" shows "emeasure M X = \ X" proof - interpret sigma_algebra \ "sigma_sets \ A" by (rule sigma_algebra_sigma_sets) fact have "measure_space \ (sigma_sets \ A) \" using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets) thus ?thesis using X ms by(simp add: M emeasure_measure_of_conv sets_measure_of_conv) qed lemma emeasure_measure_of_sigma: assumes ms: "sigma_algebra \ M" "positive M \" "countably_additive M \" assumes A: "A \ M" shows "emeasure (measure_of \ M \) A = \ A" proof - interpret sigma_algebra \ M by fact have "measure_space \ (sigma_sets \ M) \" using ms sigma_sets_eq by (simp add: measure_space_def) thus ?thesis by(simp add: emeasure_measure_of_conv A) qed lemma measure_cases[cases type: measure]: obtains (measure) \ A \ where "x = Abs_measure (\, A, \)" "\a\-A. \ a = 0" "measure_space \ A \" by atomize_elim (cases x, auto) lemma sets_le_imp_space_le: "sets A \ sets B \ space A \ space B" by (auto dest: sets.sets_into_space) lemma sets_eq_imp_space_eq: "sets M = sets M' \ space M = space M'" by (auto intro!: antisym sets_le_imp_space_le) lemma emeasure_notin_sets: "A \ sets M \ emeasure M A = 0" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) lemma emeasure_neq_0_sets: "emeasure M A \ 0 \ A \ sets M" using emeasure_notin_sets[of A M] by blast lemma measure_notin_sets: "A \ sets M \ measure M A = 0" by (simp add: measure_def emeasure_notin_sets zero_ennreal.rep_eq) lemma measure_eqI: fixes M N :: "'a measure" assumes "sets M = sets N" and eq: "\A. A \ sets M \ emeasure M A = emeasure N A" shows "M = N" proof (cases M N rule: measure_cases[case_product measure_cases]) case (measure_measure \ A \ \' A' \') interpret M: sigma_algebra \ A using measure_measure by (auto simp: measure_space_def) interpret N: sigma_algebra \' A' using measure_measure by (auto simp: measure_space_def) have "A = sets M" "A' = sets N" using measure_measure by (simp_all add: sets_def Abs_measure_inverse) with \sets M = sets N\ have AA': "A = A'" by simp moreover from M.top N.top M.space_closed N.space_closed AA' have "\ = \'" by auto moreover { fix B have "\ B = \' B" proof cases assume "B \ A" with eq \A = sets M\ have "emeasure M B = emeasure N B" by simp with measure_measure show "\ B = \' B" by (simp add: emeasure_def Abs_measure_inverse) next assume "B \ A" with \A = sets M\ \A' = sets N\ \A = A'\ have "B \ sets M" "B \ sets N" by auto then have "emeasure M B = 0" "emeasure N B = 0" by (simp_all add: emeasure_notin_sets) with measure_measure show "\ B = \' B" by (simp add: emeasure_def Abs_measure_inverse) qed } then have "\ = \'" by auto ultimately show "M = N" by (simp add: measure_measure) qed lemma sigma_eqI: assumes [simp]: "M \ Pow \" "N \ Pow \" "sigma_sets \ M = sigma_sets \ N" shows "sigma \ M = sigma \ N" by (rule measure_eqI) (simp_all add: emeasure_sigma) subsubsection \Measurable functions\ definition\<^marker>\tag important\ measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" (infixr "\\<^sub>M" 60) where "measurable A B = {f \ space A \ space B. \y \ sets B. f -` y \ space A \ sets A}" lemma measurableI: "(\x. x \ space M \ f x \ space N) \ (\A. A \ sets N \ f -` A \ space M \ sets M) \ f \ measurable M N" by (auto simp: measurable_def) lemma measurable_space: "f \ measurable M A \ x \ space M \ f x \ space A" unfolding measurable_def by auto lemma measurable_sets: "f \ measurable M A \ S \ sets A \ f -` S \ space M \ sets M" unfolding measurable_def by auto lemma measurable_sets_Collect: assumes f: "f \ measurable M N" and P: "{x\space N. P x} \ sets N" shows "{x\space M. P (f x)} \ sets M" proof - have "f -` {x \ space N. P x} \ space M = {x\space M. P (f x)}" using measurable_space[OF f] by auto with measurable_sets[OF f P] show ?thesis by simp qed lemma measurable_sigma_sets: assumes B: "sets N = sigma_sets \ A" "A \ Pow \" and f: "f \ space M \ \" and ba: "\y. y \ A \ (f -` y) \ space M \ sets M" shows "f \ measurable M N" proof - interpret A: sigma_algebra \ "sigma_sets \ A" using B(2) by (rule sigma_algebra_sigma_sets) from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \: "\ = space N" by force { fix X assume "X \ sigma_sets \ A" then have "f -` X \ space M \ sets M \ X \ \" proof induct case (Basic a) then show ?case by (auto simp add: ba) (metis B(2) subsetD PowD) next case (Compl a) have [simp]: "f -` \ \ space M = space M" by (auto simp add: funcset_mem [OF f]) then show ?case by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl) next case (Union a) then show ?case by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast qed auto } with f show ?thesis by (auto simp add: measurable_def B \) qed lemma measurable_measure_of: assumes B: "N \ Pow \" and f: "f \ space M \ \" and ba: "\y. y \ N \ (f -` y) \ space M \ sets M" shows "f \ measurable M (measure_of \ N \)" proof - have "sets (measure_of \ N \) = sigma_sets \ N" using B by (rule sets_measure_of) from this assms show ?thesis by (rule measurable_sigma_sets) qed lemma measurable_iff_measure_of: assumes "N \ Pow \" "f \ space M \ \" shows "f \ measurable M (measure_of \ N \) \ (\A\N. f -` A \ space M \ sets M)" by (metis assms in_measure_of measurable_measure_of assms measurable_sets) lemma measurable_cong_sets: assumes sets: "sets M = sets M'" "sets N = sets N'" shows "measurable M N = measurable M' N'" using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def) lemma measurable_cong: assumes "\w. w \ space M \ f w = g w" shows "f \ measurable M M' \ g \ measurable M M'" unfolding measurable_def using assms by (simp cong: vimage_inter_cong Pi_cong) lemma measurable_cong': assumes "\w. w \ space M =simp=> f w = g w" shows "f \ measurable M M' \ g \ measurable M M'" unfolding measurable_def using assms by (simp cong: vimage_inter_cong Pi_cong add: simp_implies_def) lemma measurable_cong_simp: "M = N \ M' = N' \ (\w. w \ space M \ f w = g w) \ f \ measurable M M' \ g \ measurable N N'" by (metis measurable_cong) lemma measurable_compose: assumes f: "f \ measurable M N" and g: "g \ measurable N L" shows "(\x. g (f x)) \ measurable M L" proof - have "\A. (\x. g (f x)) -` A \ space M = f -` (g -` A \ space N) \ space M" using measurable_space[OF f] by auto with measurable_space[OF f] measurable_space[OF g] show ?thesis by (auto intro: measurable_sets[OF f] measurable_sets[OF g] simp del: vimage_Int simp add: measurable_def) qed lemma measurable_comp: "f \ measurable M N \ g \ measurable N L \ g \ f \ measurable M L" using measurable_compose[of f M N g L] by (simp add: comp_def) lemma measurable_const: "c \ space M' \ (\x. c) \ measurable M M'" by (auto simp add: measurable_def) lemma measurable_ident: "id \ measurable M M" by (auto simp add: measurable_def) lemma measurable_id: "(\x. x) \ measurable M M" by (simp add: measurable_def) lemma measurable_ident_sets: assumes eq: "sets M = sets M'" shows "(\x. x) \ measurable M M'" using measurable_ident[of M] unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] . lemma sets_Least: assumes meas: "\i::nat. {x\space M. P i x} \ M" shows "(\x. LEAST j. P j x) -` A \ space M \ sets M" proof - { fix i have "(\x. LEAST j. P j x) -` {i} \ space M \ sets M" proof cases assume i: "(LEAST j. False) = i" have "(\x. LEAST j. P j x) -` {i} \ space M = {x\space M. P i x} \ (space M - (\jspace M. P j x})) \ (space M - (\i. {x\space M. P i x}))" by (simp add: set_eq_iff, safe) (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality) with meas show ?thesis by (auto intro!: sets.Int) next assume i: "(LEAST j. False) \ i" then have "(\x. LEAST j. P j x) -` {i} \ space M = {x\space M. P i x} \ (space M - (\jspace M. P j x}))" proof (simp add: set_eq_iff, safe) fix x assume neq: "(LEAST j. False) \ (LEAST j. P j x)" have "\j. P j x" by (rule ccontr) (insert neq, auto) then show "P (LEAST j. P j x) x" by (rule LeastI_ex) qed (auto dest: Least_le intro!: Least_equality) with meas show ?thesis by auto qed } then have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) \ sets M" by (intro sets.countable_UN) auto moreover have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) = (\x. LEAST j. P j x) -` A \ space M" by auto ultimately show ?thesis by auto qed lemma measurable_mono1: "M' \ Pow \ \ M \ M' \ measurable (measure_of \ M \) N \ measurable (measure_of \ M' \') N" using measure_of_subset[of M' \ M] by (auto simp add: measurable_def) subsubsection \Counting space\ definition\<^marker>\tag important\ count_space :: "'a set \ 'a measure" where "count_space \ = measure_of \ (Pow \) (\A. if finite A then of_nat (card A) else \)" lemma shows space_count_space[simp]: "space (count_space \) = \" and sets_count_space[simp]: "sets (count_space \) = Pow \" using sigma_sets_into_sp[of "Pow \" \] by (auto simp: count_space_def) lemma measurable_count_space_eq1[simp]: "f \ measurable (count_space A) M \ f \ A \ space M" unfolding measurable_def by simp lemma measurable_compose_countable': assumes f: "\i. i \ I \ (\x. f i x) \ measurable M N" and g: "g \ measurable M (count_space I)" and I: "countable I" shows "(\x. f (g x) x) \ measurable M N" unfolding measurable_def proof safe fix x assume "x \ space M" then show "f (g x) x \ space N" using measurable_space[OF f] g[THEN measurable_space] by auto next fix A assume A: "A \ sets N" have "(\x. f (g x) x) -` A \ space M = (\i\I. (g -` {i} \ space M) \ (f i -` A \ space M))" using measurable_space[OF g] by auto also have "\ \ sets M" using f[THEN measurable_sets, OF _ A] g[THEN measurable_sets] by (auto intro!: sets.countable_UN' I intro: sets.Int[OF measurable_sets measurable_sets]) finally show "(\x. f (g x) x) -` A \ space M \ sets M" . qed lemma measurable_count_space_eq_countable: assumes "countable A" shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" proof - { fix X assume "X \ A" "f \ space M \ A" with \countable A\ have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "countable X" by (auto dest: countable_subset) moreover assume "\a\A. f -` {a} \ space M \ sets M" ultimately have "f -` X \ space M \ sets M" using \X \ A\ by (auto intro!: sets.countable_UN' simp del: UN_simps) } then show ?thesis unfolding measurable_def by auto qed lemma measurable_count_space_eq2: "finite A \ f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" by (intro measurable_count_space_eq_countable countable_finite) lemma measurable_count_space_eq2_countable: fixes f :: "'a => 'c::countable" shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" by (intro measurable_count_space_eq_countable countableI_type) lemma measurable_compose_countable: assumes f: "\i::'i::countable. (\x. f i x) \ measurable M N" and g: "g \ measurable M (count_space UNIV)" shows "(\x. f (g x) x) \ measurable M N" by (rule measurable_compose_countable'[OF assms]) auto lemma measurable_count_space_const: "(\x. c) \ measurable M (count_space UNIV)" by (simp add: measurable_const) lemma measurable_count_space: "f \ measurable (count_space A) (count_space UNIV)" by simp lemma measurable_compose_rev: assumes f: "f \ measurable L N" and g: "g \ measurable M L" shows "(\x. f (g x)) \ measurable M N" using measurable_compose[OF g f] . lemma measurable_empty_iff: "space N = {} \ f \ measurable M N \ space M = {}" by (auto simp add: measurable_def Pi_iff) subsubsection\<^marker>\tag unimportant\ \Extend measure\ definition extend_measure :: "'a set \ 'b set \ ('b \ 'a set) \ ('b \ ennreal) \ 'a measure" where "extend_measure \ I G \ = (if (\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') \ \ (\i\I. \ i = 0) then measure_of \ (G`I) (SOME \'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') else measure_of \ (G`I) (\_. 0))" lemma space_extend_measure: "G ` I \ Pow \ \ space (extend_measure \ I G \) = \" unfolding extend_measure_def by simp lemma sets_extend_measure: "G ` I \ Pow \ \ sets (extend_measure \ I G \) = sigma_sets \ (G`I)" unfolding extend_measure_def by simp lemma emeasure_extend_measure: assumes M: "M = extend_measure \ I G \" and eq: "\i. i \ I \ \' (G i) = \ i" and ms: "G ` I \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" and "i \ I" shows "emeasure M (G i) = \ i" proof cases assume *: "(\i\I. \ i = 0)" with M have M_eq: "M = measure_of \ (G`I) (\_. 0)" by (simp add: extend_measure_def) from measure_space_0[OF ms(1)] ms \i\I\ have "emeasure M (G i) = 0" by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure) with \i\I\ * show ?thesis by simp next define P where "P \' \ (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" for \' assume "\ (\i\I. \ i = 0)" moreover have "measure_space (space M) (sets M) \'" using ms unfolding measure_space_def by auto standard with ms eq have "\\'. P \'" unfolding P_def by (intro exI[of _ \']) (auto simp add: M space_extend_measure sets_extend_measure) ultimately have M_eq: "M = measure_of \ (G`I) (Eps P)" by (simp add: M extend_measure_def P_def[symmetric]) from \\\'. P \'\ have P: "P (Eps P)" by (rule someI_ex) show "emeasure M (G i) = \ i" proof (subst emeasure_measure_of[OF M_eq]) have sets_M: "sets M = sigma_sets \ (G`I)" using M_eq ms by (auto simp: sets_extend_measure) then show "G i \ sets M" using \i \ I\ by auto show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \ i" using P \i\I\ by (auto simp add: sets_M measure_space_def P_def) qed fact qed lemma emeasure_extend_measure_Pair: assumes M: "M = extend_measure \ {(i, j). I i j} (\(i, j). G i j) (\(i, j). \ i j)" and eq: "\i j. I i j \ \' (G i j) = \ i j" and ms: "\i j. I i j \ G i j \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" and "I i j" shows "emeasure M (G i j) = \ i j" using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) \I i j\ by (auto simp: subset_eq) subsection \The smallest \\\-algebra regarding a function\ definition\<^marker>\tag important\ vimage_algebra :: "'a set \ ('a \ 'b) \ 'b measure \ 'a measure" where "vimage_algebra X f M = sigma X {f -` A \ X | A. A \ sets M}" lemma space_vimage_algebra[simp]: "space (vimage_algebra X f M) = X" unfolding vimage_algebra_def by (rule space_measure_of) auto lemma sets_vimage_algebra: "sets (vimage_algebra X f M) = sigma_sets X {f -` A \ X | A. A \ sets M}" unfolding vimage_algebra_def by (rule sets_measure_of) auto lemma sets_vimage_algebra2: "f \ X \ space M \ sets (vimage_algebra X f M) = {f -` A \ X | A. A \ sets M}" using sigma_sets_vimage_commute[of f X "space M" "sets M"] unfolding sets_vimage_algebra sets.sigma_sets_eq by simp lemma sets_vimage_algebra_cong: "sets M = sets N \ sets (vimage_algebra X f M) = sets (vimage_algebra X f N)" by (simp add: sets_vimage_algebra) lemma vimage_algebra_cong: assumes "X = Y" assumes "\x. x \ Y \ f x = g x" assumes "sets M = sets N" shows "vimage_algebra X f M = vimage_algebra Y g N" by (auto simp: vimage_algebra_def assms intro!: arg_cong2[where f=sigma]) lemma in_vimage_algebra: "A \ sets M \ f -` A \ X \ sets (vimage_algebra X f M)" by (auto simp: vimage_algebra_def) lemma sets_image_in_sets: assumes N: "space N = X" assumes f: "f \ measurable N M" shows "sets (vimage_algebra X f M) \ sets N" unfolding sets_vimage_algebra N[symmetric] by (rule sets.sigma_sets_subset) (auto intro!: measurable_sets f) lemma measurable_vimage_algebra1: "f \ X \ space M \ f \ measurable (vimage_algebra X f M) M" unfolding measurable_def by (auto intro: in_vimage_algebra) lemma measurable_vimage_algebra2: assumes g: "g \ space N \ X" and f: "(\x. f (g x)) \ measurable N M" shows "g \ measurable N (vimage_algebra X f M)" unfolding vimage_algebra_def proof (rule measurable_measure_of) fix A assume "A \ {f -` A \ X | A. A \ sets M}" then obtain Y where Y: "Y \ sets M" and A: "A = f -` Y \ X" by auto then have "g -` A \ space N = (\x. f (g x)) -` Y \ space N" using g by auto also have "\ \ sets N" using f Y by (rule measurable_sets) finally show "g -` A \ space N \ sets N" . qed (insert g, auto) lemma vimage_algebra_sigma: assumes X: "X \ Pow \'" and f: "f \ \ \ \'" shows "vimage_algebra \ f (sigma \' X) = sigma \ {f -` A \ \ | A. A \ X }" (is "?V = ?S") proof (rule measure_eqI) have \: "{f -` A \ \ |A. A \ X} \ Pow \" by auto show "sets ?V = sets ?S" using sigma_sets_vimage_commute[OF f, of X] by (simp add: space_measure_of_conv f sets_vimage_algebra2 \ X) qed (simp add: vimage_algebra_def emeasure_sigma) lemma vimage_algebra_vimage_algebra_eq: assumes *: "f \ X \ Y" "g \ Y \ space M" shows "vimage_algebra X f (vimage_algebra Y g M) = vimage_algebra X (\x. g (f x)) M" (is "?VV = ?V") proof (rule measure_eqI) have "(\x. g (f x)) \ X \ space M" "\A. A \ f -` Y \ X = A \ X" using * by auto with * show "sets ?VV = sets ?V" by (simp add: sets_vimage_algebra2 vimage_comp comp_def flip: ex_simps) qed (simp add: vimage_algebra_def emeasure_sigma) subsubsection \Restricted Space Sigma Algebra\ definition restrict_space :: "'a measure \ 'a set \ 'a measure" where "restrict_space M \ = measure_of (\ \ space M) (((\) \) ` sets M) (emeasure M)" lemma space_restrict_space: "space (restrict_space M \) = \ \ space M" using sets.sets_into_space unfolding restrict_space_def by (subst space_measure_of) auto lemma space_restrict_space2 [simp]: "\ \ sets M \ space (restrict_space M \) = \" by (simp add: space_restrict_space sets.sets_into_space) lemma sets_restrict_space: "sets (restrict_space M \) = ((\) \) ` sets M" unfolding restrict_space_def proof (subst sets_measure_of) show "(\) \ ` sets M \ Pow (\ \ space M)" by (auto dest: sets.sets_into_space) have "sigma_sets (\ \ space M) {((\x. x) -` X) \ (\ \ space M) | X. X \ sets M} = (\X. X \ (\ \ space M)) ` sets M" by (subst sigma_sets_vimage_commute[symmetric, where \' = "space M"]) (auto simp add: sets.sigma_sets_eq) moreover have "{((\x. x) -` X) \ (\ \ space M) | X. X \ sets M} = (\X. X \ (\ \ space M)) ` sets M" by auto moreover have "(\X. X \ (\ \ space M)) ` sets M = ((\) \) ` sets M" by (intro image_cong) (auto dest: sets.sets_into_space) ultimately show "sigma_sets (\ \ space M) ((\) \ ` sets M) = (\) \ ` sets M" by simp qed lemma restrict_space_sets_cong: "A = B \ sets M = sets N \ sets (restrict_space M A) = sets (restrict_space N B)" by (auto simp: sets_restrict_space) lemma sets_restrict_space_count_space : "sets (restrict_space (count_space A) B) = sets (count_space (A \ B))" by(auto simp add: sets_restrict_space) lemma sets_restrict_UNIV[simp]: "sets (restrict_space M UNIV) = sets M" by (auto simp add: sets_restrict_space) lemma sets_restrict_restrict_space: "sets (restrict_space (restrict_space M A) B) = sets (restrict_space M (A \ B))" unfolding sets_restrict_space image_comp by (intro image_cong) auto lemma sets_restrict_space_iff: "\ \ space M \ sets M \ A \ sets (restrict_space M \) \ (A \ \ \ A \ sets M)" proof (subst sets_restrict_space, safe) fix A assume "\ \ space M \ sets M" and A: "A \ sets M" then have "(\ \ space M) \ A \ sets M" by rule also have "(\ \ space M) \ A = \ \ A" using sets.sets_into_space[OF A] by auto finally show "\ \ A \ sets M" by auto qed auto lemma sets_restrict_space_cong: "sets M = sets N \ sets (restrict_space M \) = sets (restrict_space N \)" by (simp add: sets_restrict_space) lemma restrict_space_eq_vimage_algebra: "\ \ space M \ sets (restrict_space M \) = sets (vimage_algebra \ (\x. x) M)" unfolding restrict_space_def apply (subst sets_measure_of) apply (auto simp add: image_subset_iff dest: sets.sets_into_space) [] apply (auto simp add: sets_vimage_algebra intro!: arg_cong2[where f=sigma_sets]) done lemma sets_Collect_restrict_space_iff: assumes "S \ sets M" shows "{x\space (restrict_space M S). P x} \ sets (restrict_space M S) \ {x\space M. x \ S \ P x} \ sets M" proof - have "{x\S. P x} = {x\space M. x \ S \ P x}" using sets.sets_into_space[OF assms] by auto then show ?thesis by (subst sets_restrict_space_iff) (auto simp add: space_restrict_space assms) qed lemma measurable_restrict_space1: assumes f: "f \ measurable M N" shows "f \ measurable (restrict_space M \) N" unfolding measurable_def proof (intro CollectI conjI ballI) show sp: "f \ space (restrict_space M \) \ space N" using measurable_space[OF f] by (auto simp: space_restrict_space) fix A assume "A \ sets N" have "f -` A \ space (restrict_space M \) = (f -` A \ space M) \ (\ \ space M)" by (auto simp: space_restrict_space) also have "\ \ sets (restrict_space M \)" unfolding sets_restrict_space using measurable_sets[OF f \A \ sets N\] by blast finally show "f -` A \ space (restrict_space M \) \ sets (restrict_space M \)" . qed lemma measurable_restrict_space2_iff: "f \ measurable M (restrict_space N \) \ (f \ measurable M N \ f \ space M \ \)" proof - have "\A. f \ space M \ \ \ f -` \ \ f -` A \ space M = f -` A \ space M" by auto then show ?thesis by (auto simp: measurable_def space_restrict_space Pi_Int[symmetric] sets_restrict_space) qed lemma measurable_restrict_space2: "f \ space M \ \ \ f \ measurable M N \ f \ measurable M (restrict_space N \)" by (simp add: measurable_restrict_space2_iff) lemma measurable_piecewise_restrict: assumes I: "countable C" and X: "\\. \ \ C \ \ \ space M \ sets M" "space M \ \C" and f: "\\. \ \ C \ f \ measurable (restrict_space M \) N" shows "f \ measurable M N" proof (rule measurableI) fix x assume "x \ space M" with X obtain \ where "\ \ C" "x \ \" "x \ space M" by auto then show "f x \ space N" by (auto simp: space_restrict_space intro: f measurable_space) next fix A assume A: "A \ sets N" have "f -` A \ space M = (\\\C. (f -` A \ (\ \ space M)))" using X by (auto simp: subset_eq) also have "\ \ sets M" using measurable_sets[OF f A] X I by (intro sets.countable_UN') (auto simp: sets_restrict_space_iff space_restrict_space) finally show "f -` A \ space M \ sets M" . qed lemma measurable_piecewise_restrict_iff: "countable C \ (\\. \ \ C \ \ \ space M \ sets M) \ space M \ (\C) \ f \ measurable M N \ (\\\C. f \ measurable (restrict_space M \) N)" by (auto intro: measurable_piecewise_restrict measurable_restrict_space1) lemma measurable_If_restrict_space_iff: "{x\space M. P x} \ sets M \ (\x. if P x then f x else g x) \ measurable M N \ (f \ measurable (restrict_space M {x. P x}) N \ g \ measurable (restrict_space M {x. \ P x}) N)" by (subst measurable_piecewise_restrict_iff[where C="{{x. P x}, {x. \ P x}}"]) (auto simp: Int_def sets.sets_Collect_neg space_restrict_space conj_commute[of _ "x \ space M" for x] cong: measurable_cong') lemma measurable_If: "f \ measurable M M' \ g \ measurable M M' \ {x\space M. P x} \ sets M \ (\x. if P x then f x else g x) \ measurable M M'" unfolding measurable_If_restrict_space_iff by (auto intro: measurable_restrict_space1) lemma measurable_If_set: assumes measure: "f \ measurable M M'" "g \ measurable M M'" assumes P: "A \ space M \ sets M" shows "(\x. if x \ A then f x else g x) \ measurable M M'" proof (rule measurable_If[OF measure]) have "{x \ space M. x \ A} = A \ space M" by auto thus "{x \ space M. x \ A} \ sets M" using \A \ space M \ sets M\ by auto qed lemma measurable_restrict_space_iff: "\ \ space M \ sets M \ c \ space N \ f \ measurable (restrict_space M \) N \ (\x. if x \ \ then f x else c) \ measurable M N" by (subst measurable_If_restrict_space_iff) (simp_all add: Int_def conj_commute measurable_const) lemma restrict_space_singleton: "{x} \ sets M \ sets (restrict_space M {x}) = sets (count_space {x})" using sets_restrict_space_iff[of "{x}" M] by (auto simp add: sets_restrict_space_iff dest!: subset_singletonD) lemma measurable_restrict_countable: assumes X[intro]: "countable X" assumes sets[simp]: "\x. x \ X \ {x} \ sets M" assumes space[simp]: "\x. x \ X \ f x \ space N" assumes f: "f \ measurable (restrict_space M (- X)) N" shows "f \ measurable M N" using f sets.countable[OF sets X] by (intro measurable_piecewise_restrict[where M=M and C="{- X} \ ((\x. {x}) ` X)"]) (auto simp: Diff_Int_distrib2 Compl_eq_Diff_UNIV Int_insert_left sets.Diff restrict_space_singleton simp del: sets_count_space cong: measurable_cong_sets) lemma measurable_discrete_difference: assumes f: "f \ measurable M N" assumes X: "countable X" "\x. x \ X \ {x} \ sets M" "\x. x \ X \ g x \ space N" assumes eq: "\x. x \ space M \ x \ X \ f x = g x" shows "g \ measurable M N" by (rule measurable_restrict_countable[OF X]) (auto simp: eq[symmetric] space_restrict_space cong: measurable_cong' intro: f measurable_restrict_space1) lemma measurable_count_space_extend: "A \ B \ f \ space M \ A \ f \ M \\<^sub>M count_space B \ f \ M \\<^sub>M count_space A" by (auto simp: measurable_def) end diff --git a/src/HOL/Analysis/Summation_Tests.thy b/src/HOL/Analysis/Summation_Tests.thy --- a/src/HOL/Analysis/Summation_Tests.thy +++ b/src/HOL/Analysis/Summation_Tests.thy @@ -1,901 +1,901 @@ (* Title: HOL/Analysis/Summation_Tests.thy Author: Manuel Eberl, TU München *) section \Radius of Convergence and Summation Tests\ theory Summation_Tests imports Complex_Main "HOL-Library.Discrete" "HOL-Library.Extended_Real" "HOL-Library.Liminf_Limsup" Extended_Real_Limits begin text \ The definition of the radius of convergence of a power series, various summability tests, lemmas to compute the radius of convergence etc. \ subsection \Convergence tests for infinite sums\ subsubsection \Root test\ lemma limsup_root_powser: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" shows "limsup (\n. ereal (root n (norm (f n * z ^ n)))) = limsup (\n. ereal (root n (norm (f n)))) * ereal (norm z)" proof - have A: "(\n. ereal (root n (norm (f n * z ^ n)))) = (\n. ereal (root n (norm (f n))) * ereal (norm z))" (is "?g = ?h") proof fix n show "?g n = ?h n" by (cases "n = 0") (simp_all add: norm_mult real_root_mult real_root_pos2 norm_power) qed show ?thesis by (subst A, subst limsup_ereal_mult_right) simp_all qed lemma limsup_root_limit: assumes "(\n. ereal (root n (norm (f n)))) \ l" (is "?g \ _") shows "limsup (\n. ereal (root n (norm (f n)))) = l" proof - from assms have "convergent ?g" "lim ?g = l" unfolding convergent_def by (blast intro: limI)+ with convergent_limsup_cl show ?thesis by force qed lemma limsup_root_limit': assumes "(\n. root n (norm (f n))) \ l" shows "limsup (\n. ereal (root n (norm (f n)))) = ereal l" by (intro limsup_root_limit tendsto_ereal assms) theorem root_test_convergence': fixes f :: "nat \ 'a :: banach" defines "l \ limsup (\n. ereal (root n (norm (f n))))" assumes l: "l < 1" shows "summable f" proof - have "0 = limsup (\n. 0)" by (simp add: Limsup_const) also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally have "l \ 0" by simp with l obtain l' where l': "l = ereal l'" by (cases l) simp_all define c where "c = (1 - l') / 2" from l and \l \ 0\ have c: "l + c > l" "l' + c \ 0" "l' + c < 1" unfolding c_def by (simp_all add: field_simps l') have "\C>l. eventually (\n. ereal (root n (norm (f n))) < C) sequentially" by (subst Limsup_le_iff[symmetric]) (simp add: l_def) with c have "eventually (\n. ereal (root n (norm (f n))) < l + ereal c) sequentially" by simp with eventually_gt_at_top[of "0::nat"] have "eventually (\n. norm (f n) \ (l' + c) ^ n) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" assume "ereal (root n (norm (f n))) < l + ereal c" hence "root n (norm (f n)) \ l' + c" by (simp add: l') with c n have "root n (norm (f n)) ^ n \ (l' + c) ^ n" by (intro power_mono) (simp_all add: real_root_ge_zero) also from n have "root n (norm (f n)) ^ n = norm (f n)" by simp finally show "norm (f n) \ (l' + c) ^ n" by simp qed thus ?thesis by (rule summable_comparison_test_ev[OF _ summable_geometric]) (simp add: c) qed theorem root_test_divergence: fixes f :: "nat \ 'a :: banach" defines "l \ limsup (\n. ereal (root n (norm (f n))))" assumes l: "l > 1" shows "\summable f" proof assume "summable f" hence bounded: "Bseq f" by (simp add: summable_imp_Bseq) have "0 = limsup (\n. 0)" by (simp add: Limsup_const) also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally have l_nonneg: "l \ 0" by simp define c where "c = (if l = \ then 2 else 1 + (real_of_ereal l - 1) / 2)" from l l_nonneg consider "l = \" | "\l'. l = ereal l'" by (cases l) simp_all hence c: "c > 1 \ ereal c < l" by cases (insert l, auto simp: c_def field_simps) have unbounded: "\bdd_above {n. root n (norm (f n)) > c}" proof assume "bdd_above {n. root n (norm (f n)) > c}" then obtain N where "\n. root n (norm (f n)) > c \ n \ N" unfolding bdd_above_def by blast hence "\N. \n\N. root n (norm (f n)) \ c" by (intro exI[of _ "N + 1"]) (force simp: not_less_eq_eq[symmetric]) hence "eventually (\n. root n (norm (f n)) \ c) sequentially" by (auto simp: eventually_at_top_linorder) hence "l \ c" unfolding l_def by (intro Limsup_bounded) simp_all with c show False by auto qed from bounded obtain K where K: "K > 0" "\n. norm (f n) \ K" using BseqE by blast define n where "n = nat \log c K\" from unbounded have "\m>n. c < root m (norm (f m))" unfolding bdd_above_def by (auto simp: not_le) - then guess m by (elim exE conjE) note m = this + then obtain m where m: "n < m" "c < root m (norm (f m))" by auto from c K have "K = c powr log c K" by (simp add: powr_def log_def) also from c have "c powr log c K \ c powr real n" unfolding n_def by (intro powr_mono, linarith, simp) finally have "K \ c ^ n" using c by (simp add: powr_realpow) also from c m have "c ^ n < c ^ m" by simp also from c m have "c ^ m < root m (norm (f m)) ^ m" by (intro power_strict_mono) simp_all also from m have "... = norm (f m)" by simp finally show False using K(2)[of m] by simp qed subsubsection \Cauchy's condensation test\ context fixes f :: "nat \ real" begin private lemma condensation_inequality: assumes mono: "\m n. 0 < m \ m \ n \ f n \ f m" shows "(\k=1.. (\k=1..k=1.. (\k=1..k=1..<2^n. f (2 ^ Discrete.log k)) = (\k {2^n..<(2^Suc n :: nat)}" by auto also have "(\k\\. f (2 ^ Discrete.log k)) = (\kk = 2^n..<2^Suc n. f (2^Discrete.log k))" by (subst sum.union_disjoint) (insert Suc, auto) also have "Discrete.log k = n" if "k \ {2^n..<2^Suc n}" for k using that by (intro Discrete.log_eqI) simp_all hence "(\k = 2^n..<2^Suc n. f (2^Discrete.log k)) = (\(_::nat) = 2^n..<2^Suc n. f (2^n))" by (intro sum.cong) simp_all also have "\ = 2^n * f (2^n)" by (simp) finally show ?case by simp qed simp private lemma condensation_condense2: "(\k=1..<2^n. f (2 * 2 ^ Discrete.log k)) = (\k {2^n..<(2^Suc n :: nat)}" by auto also have "(\k\\. f (2 * 2 ^ Discrete.log k)) = (\kk = 2^n..<2^Suc n. f (2 * 2^Discrete.log k))" by (subst sum.union_disjoint) (insert Suc, auto) also have "Discrete.log k = n" if "k \ {2^n..<2^Suc n}" for k using that by (intro Discrete.log_eqI) simp_all hence "(\k = 2^n..<2^Suc n. f (2*2^Discrete.log k)) = (\(_::nat) = 2^n..<2^Suc n. f (2^Suc n))" by (intro sum.cong) simp_all also have "\ = 2^n * f (2^Suc n)" by (simp) finally show ?case by simp qed simp theorem condensation_test: assumes mono: "\m. 0 < m \ f (Suc m) \ f m" assumes nonneg: "\n. f n \ 0" shows "summable f \ summable (\n. 2^n * f (2^n))" proof - define f' where "f' n = (if n = 0 then 0 else f n)" for n from mono have mono': "decseq (\n. f (Suc n))" by (intro decseq_SucI) simp hence mono': "f n \ f m" if "m \ n" "m > 0" for m n using that decseqD[OF mono', of "m - 1" "n - 1"] by simp have "(\n. f (Suc n)) = (\n. f' (Suc n))" by (intro ext) (simp add: f'_def) hence "summable f \ summable f'" by (subst (1 2) summable_Suc_iff [symmetric]) (simp only:) also have "\ \ convergent (\n. \kn. \kn. \k Bseq (\n. \k \ Bseq (\n. \k=1.. \ Bseq (\n. \k=1.. \ Bseq (\n. \k=1..<2^n. f k)" by (rule nonneg_incseq_Bseq_subseq_iff[symmetric]) (auto intro!: sum_nonneg incseq_SucI nonneg simp: strict_mono_def) also have "\ \ Bseq (\n. \kn. \k=1..<2^n. f k)" have "eventually (\n. norm (\k norm (\k=1..<2^n. f k)) sequentially" proof (intro always_eventually allI) fix n :: nat have "norm (\kk \ (\k=1..<2^n. f k)" by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono') also have "\ = norm \" unfolding real_norm_def by (intro abs_of_nonneg[symmetric] sum_nonneg ballI mult_nonneg_nonneg nonneg) finally show "norm (\k norm (\k=1..<2^n. f k)" . qed from this and A have "Bseq (\n. \kn. \kn. (\k=Suc 0..n. (\k=0..n. (\kn. (\kn. norm (\k=1..<2^n. f k) \ norm (\kk=1..<2^n. f k) = (\k=1..<2^n. f k)" unfolding real_norm_def by (intro abs_of_nonneg sum_nonneg ballI mult_nonneg_nonneg nonneg) also have "\ \ (\k = norm \" unfolding real_norm_def by (intro abs_of_nonneg [symmetric] sum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all finally show "norm (\k=1..<2^n. f k) \ norm (\kn. \k=1..<2^n. f k)" by (rule Bseq_eventually_mono) qed also have "monoseq (\n. (\kn. (\k convergent (\n. (\k \ summable (\k. 2^k * f (2^k))" by (simp only: summable_iff_convergent) finally show ?thesis . qed end subsubsection \Summability of powers\ lemma abs_summable_complex_powr_iff: "summable (\n. norm (exp (of_real (ln (of_nat n)) * s))) \ Re s < -1" proof (cases "Re s \ 0") let ?l = "\n. complex_of_real (ln (of_nat n))" case False have "eventually (\n. norm (1 :: real) \ norm (exp (?l n * s))) sequentially" apply (rule eventually_mono [OF eventually_gt_at_top[of "0::nat"]]) using False ge_one_powr_ge_zero by auto from summable_comparison_test_ev[OF this] False show ?thesis by (auto simp: summable_const_iff) next let ?l = "\n. complex_of_real (ln (of_nat n))" case True hence "summable (\n. norm (exp (?l n * s))) \ summable (\n. 2^n * norm (exp (?l (2^n) * s)))" by (intro condensation_test) (auto intro!: mult_right_mono_neg) also have "(\n. 2^n * norm (exp (?l (2^n) * s))) = (\n. (2 powr (Re s + 1)) ^ n)" proof fix n :: nat have "2^n * norm (exp (?l (2^n) * s)) = exp (real n * ln 2) * exp (real n * ln 2 * Re s)" using True by (subst exp_of_nat_mult) (simp add: ln_realpow algebra_simps) also have "\ = exp (real n * (ln 2 * (Re s + 1)))" by (simp add: algebra_simps exp_add) also have "\ = exp (ln 2 * (Re s + 1)) ^ n" by (subst exp_of_nat_mult) simp also have "exp (ln 2 * (Re s + 1)) = 2 powr (Re s + 1)" by (simp add: powr_def) finally show "2^n * norm (exp (?l (2^n) * s)) = (2 powr (Re s + 1)) ^ n" . qed also have "summable \ \ 2 powr (Re s + 1) < 2 powr 0" by (subst summable_geometric_iff) simp also have "\ \ Re s < -1" by (subst powr_less_cancel_iff) (simp, linarith) finally show ?thesis . qed theorem summable_complex_powr_iff: assumes "Re s < -1" shows "summable (\n. exp (of_real (ln (of_nat n)) * s))" by (rule summable_norm_cancel, subst abs_summable_complex_powr_iff) fact lemma summable_real_powr_iff: "summable (\n. of_nat n powr s :: real) \ s < -1" proof - from eventually_gt_at_top[of "0::nat"] have "summable (\n. of_nat n powr s) \ summable (\n. exp (ln (of_nat n) * s))" by (intro summable_cong) (auto elim!: eventually_mono simp: powr_def) also have "\ \ s < -1" using abs_summable_complex_powr_iff[of "of_real s"] by simp finally show ?thesis . qed lemma inverse_power_summable: assumes s: "s \ 2" shows "summable (\n. inverse (of_nat n ^ s :: 'a :: {real_normed_div_algebra,banach}))" proof (rule summable_norm_cancel, subst summable_cong) from eventually_gt_at_top[of "0::nat"] show "eventually (\n. norm (inverse (of_nat n ^ s:: 'a)) = real_of_nat n powr (-real s)) at_top" by eventually_elim (simp add: norm_inverse norm_power powr_minus powr_realpow) qed (insert s summable_real_powr_iff[of "-s"], simp_all) lemma not_summable_harmonic: "\summable (\n. inverse (of_nat n) :: 'a :: real_normed_field)" proof assume "summable (\n. inverse (of_nat n) :: 'a)" hence "convergent (\n. norm (of_real (\kn. abs (\kn. abs (\kn. \k \ summable (\k. inverse (of_nat k) :: real)" by (simp add: summable_iff_convergent) finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus) qed subsubsection \Kummer's test\ theorem kummers_test_convergence: fixes f p :: "nat \ real" assumes pos_f: "eventually (\n. f n > 0) sequentially" assumes nonneg_p: "eventually (\n. p n \ 0) sequentially" defines "l \ liminf (\n. ereal (p n * f n / f (Suc n) - p (Suc n)))" assumes l: "l > 0" shows "summable f" unfolding summable_iff_convergent' proof - define r where "r = (if l = \ then 1 else real_of_ereal l / 2)" from l have "r > 0 \ of_real r < l" by (cases l) (simp_all add: r_def) hence r: "r > 0" "of_real r < l" by simp_all hence "eventually (\n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially" unfolding l_def by (force dest: less_LiminfD) moreover from pos_f have "eventually (\n. f (Suc n) > 0) sequentially" by (subst eventually_sequentially_Suc) ultimately have "eventually (\n. p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)) sequentially" by eventually_elim (simp add: field_simps) from eventually_conj[OF pos_f eventually_conj[OF nonneg_p this]] obtain m where m: "\n. n \ m \ f n > 0" "\n. n \ m \ p n \ 0" "\n. n \ m \ p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)" unfolding eventually_at_top_linorder by blast let ?c = "(norm (\k\m. r * f k) + p m * f m) / r" have "Bseq (\n. (\k\n + Suc m. f k))" proof (rule BseqI') fix k :: nat define n where "n = k + Suc m" have n: "n > m" by (simp add: n_def) from r have "r * norm (\k\n. f k) = norm (\k\n. r * f k)" by (simp add: sum_distrib_left[symmetric] abs_mult) also from n have "{..n} = {..m} \ {Suc m..n}" by auto hence "(\k\n. r * f k) = (\k\{..m} \ {Suc m..n}. r * f k)" by (simp only:) also have "\ = (\k\m. r * f k) + (\k=Suc m..n. r * f k)" by (subst sum.union_disjoint) auto also have "norm \ \ norm (\k\m. r * f k) + norm (\k=Suc m..n. r * f k)" by (rule norm_triangle_ineq) also from r less_imp_le[OF m(1)] have "(\k=Suc m..n. r * f k) \ 0" by (intro sum_nonneg) auto hence "norm (\k=Suc m..n. r * f k) = (\k=Suc m..n. r * f k)" by simp also have "(\k=Suc m..n. r * f k) = (\k=m.. \ (\k=m.. = -(\k=m.. = p m * f m - p n * f n" by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst sum_Suc_diff) simp_all also from less_imp_le[OF m(1)] m(2) n have "\ \ p m * f m" by simp finally show "norm (\k\n. f k) \ (norm (\k\m. r * f k) + p m * f m) / r" using r by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac) qed moreover have "(\k\n. f k) \ (\k\n'. f k)" if "Suc m \ n" "n \ n'" for n n' using less_imp_le[OF m(1)] that by (intro sum_mono2) auto ultimately show "convergent (\n. \k\n. f k)" by (rule Bseq_monoseq_convergent'_inc) qed theorem kummers_test_divergence: fixes f p :: "nat \ real" assumes pos_f: "eventually (\n. f n > 0) sequentially" assumes pos_p: "eventually (\n. p n > 0) sequentially" assumes divergent_p: "\summable (\n. inverse (p n))" defines "l \ limsup (\n. ereal (p n * f n / f (Suc n) - p (Suc n)))" assumes l: "l < 0" shows "\summable f" proof assume "summable f" from eventually_conj[OF pos_f eventually_conj[OF pos_p Limsup_lessD[OF l[unfolded l_def]]]] obtain N where N: "\n. n \ N \ p n > 0" "\n. n \ N \ f n > 0" "\n. n \ N \ p n * f n / f (Suc n) - p (Suc n) < 0" by (auto simp: eventually_at_top_linorder) hence A: "p n * f n < p (Suc n) * f (Suc n)" if "n \ N" for n using that N[of n] N[of "Suc n"] by (simp add: field_simps) have B: "p n * f n \ p N * f N" if "n \ N" for n using that and A by (induction n rule: dec_induct) (auto intro!: less_imp_le elim!: order.trans) have "eventually (\n. norm (p N * f N * inverse (p n)) \ f n) sequentially" apply (rule eventually_mono [OF eventually_ge_at_top[of N]]) using B N by (auto simp: field_simps abs_of_pos) from this and \summable f\ have "summable (\n. p N * f N * inverse (p n))" by (rule summable_comparison_test_ev) from summable_mult[OF this, of "inverse (p N * f N)"] N[OF le_refl] have "summable (\n. inverse (p n))" by (simp add: field_split_simps) with divergent_p show False by contradiction qed subsubsection \Ratio test\ theorem ratio_test_convergence: fixes f :: "nat \ real" assumes pos_f: "eventually (\n. f n > 0) sequentially" defines "l \ liminf (\n. ereal (f n / f (Suc n)))" assumes l: "l > 1" shows "summable f" proof (rule kummers_test_convergence[OF pos_f]) note l also have "l = liminf (\n. ereal (f n / f (Suc n) - 1)) + 1" by (subst Liminf_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def) finally show "liminf (\n. ereal (1 * f n / f (Suc n) - 1)) > 0" by (cases "liminf (\n. ereal (1 * f n / f (Suc n) - 1))") simp_all qed simp theorem ratio_test_divergence: fixes f :: "nat \ real" assumes pos_f: "eventually (\n. f n > 0) sequentially" defines "l \ limsup (\n. ereal (f n / f (Suc n)))" assumes l: "l < 1" shows "\summable f" proof (rule kummers_test_divergence[OF pos_f]) have "limsup (\n. ereal (f n / f (Suc n) - 1)) + 1 = l" by (subst Limsup_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def) also note l finally show "limsup (\n. ereal (1 * f n / f (Suc n) - 1)) < 0" by (cases "limsup (\n. ereal (1 * f n / f (Suc n) - 1))") simp_all qed (simp_all add: summable_const_iff) subsubsection \Raabe's test\ theorem raabes_test_convergence: fixes f :: "nat \ real" assumes pos: "eventually (\n. f n > 0) sequentially" defines "l \ liminf (\n. ereal (of_nat n * (f n / f (Suc n) - 1)))" assumes l: "l > 1" shows "summable f" proof (rule kummers_test_convergence) let ?l' = "liminf (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))" have "1 < l" by fact also have "l = liminf (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)" by (simp add: l_def algebra_simps) also have "\ = ?l' + 1" by (subst Liminf_add_ereal_right) simp_all finally show "?l' > 0" by (cases ?l') (simp_all add: algebra_simps) qed (simp_all add: pos) theorem raabes_test_divergence: fixes f :: "nat \ real" assumes pos: "eventually (\n. f n > 0) sequentially" defines "l \ limsup (\n. ereal (of_nat n * (f n / f (Suc n) - 1)))" assumes l: "l < 1" shows "\summable f" proof (rule kummers_test_divergence) let ?l' = "limsup (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))" note l also have "l = limsup (\n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)" by (simp add: l_def algebra_simps) also have "\ = ?l' + 1" by (subst Limsup_add_ereal_right) simp_all finally show "?l' < 0" by (cases ?l') (simp_all add: algebra_simps) qed (insert pos eventually_gt_at_top[of "0::nat"] not_summable_harmonic, simp_all) subsection \Radius of convergence\ text \ The radius of convergence of a power series. This value always exists, ranges from \<^term>\0::ereal\ to \<^term>\\::ereal\, and the power series is guaranteed to converge for all inputs with a norm that is smaller than that radius and to diverge for all inputs with a norm that is greater. \ definition\<^marker>\tag important\ conv_radius :: "(nat \ 'a :: banach) \ ereal" where "conv_radius f = inverse (limsup (\n. ereal (root n (norm (f n)))))" lemma conv_radius_cong_weak [cong]: "(\n. f n = g n) \ conv_radius f = conv_radius g" by (drule ext) simp_all lemma conv_radius_nonneg: "conv_radius f \ 0" proof - have "0 = limsup (\n. 0)" by (subst Limsup_const) simp_all also have "\ \ limsup (\n. ereal (root n (norm (f n))))" by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally show ?thesis unfolding conv_radius_def by (auto simp: ereal_inverse_nonneg_iff) qed lemma conv_radius_zero [simp]: "conv_radius (\_. 0) = \" by (auto simp: conv_radius_def zero_ereal_def [symmetric] Limsup_const) lemma conv_radius_altdef: "conv_radius f = liminf (\n. inverse (ereal (root n (norm (f n)))))" by (subst Liminf_inverse_ereal) (simp_all add: real_root_ge_zero conv_radius_def) lemma conv_radius_cong': assumes "eventually (\x. f x = g x) sequentially" shows "conv_radius f = conv_radius g" unfolding conv_radius_altdef by (intro Liminf_eq eventually_mono [OF assms]) auto lemma conv_radius_cong: assumes "eventually (\x. norm (f x) = norm (g x)) sequentially" shows "conv_radius f = conv_radius g" unfolding conv_radius_altdef by (intro Liminf_eq eventually_mono [OF assms]) auto theorem abs_summable_in_conv_radius: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "ereal (norm z) < conv_radius f" shows "summable (\n. norm (f n * z ^ n))" proof (rule root_test_convergence') define l where "l = limsup (\n. ereal (root n (norm (f n))))" have "0 = limsup (\n. 0)" by (simp add: Limsup_const) also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally have l_nonneg: "l \ 0" . have "limsup (\n. root n (norm (f n * z^n))) = l * ereal (norm z)" unfolding l_def by (rule limsup_root_powser) also from l_nonneg consider "l = 0" | "l = \" | "\l'. l = ereal l' \ l' > 0" by (cases "l") (auto simp: less_le) hence "l * ereal (norm z) < 1" proof cases assume "l = \" hence "conv_radius f = 0" unfolding conv_radius_def l_def by simp with assms show ?thesis by simp next assume "\l'. l = ereal l' \ l' > 0" - then guess l' by (elim exE conjE) note l' = this + then obtain l' where l': "l = ereal l'" "0 < l'" by auto hence "l \ \" by auto have "l * ereal (norm z) < l * conv_radius f" by (intro ereal_mult_strict_left_mono) (simp_all add: l' assms) also have "conv_radius f = inverse l" by (simp add: conv_radius_def l_def) also from l' have "l * inverse l = 1" by simp finally show ?thesis . qed simp_all finally show "limsup (\n. ereal (root n (norm (norm (f n * z ^ n))))) < 1" by simp qed lemma summable_in_conv_radius: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "ereal (norm z) < conv_radius f" shows "summable (\n. f n * z ^ n)" by (rule summable_norm_cancel, rule abs_summable_in_conv_radius) fact+ theorem not_summable_outside_conv_radius: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "ereal (norm z) > conv_radius f" shows "\summable (\n. f n * z ^ n)" proof (rule root_test_divergence) define l where "l = limsup (\n. ereal (root n (norm (f n))))" have "0 = limsup (\n. 0)" by (simp add: Limsup_const) also have "... \ l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero) finally have l_nonneg: "l \ 0" . from assms have l_nz: "l \ 0" unfolding conv_radius_def l_def by auto have "limsup (\n. ereal (root n (norm (f n * z^n)))) = l * ereal (norm z)" unfolding l_def by (rule limsup_root_powser) also have "... > 1" proof (cases l) assume "l = \" with assms conv_radius_nonneg[of f] show ?thesis by (auto simp: zero_ereal_def[symmetric]) next fix l' assume l': "l = ereal l'" from l_nonneg l_nz have "1 = l * inverse l" by (auto simp: l' field_simps) also from l_nz have "inverse l = conv_radius f" unfolding l_def conv_radius_def by auto also from l' l_nz l_nonneg assms have "l * \ < l * ereal (norm z)" by (intro ereal_mult_strict_left_mono) (auto simp: l') finally show ?thesis . qed (insert l_nonneg, simp_all) finally show "limsup (\n. ereal (root n (norm (f n * z^n)))) > 1" . qed lemma conv_radius_geI: assumes "summable (\n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})" shows "conv_radius f \ norm z" using not_summable_outside_conv_radius[of f z] assms by (force simp: not_le[symmetric]) lemma conv_radius_leI: assumes "\summable (\n. norm (f n * z ^ n :: 'a :: {banach, real_normed_div_algebra}))" shows "conv_radius f \ norm z" using abs_summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric]) lemma conv_radius_leI': assumes "\summable (\n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})" shows "conv_radius f \ norm z" using summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric]) lemma conv_radius_geI_ex: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "\r. 0 < r \ ereal r < R \ \z. norm z = r \ summable (\n. f n * z^n)" shows "conv_radius f \ R" proof (rule linorder_cases[of "conv_radius f" R]) assume R: "conv_radius f < R" with conv_radius_nonneg[of f] obtain conv_radius' where [simp]: "conv_radius f = ereal conv_radius'" by (cases "conv_radius f") simp_all define r where "r = (if R = \ then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2)" from R conv_radius_nonneg[of f] have "0 < r \ ereal r < R \ ereal r > conv_radius f" unfolding r_def by (cases R) (auto simp: r_def field_simps) with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\n. f n * z^n)" by auto with not_summable_outside_conv_radius[of f z] show ?thesis by simp qed simp_all lemma conv_radius_geI_ex': fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "\r. 0 < r \ ereal r < R \ summable (\n. f n * of_real r^n)" shows "conv_radius f \ R" proof (rule conv_radius_geI_ex) fix r assume "0 < r" "ereal r < R" with assms[of r] show "\z. norm z = r \ summable (\n. f n * z ^ n)" by (intro exI[of _ "of_real r :: 'a"]) auto qed lemma conv_radius_leI_ex: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "R \ 0" assumes "\r. 0 < r \ ereal r > R \ \z. norm z = r \ \summable (\n. norm (f n * z^n))" shows "conv_radius f \ R" proof (rule linorder_cases[of "conv_radius f" R]) assume R: "conv_radius f > R" from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all define r where "r = (if conv_radius f = \ then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2)" from R conv_radius_nonneg[of f] have "r > R \ r < conv_radius f" unfolding r_def by (cases "conv_radius f") (auto simp: r_def field_simps R') with assms(1) assms(2)[of r] R' obtain z where "norm z < conv_radius f" "\summable (\n. norm (f n * z^n))" by auto with abs_summable_in_conv_radius[of z f] show ?thesis by auto qed simp_all lemma conv_radius_leI_ex': fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "R \ 0" assumes "\r. 0 < r \ ereal r > R \ \summable (\n. f n * of_real r^n)" shows "conv_radius f \ R" proof (rule conv_radius_leI_ex) fix r assume "0 < r" "ereal r > R" with assms(2)[of r] show "\z. norm z = r \ \summable (\n. norm (f n * z ^ n))" by (intro exI[of _ "of_real r :: 'a"]) (auto dest: summable_norm_cancel) qed fact+ lemma conv_radius_eqI: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "R \ 0" assumes "\r. 0 < r \ ereal r < R \ \z. norm z = r \ summable (\n. f n * z^n)" assumes "\r. 0 < r \ ereal r > R \ \z. norm z = r \ \summable (\n. norm (f n * z^n))" shows "conv_radius f = R" by (intro antisym conv_radius_geI_ex conv_radius_leI_ex assms) lemma conv_radius_eqI': fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" assumes "R \ 0" assumes "\r. 0 < r \ ereal r < R \ summable (\n. f n * (of_real r)^n)" assumes "\r. 0 < r \ ereal r > R \ \summable (\n. norm (f n * (of_real r)^n))" shows "conv_radius f = R" proof (intro conv_radius_eqI[OF assms(1)]) fix r assume "0 < r" "ereal r < R" with assms(2)[OF this] show "\z. norm z = r \ summable (\n. f n * z ^ n)" by force next fix r assume "0 < r" "ereal r > R" with assms(3)[OF this] show "\z. norm z = r \ \summable (\n. norm (f n * z ^ n))" by force qed lemma conv_radius_zeroI: fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes "\z. z \ 0 \ \summable (\n. f n * z^n)" shows "conv_radius f = 0" proof (rule ccontr) assume "conv_radius f \ 0" with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp define r where "r = (if conv_radius f = \ then 1 else real_of_ereal (conv_radius f) / 2)" from pos have r: "ereal r > 0 \ ereal r < conv_radius f" by (cases "conv_radius f") (simp_all add: r_def) hence "summable (\n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp moreover from r and assms[of "of_real r"] have "\summable (\n. f n * of_real r ^ n)" by simp ultimately show False by contradiction qed lemma conv_radius_inftyI': fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes "\r. r > c \ \z. norm z = r \ summable (\n. f n * z^n)" shows "conv_radius f = \" proof - { fix r :: real have "max r (c + 1) > c" by (auto simp: max_def) from assms[OF this] obtain z where "norm z = max r (c + 1)" "summable (\n. f n * z^n)" by blast from conv_radius_geI[OF this(2)] this(1) have "conv_radius f \ r" by simp } from this[of "real_of_ereal (conv_radius f + 1)"] show "conv_radius f = \" by (cases "conv_radius f") simp_all qed lemma conv_radius_inftyI: fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes "\r. \z. norm z = r \ summable (\n. f n * z^n)" shows "conv_radius f = \" using assms by (rule conv_radius_inftyI') lemma conv_radius_inftyI'': fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes "\z. summable (\n. f n * z^n)" shows "conv_radius f = \" proof (rule conv_radius_inftyI') fix r :: real assume "r > 0" with assms show "\z. norm z = r \ summable (\n. f n * z^n)" by (intro exI[of _ "of_real r"]) simp qed lemma conv_radius_conv_Sup: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" shows "conv_radius f = Sup {r. \z. ereal (norm z) < r \ summable (\n. f n * z ^ n)}" proof (rule Sup_eqI [symmetric], goal_cases) case (1 r) thus ?case by (intro conv_radius_geI_ex') auto next case (2 r) from 2[of 0] have r: "r \ 0" by auto show ?case proof (intro conv_radius_leI_ex' r) fix R assume R: "R > 0" "ereal R > r" with r obtain r' where [simp]: "r = ereal r'" by (cases r) auto show "\summable (\n. f n * of_real R ^ n)" proof assume *: "summable (\n. f n * of_real R ^ n)" define R' where "R' = (R + r') / 2" from R have R': "R' > r'" "R' < R" by (simp_all add: R'_def) hence "\z. norm z < R' \ summable (\n. f n * z ^ n)" using powser_inside[OF *] by auto from 2[of R'] and this have "R' \ r'" by auto with \R' > r'\ show False by simp qed qed qed lemma conv_radius_shift: fixes f :: "nat \ 'a :: {banach, real_normed_div_algebra}" shows "conv_radius (\n. f (n + m)) = conv_radius f" unfolding conv_radius_conv_Sup summable_powser_ignore_initial_segment .. lemma conv_radius_norm [simp]: "conv_radius (\x. norm (f x)) = conv_radius f" by (simp add: conv_radius_def) lemma conv_radius_ratio_limit_ereal: fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes nz: "eventually (\n. f n \ 0) sequentially" assumes lim: "(\n. ereal (norm (f n) / norm (f (Suc n)))) \ c" shows "conv_radius f = c" proof (rule conv_radius_eqI') show "c \ 0" by (intro Lim_bounded2[OF lim]) simp_all next fix r assume r: "0 < r" "ereal r < c" let ?l = "liminf (\n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))" have "?l = liminf (\n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))" using r by (simp add: norm_mult norm_power field_split_simps) also from r have "\ = liminf (\n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)" by (intro Liminf_ereal_mult_right) simp_all also have "liminf (\n. ereal (norm (f n) / (norm (f (Suc n))))) = c" by (intro lim_imp_Liminf lim) simp finally have l: "?l = c * ereal (inverse r)" by simp from r have l': "c * ereal (inverse r) > 1" by (cases c) (simp_all add: field_simps) show "summable (\n. f n * of_real r^n)" by (rule summable_norm_cancel, rule ratio_test_convergence) (insert r nz l l', auto elim!: eventually_mono) next fix r assume r: "0 < r" "ereal r > c" let ?l = "limsup (\n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))" have "?l = limsup (\n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))" using r by (simp add: norm_mult norm_power field_split_simps) also from r have "\ = limsup (\n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)" by (intro Limsup_ereal_mult_right) simp_all also have "limsup (\n. ereal (norm (f n) / (norm (f (Suc n))))) = c" by (intro lim_imp_Limsup lim) simp finally have l: "?l = c * ereal (inverse r)" by simp from r have l': "c * ereal (inverse r) < 1" by (cases c) (simp_all add: field_simps) show "\summable (\n. norm (f n * of_real r^n))" by (rule ratio_test_divergence) (insert r nz l l', auto elim!: eventually_mono) qed lemma conv_radius_ratio_limit_ereal_nonzero: fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes nz: "c \ 0" assumes lim: "(\n. ereal (norm (f n) / norm (f (Suc n)))) \ c" shows "conv_radius f = c" proof (rule conv_radius_ratio_limit_ereal[OF _ lim], rule ccontr) assume "\eventually (\n. f n \ 0) sequentially" hence "frequently (\n. f n = 0) sequentially" by (simp add: frequently_def) hence "frequently (\n. ereal (norm (f n) / norm (f (Suc n))) = 0) sequentially" by (force elim!: frequently_elim1) hence "c = 0" by (intro limit_frequently_eq[OF _ _ lim]) auto with nz show False by contradiction qed lemma conv_radius_ratio_limit: fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes "c' = ereal c" assumes nz: "eventually (\n. f n \ 0) sequentially" assumes lim: "(\n. norm (f n) / norm (f (Suc n))) \ c" shows "conv_radius f = c'" using assms by (intro conv_radius_ratio_limit_ereal) simp_all lemma conv_radius_ratio_limit_nonzero: fixes f :: "nat \ 'a :: {banach,real_normed_div_algebra}" assumes "c' = ereal c" assumes nz: "c \ 0" assumes lim: "(\n. norm (f n) / norm (f (Suc n))) \ c" shows "conv_radius f = c'" using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all lemma conv_radius_cmult_left: assumes "c \ (0 :: 'a :: {banach, real_normed_div_algebra})" shows "conv_radius (\n. c * f n) = conv_radius f" proof - have "conv_radius (\n. c * f n) = inverse (limsup (\n. ereal (root n (norm (c * f n)))))" unfolding conv_radius_def .. also have "(\n. ereal (root n (norm (c * f n)))) = (\n. ereal (root n (norm c)) * ereal (root n (norm (f n))))" by (rule ext) (auto simp: norm_mult real_root_mult) also have "limsup \ = ereal 1 * limsup (\n. ereal (root n (norm (f n))))" using assms by (intro ereal_limsup_lim_mult tendsto_ereal LIMSEQ_root_const) auto also have "inverse \ = conv_radius f" by (simp add: conv_radius_def) finally show ?thesis . qed lemma conv_radius_cmult_right: assumes "c \ (0 :: 'a :: {banach, real_normed_div_algebra})" shows "conv_radius (\n. f n * c) = conv_radius f" proof - have "conv_radius (\n. f n * c) = conv_radius (\n. c * f n)" by (simp add: conv_radius_def norm_mult mult.commute) with conv_radius_cmult_left[OF assms, of f] show ?thesis by simp qed lemma conv_radius_mult_power: assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" shows "conv_radius (\n. c ^ n * f n) = conv_radius f / ereal (norm c)" proof - have "limsup (\n. ereal (root n (norm (c ^ n * f n)))) = limsup (\n. ereal (norm c) * ereal (root n (norm (f n))))" by (intro Limsup_eq eventually_mono [OF eventually_gt_at_top[of "0::nat"]]) (auto simp: norm_mult norm_power real_root_mult real_root_power) also have "\ = ereal (norm c) * limsup (\n. ereal (root n (norm (f n))))" using assms by (subst Limsup_ereal_mult_left[symmetric]) simp_all finally have A: "limsup (\n. ereal (root n (norm (c ^ n * f n)))) = ereal (norm c) * limsup (\n. ereal (root n (norm (f n))))" . show ?thesis using assms apply (cases "limsup (\n. ereal (root n (norm (f n)))) = 0") apply (simp add: A conv_radius_def) apply (unfold conv_radius_def A divide_ereal_def, simp add: mult.commute ereal_inverse_mult) done qed lemma conv_radius_mult_power_right: assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" shows "conv_radius (\n. f n * c ^ n) = conv_radius f / ereal (norm c)" using conv_radius_mult_power[OF assms, of f] unfolding conv_radius_def by (simp add: mult.commute norm_mult) lemma conv_radius_divide_power: assumes "c \ (0 :: 'a :: {real_normed_div_algebra,banach})" shows "conv_radius (\n. f n / c^n) = conv_radius f * ereal (norm c)" proof - from assms have "inverse c \ 0" by simp from conv_radius_mult_power_right[OF this, of f] show ?thesis by (simp add: divide_inverse divide_ereal_def assms norm_inverse power_inverse) qed lemma conv_radius_add_ge: "min (conv_radius f) (conv_radius g) \ conv_radius (\x. f x + g x :: 'a :: {banach,real_normed_div_algebra})" by (rule conv_radius_geI_ex') (auto simp: algebra_simps intro!: summable_add summable_in_conv_radius) lemma conv_radius_mult_ge: fixes f g :: "nat \ ('a :: {banach,real_normed_div_algebra})" shows "conv_radius (\x. \i\x. f i * g (x - i)) \ min (conv_radius f) (conv_radius g)" proof (rule conv_radius_geI_ex') fix r assume r: "r > 0" "ereal r < min (conv_radius f) (conv_radius g)" from r have "summable (\n. (\i\n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))" by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all thus "summable (\n. (\i\n. f i * g (n - i)) * of_real r ^ n)" by (simp add: algebra_simps of_real_def power_add [symmetric] scaleR_sum_right) qed lemma le_conv_radius_iff: fixes a :: "nat \ 'a::{real_normed_div_algebra,banach}" shows "r \ conv_radius a \ (\x. norm (x-\) < r \ summable (\i. a i * (x - \) ^ i))" apply (intro iffI allI impI summable_in_conv_radius conv_radius_geI_ex) apply (meson less_ereal.simps(1) not_le order_trans) apply (rule_tac x="of_real ra" in exI, simp) apply (metis abs_of_nonneg add_diff_cancel_left' less_eq_real_def norm_of_real) done end diff --git a/src/HOL/Combinatorics/Multiset_Permutations.thy b/src/HOL/Combinatorics/Multiset_Permutations.thy --- a/src/HOL/Combinatorics/Multiset_Permutations.thy +++ b/src/HOL/Combinatorics/Multiset_Permutations.thy @@ -1,513 +1,514 @@ (* Author: Manuel Eberl (TU München) Defines the set of permutations of a given multiset (or set), i.e. the set of all lists whose entries correspond to the multiset (resp. set). *) section \Permutations of a Multiset\ theory Multiset_Permutations imports Complex_Main Permutations begin (* TODO Move *) lemma mset_tl: "xs \ [] \ mset (tl xs) = mset xs - {#hd xs#}" by (cases xs) simp_all lemma mset_set_image_inj: assumes "inj_on f A" shows "mset_set (f ` A) = image_mset f (mset_set A)" proof (cases "finite A") case True from this and assms show ?thesis by (induction A) auto qed (insert assms, simp add: finite_image_iff) lemma multiset_remove_induct [case_names empty remove]: assumes "P {#}" "\A. A \ {#} \ (\x. x \# A \ P (A - {#x#})) \ P A" shows "P A" proof (induction A rule: full_multiset_induct) case (less A) hence IH: "P B" if "B \# A" for B using that by blast show ?case proof (cases "A = {#}") case True thus ?thesis by (simp add: assms) next case False hence "P (A - {#x#})" if "x \# A" for x using that by (intro IH) (simp add: mset_subset_diff_self) from False and this show "P A" by (rule assms) qed qed lemma map_list_bind: "map g (List.bind xs f) = List.bind xs (map g \ f)" by (simp add: List.bind_def map_concat) lemma mset_eq_mset_set_imp_distinct: "finite A \ mset_set A = mset xs \ distinct xs" proof (induction xs arbitrary: A) case (Cons x xs A) from Cons.prems(2) have "x \# mset_set A" by simp with Cons.prems(1) have [simp]: "x \ A" by simp from Cons.prems have "x \# mset_set (A - {x})" by simp also from Cons.prems have "mset_set (A - {x}) = mset_set A - {#x#}" by (subst mset_set_Diff) simp_all also have "mset_set A = mset (x#xs)" by (simp add: Cons.prems) also have "\ - {#x#} = mset xs" by simp finally have [simp]: "x \ set xs" by (simp add: in_multiset_in_set) from Cons.prems show ?case by (auto intro!: Cons.IH[of "A - {x}"] simp: mset_set_Diff) qed simp_all (* END TODO *) subsection \Permutations of a multiset\ definition permutations_of_multiset :: "'a multiset \ 'a list set" where "permutations_of_multiset A = {xs. mset xs = A}" lemma permutations_of_multisetI: "mset xs = A \ xs \ permutations_of_multiset A" by (simp add: permutations_of_multiset_def) lemma permutations_of_multisetD: "xs \ permutations_of_multiset A \ mset xs = A" by (simp add: permutations_of_multiset_def) lemma permutations_of_multiset_Cons_iff: "x # xs \ permutations_of_multiset A \ x \# A \ xs \ permutations_of_multiset (A - {#x#})" by (auto simp: permutations_of_multiset_def) lemma permutations_of_multiset_empty [simp]: "permutations_of_multiset {#} = {[]}" unfolding permutations_of_multiset_def by simp lemma permutations_of_multiset_nonempty: assumes nonempty: "A \ {#}" shows "permutations_of_multiset A = (\x\set_mset A. ((#) x) ` permutations_of_multiset (A - {#x#}))" (is "_ = ?rhs") proof safe fix xs assume "xs \ permutations_of_multiset A" hence mset_xs: "mset xs = A" by (simp add: permutations_of_multiset_def) hence "xs \ []" by (auto simp: nonempty) then obtain x xs' where xs: "xs = x # xs'" by (cases xs) simp_all with mset_xs have "x \ set_mset A" "xs' \ permutations_of_multiset (A - {#x#})" by (auto simp: permutations_of_multiset_def) with xs show "xs \ ?rhs" by auto qed (auto simp: permutations_of_multiset_def) lemma permutations_of_multiset_singleton [simp]: "permutations_of_multiset {#x#} = {[x]}" by (simp add: permutations_of_multiset_nonempty) lemma permutations_of_multiset_doubleton: "permutations_of_multiset {#x,y#} = {[x,y], [y,x]}" by (simp add: permutations_of_multiset_nonempty insert_commute) lemma rev_permutations_of_multiset [simp]: "rev ` permutations_of_multiset A = permutations_of_multiset A" proof have "rev ` rev ` permutations_of_multiset A \ rev ` permutations_of_multiset A" unfolding permutations_of_multiset_def by auto also have "rev ` rev ` permutations_of_multiset A = permutations_of_multiset A" by (simp add: image_image) finally show "permutations_of_multiset A \ rev ` permutations_of_multiset A" . next show "rev ` permutations_of_multiset A \ permutations_of_multiset A" unfolding permutations_of_multiset_def by auto qed lemma length_finite_permutations_of_multiset: "xs \ permutations_of_multiset A \ length xs = size A" by (auto simp: permutations_of_multiset_def) lemma permutations_of_multiset_lists: "permutations_of_multiset A \ lists (set_mset A)" by (auto simp: permutations_of_multiset_def) lemma finite_permutations_of_multiset [simp]: "finite (permutations_of_multiset A)" proof (rule finite_subset) show "permutations_of_multiset A \ {xs. set xs \ set_mset A \ length xs = size A}" by (auto simp: permutations_of_multiset_def) show "finite {xs. set xs \ set_mset A \ length xs = size A}" by (rule finite_lists_length_eq) simp_all qed lemma permutations_of_multiset_not_empty [simp]: "permutations_of_multiset A \ {}" proof - - from ex_mset[of A] guess xs .. + from ex_mset[of A] obtain xs where "mset xs = A" .. thus ?thesis by (auto simp: permutations_of_multiset_def) qed lemma permutations_of_multiset_image: "permutations_of_multiset (image_mset f A) = map f ` permutations_of_multiset A" proof safe fix xs assume A: "xs \ permutations_of_multiset (image_mset f A)" from ex_mset[of A] obtain ys where ys: "mset ys = A" .. with A have "mset xs = mset (map f ys)" by (simp add: permutations_of_multiset_def) - from mset_eq_permutation[OF this] guess \ . note \ = this + then obtain \ where \: "\ permutes {.. (map f ys) = xs" + by (rule mset_eq_permutation) with ys have "xs = map f (permute_list \ ys)" by (simp add: permute_list_map) moreover from \ ys have "permute_list \ ys \ permutations_of_multiset A" by (simp add: permutations_of_multiset_def) ultimately show "xs \ map f ` permutations_of_multiset A" by blast qed (auto simp: permutations_of_multiset_def) subsection \Cardinality of permutations\ text \ In this section, we prove some basic facts about the number of permutations of a multiset. \ context begin private lemma multiset_prod_fact_insert: "(\y\set_mset (A+{#x#}). fact (count (A+{#x#}) y)) = (count A x + 1) * (\y\set_mset A. fact (count A y))" proof - have "(\y\set_mset (A+{#x#}). fact (count (A+{#x#}) y)) = (\y\set_mset (A+{#x#}). (if y = x then count A x + 1 else 1) * fact (count A y))" by (intro prod.cong) simp_all also have "\ = (count A x + 1) * (\y\set_mset (A+{#x#}). fact (count A y))" by (simp add: prod.distrib) also have "(\y\set_mset (A+{#x#}). fact (count A y)) = (\y\set_mset A. fact (count A y))" by (intro prod.mono_neutral_right) (auto simp: not_in_iff) finally show ?thesis . qed private lemma multiset_prod_fact_remove: "x \# A \ (\y\set_mset A. fact (count A y)) = count A x * (\y\set_mset (A-{#x#}). fact (count (A-{#x#}) y))" using multiset_prod_fact_insert[of "A - {#x#}" x] by simp lemma card_permutations_of_multiset_aux: "card (permutations_of_multiset A) * (\x\set_mset A. fact (count A x)) = fact (size A)" proof (induction A rule: multiset_remove_induct) case (remove A) have "card (permutations_of_multiset A) = card (\x\set_mset A. (#) x ` permutations_of_multiset (A - {#x#}))" by (simp add: permutations_of_multiset_nonempty remove.hyps) also have "\ = (\x\set_mset A. card (permutations_of_multiset (A - {#x#})))" by (subst card_UN_disjoint) (auto simp: card_image) also have "\ * (\x\set_mset A. fact (count A x)) = (\x\set_mset A. card (permutations_of_multiset (A - {#x#})) * (\y\set_mset A. fact (count A y)))" by (subst sum_distrib_right) simp_all also have "\ = (\x\set_mset A. count A x * fact (size A - 1))" proof (intro sum.cong refl) fix x assume x: "x \# A" have "card (permutations_of_multiset (A - {#x#})) * (\y\set_mset A. fact (count A y)) = count A x * (card (permutations_of_multiset (A - {#x#})) * (\y\set_mset (A - {#x#}). fact (count (A - {#x#}) y)))" (is "?lhs = _") by (subst multiset_prod_fact_remove[OF x]) simp_all also note remove.IH[OF x] also from x have "size (A - {#x#}) = size A - 1" by (simp add: size_Diff_submset) finally show "?lhs = count A x * fact (size A - 1)" . qed also have "(\x\set_mset A. count A x * fact (size A - 1)) = size A * fact (size A - 1)" by (simp add: sum_distrib_right size_multiset_overloaded_eq) also from remove.hyps have "\ = fact (size A)" by (cases "size A") auto finally show ?case . qed simp_all theorem card_permutations_of_multiset: "card (permutations_of_multiset A) = fact (size A) div (\x\set_mset A. fact (count A x))" "(\x\set_mset A. fact (count A x) :: nat) dvd fact (size A)" by (simp_all flip: card_permutations_of_multiset_aux[of A]) lemma card_permutations_of_multiset_insert_aux: "card (permutations_of_multiset (A + {#x#})) * (count A x + 1) = (size A + 1) * card (permutations_of_multiset A)" proof - note card_permutations_of_multiset_aux[of "A + {#x#}"] also have "fact (size (A + {#x#})) = (size A + 1) * fact (size A)" by simp also note multiset_prod_fact_insert[of A x] also note card_permutations_of_multiset_aux[of A, symmetric] finally have "card (permutations_of_multiset (A + {#x#})) * (count A x + 1) * (\y\set_mset A. fact (count A y)) = (size A + 1) * card (permutations_of_multiset A) * (\x\set_mset A. fact (count A x))" by (simp only: mult_ac) thus ?thesis by (subst (asm) mult_right_cancel) simp_all qed lemma card_permutations_of_multiset_remove_aux: assumes "x \# A" shows "card (permutations_of_multiset A) * count A x = size A * card (permutations_of_multiset (A - {#x#}))" proof - from assms have A: "A - {#x#} + {#x#} = A" by simp from assms have B: "size A = size (A - {#x#}) + 1" by (subst A [symmetric], subst size_union) simp show ?thesis using card_permutations_of_multiset_insert_aux[of "A - {#x#}" x, unfolded A] assms by (simp add: B) qed lemma real_card_permutations_of_multiset_remove: assumes "x \# A" shows "real (card (permutations_of_multiset (A - {#x#}))) = real (card (permutations_of_multiset A) * count A x) / real (size A)" using assms by (subst card_permutations_of_multiset_remove_aux[OF assms]) auto lemma real_card_permutations_of_multiset_remove': assumes "x \# A" shows "real (card (permutations_of_multiset A)) = real (size A * card (permutations_of_multiset (A - {#x#}))) / real (count A x)" using assms by (subst card_permutations_of_multiset_remove_aux[OF assms, symmetric]) simp end subsection \Permutations of a set\ definition permutations_of_set :: "'a set \ 'a list set" where "permutations_of_set A = {xs. set xs = A \ distinct xs}" lemma permutations_of_set_altdef: "finite A \ permutations_of_set A = permutations_of_multiset (mset_set A)" by (auto simp add: permutations_of_set_def permutations_of_multiset_def mset_set_set in_multiset_in_set [symmetric] mset_eq_mset_set_imp_distinct) lemma permutations_of_setI [intro]: assumes "set xs = A" "distinct xs" shows "xs \ permutations_of_set A" using assms unfolding permutations_of_set_def by simp lemma permutations_of_setD: assumes "xs \ permutations_of_set A" shows "set xs = A" "distinct xs" using assms unfolding permutations_of_set_def by simp_all lemma permutations_of_set_lists: "permutations_of_set A \ lists A" unfolding permutations_of_set_def by auto lemma permutations_of_set_empty [simp]: "permutations_of_set {} = {[]}" by (auto simp: permutations_of_set_def) lemma UN_set_permutations_of_set [simp]: "finite A \ (\xs\permutations_of_set A. set xs) = A" using finite_distinct_list by (auto simp: permutations_of_set_def) lemma permutations_of_set_infinite: "\finite A \ permutations_of_set A = {}" by (auto simp: permutations_of_set_def) lemma permutations_of_set_nonempty: "A \ {} \ permutations_of_set A = (\x\A. (\xs. x # xs) ` permutations_of_set (A - {x}))" by (cases "finite A") (simp_all add: permutations_of_multiset_nonempty mset_set_empty_iff mset_set_Diff permutations_of_set_altdef permutations_of_set_infinite) lemma permutations_of_set_singleton [simp]: "permutations_of_set {x} = {[x]}" by (subst permutations_of_set_nonempty) auto lemma permutations_of_set_doubleton: "x \ y \ permutations_of_set {x,y} = {[x,y], [y,x]}" by (subst permutations_of_set_nonempty) (simp_all add: insert_Diff_if insert_commute) lemma rev_permutations_of_set [simp]: "rev ` permutations_of_set A = permutations_of_set A" by (cases "finite A") (simp_all add: permutations_of_set_altdef permutations_of_set_infinite) lemma length_finite_permutations_of_set: "xs \ permutations_of_set A \ length xs = card A" by (auto simp: permutations_of_set_def distinct_card) lemma finite_permutations_of_set [simp]: "finite (permutations_of_set A)" by (cases "finite A") (simp_all add: permutations_of_set_infinite permutations_of_set_altdef) lemma permutations_of_set_empty_iff [simp]: "permutations_of_set A = {} \ \finite A" unfolding permutations_of_set_def using finite_distinct_list[of A] by auto lemma card_permutations_of_set [simp]: "finite A \ card (permutations_of_set A) = fact (card A)" by (simp add: permutations_of_set_altdef card_permutations_of_multiset del: One_nat_def) lemma permutations_of_set_image_inj: assumes inj: "inj_on f A" shows "permutations_of_set (f ` A) = map f ` permutations_of_set A" by (cases "finite A") (simp_all add: permutations_of_set_infinite permutations_of_set_altdef permutations_of_multiset_image mset_set_image_inj inj finite_image_iff) lemma permutations_of_set_image_permutes: "\ permutes A \ map \ ` permutations_of_set A = permutations_of_set A" by (subst permutations_of_set_image_inj [symmetric]) (simp_all add: permutes_inj_on permutes_image) subsection \Code generation\ text \ First, we give code an implementation for permutations of lists. \ declare length_remove1 [termination_simp] fun permutations_of_list_impl where "permutations_of_list_impl xs = (if xs = [] then [[]] else List.bind (remdups xs) (\x. map ((#) x) (permutations_of_list_impl (remove1 x xs))))" fun permutations_of_list_impl_aux where "permutations_of_list_impl_aux acc xs = (if xs = [] then [acc] else List.bind (remdups xs) (\x. permutations_of_list_impl_aux (x#acc) (remove1 x xs)))" declare permutations_of_list_impl_aux.simps [simp del] declare permutations_of_list_impl.simps [simp del] lemma permutations_of_list_impl_Nil [simp]: "permutations_of_list_impl [] = [[]]" by (simp add: permutations_of_list_impl.simps) lemma permutations_of_list_impl_nonempty: "xs \ [] \ permutations_of_list_impl xs = List.bind (remdups xs) (\x. map ((#) x) (permutations_of_list_impl (remove1 x xs)))" by (subst permutations_of_list_impl.simps) simp_all lemma set_permutations_of_list_impl: "set (permutations_of_list_impl xs) = permutations_of_multiset (mset xs)" by (induction xs rule: permutations_of_list_impl.induct) (subst permutations_of_list_impl.simps, simp_all add: permutations_of_multiset_nonempty set_list_bind) lemma distinct_permutations_of_list_impl: "distinct (permutations_of_list_impl xs)" by (induction xs rule: permutations_of_list_impl.induct, subst permutations_of_list_impl.simps) (auto intro!: distinct_list_bind simp: distinct_map o_def disjoint_family_on_def) lemma permutations_of_list_impl_aux_correct': "permutations_of_list_impl_aux acc xs = map (\xs. rev xs @ acc) (permutations_of_list_impl xs)" by (induction acc xs rule: permutations_of_list_impl_aux.induct, subst permutations_of_list_impl_aux.simps, subst permutations_of_list_impl.simps) (auto simp: map_list_bind intro!: list_bind_cong) lemma permutations_of_list_impl_aux_correct: "permutations_of_list_impl_aux [] xs = map rev (permutations_of_list_impl xs)" by (simp add: permutations_of_list_impl_aux_correct') lemma distinct_permutations_of_list_impl_aux: "distinct (permutations_of_list_impl_aux acc xs)" by (simp add: permutations_of_list_impl_aux_correct' distinct_map distinct_permutations_of_list_impl inj_on_def) lemma set_permutations_of_list_impl_aux: "set (permutations_of_list_impl_aux [] xs) = permutations_of_multiset (mset xs)" by (simp add: permutations_of_list_impl_aux_correct set_permutations_of_list_impl) declare set_permutations_of_list_impl_aux [symmetric, code] value [code] "permutations_of_multiset {#1,2,3,4::int#}" text \ Now we turn to permutations of sets. We define an auxiliary version with an accumulator to avoid having to map over the results. \ function permutations_of_set_aux where "permutations_of_set_aux acc A = (if \finite A then {} else if A = {} then {acc} else (\x\A. permutations_of_set_aux (x#acc) (A - {x})))" by auto termination by (relation "Wellfounded.measure (card \ snd)") (simp_all add: card_gt_0_iff) lemma permutations_of_set_aux_altdef: "permutations_of_set_aux acc A = (\xs. rev xs @ acc) ` permutations_of_set A" proof (cases "finite A") assume "finite A" thus ?thesis proof (induction A arbitrary: acc rule: finite_psubset_induct) case (psubset A acc) show ?case proof (cases "A = {}") case False note [simp del] = permutations_of_set_aux.simps from psubset.hyps False have "permutations_of_set_aux acc A = (\y\A. permutations_of_set_aux (y#acc) (A - {y}))" by (subst permutations_of_set_aux.simps) simp_all also have "\ = (\y\A. (\xs. rev xs @ acc) ` (\xs. y # xs) ` permutations_of_set (A - {y}))" apply (rule arg_cong [of _ _ Union], rule image_cong) apply (simp_all add: image_image) apply (subst psubset) apply auto done also from False have "\ = (\xs. rev xs @ acc) ` permutations_of_set A" by (subst (2) permutations_of_set_nonempty) (simp_all add: image_UN) finally show ?thesis . qed simp_all qed qed (simp_all add: permutations_of_set_infinite) declare permutations_of_set_aux.simps [simp del] lemma permutations_of_set_aux_correct: "permutations_of_set_aux [] A = permutations_of_set A" by (simp add: permutations_of_set_aux_altdef) text \ In another refinement step, we define a version on lists. \ declare length_remove1 [termination_simp] fun permutations_of_set_aux_list where "permutations_of_set_aux_list acc xs = (if xs = [] then [acc] else List.bind xs (\x. permutations_of_set_aux_list (x#acc) (List.remove1 x xs)))" definition permutations_of_set_list where "permutations_of_set_list xs = permutations_of_set_aux_list [] xs" declare permutations_of_set_aux_list.simps [simp del] lemma permutations_of_set_aux_list_refine: assumes "distinct xs" shows "set (permutations_of_set_aux_list acc xs) = permutations_of_set_aux acc (set xs)" using assms by (induction acc xs rule: permutations_of_set_aux_list.induct) (subst permutations_of_set_aux_list.simps, subst permutations_of_set_aux.simps, simp_all add: set_list_bind) text \ The permutation lists contain no duplicates if the inputs contain no duplicates. Therefore, these functions can easily be used when working with a representation of sets by distinct lists. The same approach should generalise to any kind of set implementation that supports a monadic bind operation, and since the results are disjoint, merging should be cheap. \ lemma distinct_permutations_of_set_aux_list: "distinct xs \ distinct (permutations_of_set_aux_list acc xs)" by (induction acc xs rule: permutations_of_set_aux_list.induct) (subst permutations_of_set_aux_list.simps, auto intro!: distinct_list_bind simp: disjoint_family_on_def permutations_of_set_aux_list_refine permutations_of_set_aux_altdef) lemma distinct_permutations_of_set_list: "distinct xs \ distinct (permutations_of_set_list xs)" by (simp add: permutations_of_set_list_def distinct_permutations_of_set_aux_list) lemma permutations_of_list: "permutations_of_set (set xs) = set (permutations_of_set_list (remdups xs))" by (simp add: permutations_of_set_aux_correct [symmetric] permutations_of_set_aux_list_refine permutations_of_set_list_def) lemma permutations_of_list_code [code]: "permutations_of_set (set xs) = set (permutations_of_set_list (remdups xs))" "permutations_of_set (List.coset xs) = Code.abort (STR ''Permutation of set complement not supported'') (\_. permutations_of_set (List.coset xs))" by (simp_all add: permutations_of_list) value [code] "permutations_of_set (set ''abcd'')" end diff --git a/src/HOL/Complex_Analysis/Complex_Residues.thy b/src/HOL/Complex_Analysis/Complex_Residues.thy --- a/src/HOL/Complex_Analysis/Complex_Residues.thy +++ b/src/HOL/Complex_Analysis/Complex_Residues.thy @@ -1,540 +1,542 @@ theory Complex_Residues imports Complex_Singularities begin subsection \Definition of residues\ text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. Interactive Theorem Proving\ definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" lemma residue_cong: assumes eq: "eventually (\z. f z = g z) (at z)" and "z = z'" shows "residue f z = residue g z'" proof - from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) let ?P = "\f c e. (\\>0. \ < e \ (f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \))" have "residue f z = residue g z" unfolding residue_def proof (rule Eps_cong) fix c :: complex have "\e>0. ?P g c e" if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g proof - from that(1) obtain e where e: "e > 0" "?P f c e" by blast from that(2) obtain e' where e': "e' > 0" "\z'. z' \ z \ dist z' z < e' \ f z' = g z'" unfolding eventually_at by blast have "?P g c (min e e')" proof (intro allI exI impI, goal_cases) case (1 \) hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" using e(2) by auto thus ?case proof (rule has_contour_integral_eq) fix z' assume "z' \ path_image (circlepath z \)" hence "dist z' z < e'" and "z' \ z" using 1 by (auto simp: dist_commute) with e'(2)[of z'] show "f z' = g z'" by simp qed qed moreover from e and e' have "min e e' > 0" by auto ultimately show ?thesis by blast qed from this[OF _ eq] and this[OF _ eq'] show "(\e>0. ?P f c e) \ (\e>0. ?P g c e)" by blast qed with assms show ?thesis by simp qed lemma contour_integral_circlepath_eq: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" and e2_cball:"cball z e2 \ s" shows "f contour_integrable_on circlepath z e1" "f contour_integrable_on circlepath z e2" "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" proof - define l where "l \ linepath (z+e2) (z+e1)" have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto have "e2>0" using \e1>0\ \e1\e2\ by auto have zl_img:"z\path_image l" proof assume "z \ path_image l" then have "e2 \ cmod (e2 - e1)" using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def by (auto simp add:closed_segment_commute) thus False using \e2>0\ \e1>0\ \e1\e2\ apply (subst (asm) norm_of_real) by auto qed define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" proof - show "f contour_integrable_on circlepath z e2" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e2>0\ e2_cball by auto show "f contour_integrable_on (circlepath z e1)" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e1>0\ \e1\e2\ e2_cball by auto qed have [simp]:"f contour_integrable_on l" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def by auto then show "f contour_integrable_on l" unfolding l_def apply (intro contour_integrable_continuous_linepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) by auto qed let ?ig="\g. contour_integral g f" have "(f has_contour_integral 0) g" proof (rule Cauchy_theorem_global[OF _ f_holo]) show "open (s - {z})" using \open s\ by auto show "valid_path g" unfolding g_def l_def by auto show "pathfinish g = pathstart g" unfolding g_def l_def by auto next have path_img:"path_image g \ cball z e2" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto ultimately show ?thesis unfolding g_def l_def using \e2>0\ by (simp add: path_image_join closed_segment_commute) qed show "path_image g \ s - {z}" proof - have "z\path_image g" using zl_img unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) moreover note \cball z e2 \ s\ and path_img ultimately show ?thesis by auto qed show "winding_number g w = 0" when"w \ s - {z}" for w proof - have "winding_number g w = 0" when "w\s" using that e2_cball apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) by (auto simp add:g_def l_def) moreover have "winding_number g z=0" proof - let ?Wz="\g. winding_number g z" have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + ?Wz (reversepath l)" using \e2>0\ \e1>0\ zl_img unfolding g_def l_def by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" using zl_img apply (subst (2) winding_number_reversepath) by (auto simp add:l_def closed_segment_commute) also have "... = 0" proof - have "?Wz (circlepath z e2) = 1" using \e2>0\ by (auto intro: winding_number_circlepath_centre) moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \e1>0\ apply (subst winding_number_reversepath) by (auto intro: winding_number_circlepath_centre) ultimately show ?thesis by auto qed finally show ?thesis . qed ultimately show ?thesis using that by auto qed qed then have "0 = ?ig g" using contour_integral_unique by simp also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + ?ig (reversepath l)" unfolding g_def by (auto simp add:contour_integrable_reversepath_eq) also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" by (auto simp add:contour_integral_reversepath) finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" by simp qed lemma base_residue: assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" and r_cball:"cball z r \ s" shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" proof - obtain e where "e>0" and e_cball:"cball z e \ s" using open_contains_cball[of s] \open s\ \z\s\ by auto define c where "c \ 2 * pi * \" define i where "i \ contour_integral (circlepath z e) f / c" have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ proof - have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" "f contour_integrable_on circlepath z \" "f contour_integrable_on circlepath z e" using \\ by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ then show ?thesis unfolding i_def c_def by (auto intro:has_contour_integral_integral) qed then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" unfolding residue_def c_def apply (rule_tac someI[of _ i],intro exI[where x=e]) by (auto simp add:\e>0\ c_def) then obtain e' where "e'>0" and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" by auto let ?int="\e. contour_integral (circlepath z e) f" define \ where "\ \ Min {r,e'} / 2" have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto have "(f has_contour_integral c * (residue f z)) (circlepath z \)" using e'_def[rule_format,OF \\>0\ \\] . then show ?thesis unfolding c_def using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) qed lemma residue_holo: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s" shows "residue f z = 0" proof - define c where "c \ 2 * pi * \" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f has_contour_integral c*residue f z) (circlepath z e)" using f_holo by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) moreover have "(f has_contour_integral 0) (circlepath z e)" using f_holo e_cball \e>0\ by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) ultimately have "c*residue f z =0" using has_contour_integral_unique by blast thus ?thesis unfolding c_def by auto qed lemma residue_const:"residue (\_. c) z = 0" by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) lemma residue_add: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z + g z) z= residue f z + residue g z" proof - define c where "c \ 2 * pi * \" define fg where "fg \ (\z. f z+g z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(fg has_contour_integral c * residue fg z) (circlepath z e)" unfolding fg_def using f_holo g_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro:holomorphic_intros) moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" unfolding fg_def using f_holo g_holo by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) ultimately have "c*(residue f z + residue g z) = c * residue fg z" using has_contour_integral_unique by (auto simp add:distrib_left) thus ?thesis unfolding fg_def by (auto simp add:c_def) qed lemma residue_lmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. c * (f z)) z= c * residue f z" proof (cases "c=0") case True thus ?thesis using residue_const by auto next case False define c' where "c' \ 2 * pi * \" define f' where "f' \ (\z. c * (f z))" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) by (auto intro:holomorphic_intros) moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" unfolding f'_def using f_holo by (auto intro: has_contour_integral_lmul base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) ultimately have "c' * residue f' z = c * (c' * residue f z)" using has_contour_integral_unique by auto thus ?thesis unfolding f'_def c'_def using False by (auto simp add:field_simps) qed lemma residue_rmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) * c) z= residue f z * c" using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) lemma residue_div: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) / c) z= residue f z / c " using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) lemma residue_neg: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. - (f z)) z= - residue f z" using residue_lmul[OF assms,of "-1"] by auto lemma residue_diff: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z - g z) z= residue f z - residue g z" using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] by (auto intro:holomorphic_intros g_holo) lemma residue_simple: assumes "open s" "z\s" and f_holo:"f holomorphic_on s" shows "residue (\w. f w / (w - z)) z = f z" proof - define c where "c \ 2 * pi * \" define f' where "f' \ \w. f w / (w - z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c * f z) (circlepath z e)" unfolding f'_def c_def using \e>0\ f_holo e_cball by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro!:holomorphic_intros) ultimately have "c * f z = c * residue f' z" using has_contour_integral_unique by blast thus ?thesis unfolding c_def f'_def by auto qed lemma residue_simple': assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" and lim: "((\w. f w * (w - z)) \ c) (at z)" shows "residue f z = c" proof - define g where "g = (\w. if w = z then c else f w * (w - z))" from holo have "(\w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P") by (force intro: holomorphic_intros) also have "?P \ g holomorphic_on (s - {z})" by (intro holomorphic_cong refl) (simp_all add: g_def) finally have *: "g holomorphic_on (s - {z})" . note lim also have "(\w. f w * (w - z)) \z\ c \ g \z\ g z" by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter) finally have **: "g \z\ g z" . have g_holo: "g holomorphic_on s" by (rule no_isolated_singularity'[where K = "{z}"]) (insert assms * **, simp_all add: at_within_open_NO_MATCH) from s and this have "residue (\w. g w / (w - z)) z = g z" by (rule residue_simple) also have "\\<^sub>F za in at z. g za / (za - z) = f za" unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def) hence "residue (\w. g w / (w - z)) z = residue f z" by (intro residue_cong refl) finally show ?thesis by (simp add: g_def) qed lemma residue_holomorphic_over_power: assumes "open A" "z0 \ A" "f holomorphic_on A" shows "residue (\z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n" proof - let ?f = "\z. f z / (z - z0) ^ Suc n" from assms(1,2) obtain r where r: "r > 0" "cball z0 r \ A" by (auto simp: open_contains_cball) have "(?f has_contour_integral 2 * pi * \ * residue ?f z0) (circlepath z0 r)" using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros) moreover have "(?f has_contour_integral 2 * pi * \ / fact n * (deriv ^^ n) f z0) (circlepath z0 r)" using assms r by (intro Cauchy_has_contour_integral_higher_derivative_circlepath) (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on) ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" by (rule has_contour_integral_unique) thus ?thesis by (simp add: field_simps) qed lemma residue_holomorphic_over_power': assumes "open A" "0 \ A" "f holomorphic_on A" shows "residue (\z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" using residue_holomorphic_over_power[OF assms] by simp theorem residue_fps_expansion_over_power_at_0: assumes "f has_fps_expansion F" shows "residue (\z. f z / z ^ Suc n) 0 = fps_nth F n" proof - - from has_fps_expansion_imp_holomorphic[OF assms] guess s . note s = this - have "residue (\z. f z / (z - 0) ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" - using assms s unfolding has_fps_expansion_def + from has_fps_expansion_imp_holomorphic[OF assms] obtain s + where "open s" "0 \ s" "f holomorphic_on s" "\z. z \ s \ f z = eval_fps F z" + by auto + with assms have "residue (\z. f z / (z - 0) ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" + unfolding has_fps_expansion_def by (intro residue_holomorphic_over_power[of s]) (auto simp: zero_ereal_def) also from assms have "\ = fps_nth F n" by (subst fps_nth_fps_expansion) auto finally show ?thesis by simp qed lemma residue_pole_order: fixes f::"complex \ complex" and z::complex defines "n \ nat (- zorder f z)" and "h \ zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and pole:"is_pole f z" shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ ball z e" and h_holo: "h holomorphic_on cball z r" and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" proof - obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" "(\w\cball z r - {z}. f w = h w / (w - z) ^ n \ h w \ 0)" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\,folded n_def h_def] by auto have "n>0" using \zorder f z < 0\ unfolding n_def by simp moreover have "(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" using \h z\0\ r(6) by blast ultimately show ?thesis using r(3,4,5) that by blast qed have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" using h_divide by simp define c where "c \ 2 * pi * \" define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" define h' where "h' \ \u. h u / (u - z) ^ n" have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" unfolding h'_def proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", folded c_def Suc_pred'[OF \n>0\]]) show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp show "h holomorphic_on ball z r" using h_holo by auto show " z \ ball z r" using \r>0\ by auto qed then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto then have "(f has_contour_integral c * der_f) (circlepath z r)" proof (elim has_contour_integral_eq) fix x assume "x \ path_image (circlepath z r)" hence "x\cball z r - {z}" using \r>0\ by auto then show "h' x = f x" using h_divide unfolding h'_def by auto qed moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] unfolding c_def by simp ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast hence "der_f = residue f z" unfolding c_def by auto thus ?thesis unfolding der_f_def by auto qed lemma residue_simple_pole: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" shows "residue f z0 = zor_poly f z0 z0" using assms by (subst residue_pole_order) simp_all lemma residue_simple_pole_limit: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" assumes "((\x. f (g x) * (g x - z0)) \ c) F" assumes "filterlim g (at z0) F" "F \ bot" shows "residue f z0 = c" proof - have "residue f z0 = zor_poly f z0 z0" by (rule residue_simple_pole assms)+ also have "\ = c" apply (rule zor_poly_pole_eqI) using assms by auto finally show ?thesis . qed lemma assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and "open s" "connected s" "z \ s" assumes g_deriv:"(g has_field_derivative g') (at z)" assumes "f z \ 0" "g z = 0" "g' \ 0" shows porder_simple_pole_deriv: "zorder (\w. f w / g w) z = - 1" and residue_simple_pole_deriv: "residue (\w. f w / g w) z = f z / g'" proof - have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z" using isolated_singularity_at_holomorphic[OF _ \open s\ \z\s\] f_holo g_holo by (meson Diff_subset holomorphic_on_subset)+ have [simp]:"not_essential f z" "not_essential g z" unfolding not_essential_def using f_holo g_holo assms(3,5) by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+ have g_nconst:"\\<^sub>F w in at z. g w \0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at z. g w \ 0)" then have "\\<^sub>F w in nhds z. g w = 0" unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball) then have "deriv g z = deriv (\_. 0) z" by (intro deriv_cong_ev) auto then have "deriv g z = 0" by auto then have "g' = 0" using g_deriv DERIV_imp_deriv by blast then show False using \g'\0\ by auto qed have "zorder (\w. f w / g w) z = zorder f z - zorder g z" proof - have "\\<^sub>F w in at z. f w \0 \ w\s" apply (rule non_zero_neighbour_alt) using assms by auto with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" by (elim frequently_rev_mp eventually_rev_mp,auto) then show ?thesis using zorder_divide[of f z g] by auto qed moreover have "zorder f z=0" apply (rule zorder_zero_eqI[OF f_holo \open s\ \z\s\]) using \f z\0\ by auto moreover have "zorder g z=1" apply (rule zorder_zero_eqI[OF g_holo \open s\ \z\s\]) subgoal using assms(8) by auto subgoal using DERIV_imp_deriv assms(9) g_deriv by auto subgoal by simp done ultimately show "zorder (\w. f w / g w) z = - 1" by auto show "residue (\w. f w / g w) z = f z / g'" proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified]) show "zorder (\w. f w / g w) z = - 1" by fact show "isolated_singularity_at (\w. f w / g w) z" by (auto intro: singularity_intros) show "is_pole (\w. f w / g w) z" proof (rule is_pole_divide) have "\\<^sub>F x in at z. g x \ 0" apply (rule non_zero_neighbour) using g_nconst by auto moreover have "g \z\ 0" using DERIV_isCont assms(8) continuous_at g_deriv by force ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp show "isCont f z" using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on by auto show "f z \ 0" by fact qed show "filterlim id (at z) (at z)" by (simp add: filterlim_iff) have "((\w. (f w * (w - z)) / g w) \ f z / g') (at z)" proof (rule lhopital_complex_simple) show "((\w. f w * (w - z)) has_field_derivative f z) (at z)" using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo]) show "(g has_field_derivative g') (at z)" by fact qed (insert assms, auto) then show "((\w. (f w / g w) * (w - z)) \ f z / g') (at z)" by (simp add: field_split_simps) qed qed subsection \Poles and residues of some well-known functions\ (* TODO: add more material here for other functions *) lemma is_pole_Gamma: "is_pole Gamma (-of_nat n)" unfolding is_pole_def using Gamma_poles . lemma Gamma_residue: "residue Gamma (-of_nat n) = (-1) ^ n / fact n" proof (rule residue_simple') show "open (- (\\<^sub>\\<^sub>0 - {-of_nat n}) :: complex set)" by (intro open_Compl closed_subset_Ints) auto show "Gamma holomorphic_on (- (\\<^sub>\\<^sub>0 - {-of_nat n}) - {- of_nat n})" by (rule holomorphic_Gamma) auto show "(\w. Gamma w * (w - (-of_nat n))) \(-of_nat n)\ (- 1) ^ n / fact n" using Gamma_residues[of n] by simp qed auto end diff --git a/src/HOL/Computational_Algebra/Factorial_Ring.thy b/src/HOL/Computational_Algebra/Factorial_Ring.thy --- a/src/HOL/Computational_Algebra/Factorial_Ring.thy +++ b/src/HOL/Computational_Algebra/Factorial_Ring.thy @@ -1,2238 +1,2260 @@ (* Title: HOL/Computational_Algebra/Factorial_Ring.thy Author: Manuel Eberl, TU Muenchen Author: Florian Haftmann, TU Muenchen *) section \Factorial (semi)rings\ theory Factorial_Ring imports Main "HOL-Library.Multiset" begin subsection \Irreducible and prime elements\ context comm_semiring_1 begin definition irreducible :: "'a \ bool" where "irreducible p \ p \ 0 \ \p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1)" lemma not_irreducible_zero [simp]: "\irreducible 0" by (simp add: irreducible_def) lemma irreducible_not_unit: "irreducible p \ \p dvd 1" by (simp add: irreducible_def) lemma not_irreducible_one [simp]: "\irreducible 1" by (simp add: irreducible_def) lemma irreducibleI: "p \ 0 \ \p dvd 1 \ (\a b. p = a * b \ a dvd 1 \ b dvd 1) \ irreducible p" by (simp add: irreducible_def) lemma irreducibleD: "irreducible p \ p = a * b \ a dvd 1 \ b dvd 1" by (simp add: irreducible_def) lemma irreducible_mono: assumes irr: "irreducible b" and "a dvd b" "\a dvd 1" shows "irreducible a" proof (rule irreducibleI) fix c d assume "a = c * d" from assms obtain k where [simp]: "b = a * k" by auto from \a = c * d\ have "b = c * d * k" by simp hence "c dvd 1 \ (d * k) dvd 1" using irreducibleD[OF irr, of c "d * k"] by (auto simp: mult.assoc) thus "c dvd 1 \ d dvd 1" by auto qed (use assms in \auto simp: irreducible_def\) definition prime_elem :: "'a \ bool" where "prime_elem p \ p \ 0 \ \p dvd 1 \ (\a b. p dvd (a * b) \ p dvd a \ p dvd b)" lemma not_prime_elem_zero [simp]: "\prime_elem 0" by (simp add: prime_elem_def) lemma prime_elem_not_unit: "prime_elem p \ \p dvd 1" by (simp add: prime_elem_def) lemma prime_elemI: "p \ 0 \ \p dvd 1 \ (\a b. p dvd (a * b) \ p dvd a \ p dvd b) \ prime_elem p" by (simp add: prime_elem_def) lemma prime_elem_dvd_multD: "prime_elem p \ p dvd (a * b) \ p dvd a \ p dvd b" by (simp add: prime_elem_def) lemma prime_elem_dvd_mult_iff: "prime_elem p \ p dvd (a * b) \ p dvd a \ p dvd b" by (auto simp: prime_elem_def) lemma not_prime_elem_one [simp]: "\ prime_elem 1" by (auto dest: prime_elem_not_unit) lemma prime_elem_not_zeroI: assumes "prime_elem p" shows "p \ 0" using assms by (auto intro: ccontr) lemma prime_elem_dvd_power: "prime_elem p \ p dvd x ^ n \ p dvd x" by (induction n) (auto dest: prime_elem_dvd_multD intro: dvd_trans[of _ 1]) lemma prime_elem_dvd_power_iff: "prime_elem p \ n > 0 \ p dvd x ^ n \ p dvd x" by (auto dest: prime_elem_dvd_power intro: dvd_trans) lemma prime_elem_imp_nonzero [simp]: "ASSUMPTION (prime_elem x) \ x \ 0" unfolding ASSUMPTION_def by (rule prime_elem_not_zeroI) lemma prime_elem_imp_not_one [simp]: "ASSUMPTION (prime_elem x) \ x \ 1" unfolding ASSUMPTION_def by auto end lemma (in normalization_semidom) irreducible_cong: assumes "normalize a = normalize b" shows "irreducible a \ irreducible b" proof (cases "a = 0 \ a dvd 1") case True hence "\irreducible a" by (auto simp: irreducible_def) from True have "normalize a = 0 \ normalize a dvd 1" by auto also note assms finally have "b = 0 \ b dvd 1" by simp hence "\irreducible b" by (auto simp: irreducible_def) with \\irreducible a\ show ?thesis by simp next case False hence b: "b \ 0" "\is_unit b" using assms by (auto simp: is_unit_normalize[of b]) show ?thesis proof assume "irreducible a" thus "irreducible b" by (rule irreducible_mono) (use assms False b in \auto dest: associatedD2\) next assume "irreducible b" thus "irreducible a" by (rule irreducible_mono) (use assms False b in \auto dest: associatedD1\) qed qed lemma (in normalization_semidom) associatedE1: assumes "normalize a = normalize b" obtains u where "is_unit u" "a = u * b" proof (cases "a = 0") case [simp]: False from assms have [simp]: "b \ 0" by auto show ?thesis proof (rule that) show "is_unit (unit_factor a div unit_factor b)" by auto have "unit_factor a div unit_factor b * b = unit_factor a * (b div unit_factor b)" using \b \ 0\ unit_div_commute unit_div_mult_swap unit_factor_is_unit by metis also have "b div unit_factor b = normalize b" by simp finally show "a = unit_factor a div unit_factor b * b" by (metis assms unit_factor_mult_normalize) qed next case [simp]: True hence [simp]: "b = 0" using assms[symmetric] by auto show ?thesis by (intro that[of 1]) auto qed lemma (in normalization_semidom) associatedE2: assumes "normalize a = normalize b" obtains u where "is_unit u" "b = u * a" proof - from assms have "normalize b = normalize a" by simp then obtain u where "is_unit u" "b = u * a" by (elim associatedE1) thus ?thesis using that by blast qed (* TODO Move *) lemma (in normalization_semidom) normalize_power_normalize: "normalize (normalize x ^ n) = normalize (x ^ n)" proof (induction n) case (Suc n) have "normalize (normalize x ^ Suc n) = normalize (x * normalize (normalize x ^ n))" by simp also note Suc.IH finally show ?case by simp qed auto context algebraic_semidom begin lemma prime_elem_imp_irreducible: assumes "prime_elem p" shows "irreducible p" proof (rule irreducibleI) fix a b assume p_eq: "p = a * b" with assms have nz: "a \ 0" "b \ 0" by auto from p_eq have "p dvd a * b" by simp with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \p = a * b\ have "a * b dvd 1 * b \ a * b dvd a * 1" by auto thus "a dvd 1 \ b dvd 1" by (simp only: dvd_times_left_cancel_iff[OF nz(1)] dvd_times_right_cancel_iff[OF nz(2)]) qed (insert assms, simp_all add: prime_elem_def) lemma (in algebraic_semidom) unit_imp_no_irreducible_divisors: assumes "is_unit x" "irreducible p" shows "\p dvd x" proof (rule notI) assume "p dvd x" with \is_unit x\ have "is_unit p" by (auto intro: dvd_trans) with \irreducible p\ show False by (simp add: irreducible_not_unit) qed lemma unit_imp_no_prime_divisors: assumes "is_unit x" "prime_elem p" shows "\p dvd x" using unit_imp_no_irreducible_divisors[OF assms(1) prime_elem_imp_irreducible[OF assms(2)]] . lemma prime_elem_mono: assumes "prime_elem p" "\q dvd 1" "q dvd p" shows "prime_elem q" proof - from \q dvd p\ obtain r where r: "p = q * r" by (elim dvdE) hence "p dvd q * r" by simp with \prime_elem p\ have "p dvd q \ p dvd r" by (rule prime_elem_dvd_multD) hence "p dvd q" proof assume "p dvd r" then obtain s where s: "r = p * s" by (elim dvdE) from r have "p * 1 = p * (q * s)" by (subst (asm) s) (simp add: mult_ac) with \prime_elem p\ have "q dvd 1" by (subst (asm) mult_cancel_left) auto with \\q dvd 1\ show ?thesis by contradiction qed show ?thesis proof (rule prime_elemI) fix a b assume "q dvd (a * b)" with \p dvd q\ have "p dvd (a * b)" by (rule dvd_trans) with \prime_elem p\ have "p dvd a \ p dvd b" by (rule prime_elem_dvd_multD) with \q dvd p\ show "q dvd a \ q dvd b" by (blast intro: dvd_trans) qed (insert assms, auto) qed lemma irreducibleD': assumes "irreducible a" "b dvd a" shows "a dvd b \ is_unit b" proof - from assms obtain c where c: "a = b * c" by (elim dvdE) from irreducibleD[OF assms(1) this] have "is_unit b \ is_unit c" . thus ?thesis by (auto simp: c mult_unit_dvd_iff) qed lemma irreducibleI': assumes "a \ 0" "\is_unit a" "\b. b dvd a \ a dvd b \ is_unit b" shows "irreducible a" proof (rule irreducibleI) fix b c assume a_eq: "a = b * c" hence "a dvd b \ is_unit b" by (intro assms) simp_all thus "is_unit b \ is_unit c" proof assume "a dvd b" hence "b * c dvd b * 1" by (simp add: a_eq) moreover from \a \ 0\ a_eq have "b \ 0" by auto ultimately show ?thesis by (subst (asm) dvd_times_left_cancel_iff) auto qed blast qed (simp_all add: assms(1,2)) lemma irreducible_altdef: "irreducible x \ x \ 0 \ \is_unit x \ (\b. b dvd x \ x dvd b \ is_unit b)" using irreducibleI'[of x] irreducibleD'[of x] irreducible_not_unit[of x] by auto lemma prime_elem_multD: assumes "prime_elem (a * b)" shows "is_unit a \ is_unit b" proof - from assms have "a \ 0" "b \ 0" by (auto dest!: prime_elem_not_zeroI) moreover from assms prime_elem_dvd_multD [of "a * b"] have "a * b dvd a \ a * b dvd b" by auto ultimately show ?thesis using dvd_times_left_cancel_iff [of a b 1] dvd_times_right_cancel_iff [of b a 1] by auto qed lemma prime_elemD2: assumes "prime_elem p" and "a dvd p" and "\ is_unit a" shows "p dvd a" proof - from \a dvd p\ obtain b where "p = a * b" .. with \prime_elem p\ prime_elem_multD \\ is_unit a\ have "is_unit b" by auto with \p = a * b\ show ?thesis by (auto simp add: mult_unit_dvd_iff) qed lemma prime_elem_dvd_prod_msetE: assumes "prime_elem p" assumes dvd: "p dvd prod_mset A" obtains a where "a \# A" and "p dvd a" proof - from dvd have "\a. a \# A \ p dvd a" proof (induct A) case empty then show ?case using \prime_elem p\ by (simp add: prime_elem_not_unit) next case (add a A) then have "p dvd a * prod_mset A" by simp with \prime_elem p\ consider (A) "p dvd prod_mset A" | (B) "p dvd a" by (blast dest: prime_elem_dvd_multD) then show ?case proof cases case B then show ?thesis by auto next case A with add.hyps obtain b where "b \# A" "p dvd b" by auto then show ?thesis by auto qed qed with that show thesis by blast qed context begin private lemma prime_elem_powerD: assumes "prime_elem (p ^ n)" shows "prime_elem p \ n = 1" proof (cases n) case (Suc m) note assms also from Suc have "p ^ n = p * p^m" by simp finally have "is_unit p \ is_unit (p^m)" by (rule prime_elem_multD) moreover from assms have "\is_unit p" by (simp add: prime_elem_def is_unit_power_iff) ultimately have "is_unit (p ^ m)" by simp with \\is_unit p\ have "m = 0" by (simp add: is_unit_power_iff) with Suc assms show ?thesis by simp qed (insert assms, simp_all) lemma prime_elem_power_iff: "prime_elem (p ^ n) \ prime_elem p \ n = 1" by (auto dest: prime_elem_powerD) end lemma irreducible_mult_unit_left: "is_unit a \ irreducible (a * p) \ irreducible p" by (auto simp: irreducible_altdef mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff dvd_mult_unit_iff) lemma prime_elem_mult_unit_left: "is_unit a \ prime_elem (a * p) \ prime_elem p" by (auto simp: prime_elem_def mult.commute[of a] is_unit_mult_iff mult_unit_dvd_iff) lemma prime_elem_dvd_cases: assumes pk: "p*k dvd m*n" and p: "prime_elem p" shows "(\x. k dvd x*n \ m = p*x) \ (\y. k dvd m*y \ n = p*y)" proof - have "p dvd m*n" using dvd_mult_left pk by blast then consider "p dvd m" | "p dvd n" using p prime_elem_dvd_mult_iff by blast then show ?thesis proof cases case 1 then obtain a where "m = p * a" by (metis dvd_mult_div_cancel) then have "\x. k dvd x * n \ m = p * x" using p pk by (auto simp: mult.assoc) then show ?thesis .. next case 2 then obtain b where "n = p * b" by (metis dvd_mult_div_cancel) with p pk have "\y. k dvd m*y \ n = p*y" by (metis dvd_mult_right dvd_times_left_cancel_iff mult.left_commute mult_zero_left) then show ?thesis .. qed qed lemma prime_elem_power_dvd_prod: assumes pc: "p^c dvd m*n" and p: "prime_elem p" shows "\a b. a+b = c \ p^a dvd m \ p^b dvd n" using pc proof (induct c arbitrary: m n) case 0 show ?case by simp next case (Suc c) consider x where "p^c dvd x*n" "m = p*x" | y where "p^c dvd m*y" "n = p*y" using prime_elem_dvd_cases [of _ "p^c", OF _ p] Suc.prems by force then show ?case proof cases case (1 x) with Suc.hyps[of x n] obtain a b where "a + b = c \ p ^ a dvd x \ p ^ b dvd n" by blast with 1 have "Suc a + b = Suc c \ p ^ Suc a dvd m \ p ^ b dvd n" by (auto intro: mult_dvd_mono) thus ?thesis by blast next case (2 y) with Suc.hyps[of m y] obtain a b where "a + b = c \ p ^ a dvd m \ p ^ b dvd y" by blast with 2 have "a + Suc b = Suc c \ p ^ a dvd m \ p ^ Suc b dvd n" by (auto intro: mult_dvd_mono) with Suc.hyps [of m y] show "\a b. a + b = Suc c \ p ^ a dvd m \ p ^ b dvd n" by blast qed qed lemma prime_elem_power_dvd_cases: assumes "p ^ c dvd m * n" and "a + b = Suc c" and "prime_elem p" shows "p ^ a dvd m \ p ^ b dvd n" proof - from assms obtain r s where "r + s = c \ p ^ r dvd m \ p ^ s dvd n" by (blast dest: prime_elem_power_dvd_prod) moreover with assms have "a \ r \ b \ s" by arith ultimately show ?thesis by (auto intro: power_le_dvd) qed lemma prime_elem_not_unit' [simp]: "ASSUMPTION (prime_elem x) \ \is_unit x" unfolding ASSUMPTION_def by (rule prime_elem_not_unit) lemma prime_elem_dvd_power_iff: assumes "prime_elem p" shows "p dvd a ^ n \ p dvd a \ n > 0" using assms by (induct n) (auto dest: prime_elem_not_unit prime_elem_dvd_multD) lemma prime_power_dvd_multD: assumes "prime_elem p" assumes "p ^ n dvd a * b" and "n > 0" and "\ p dvd a" shows "p ^ n dvd b" using \p ^ n dvd a * b\ and \n > 0\ proof (induct n arbitrary: b) case 0 then show ?case by simp next case (Suc n) show ?case proof (cases "n = 0") case True with Suc \prime_elem p\ \\ p dvd a\ show ?thesis by (simp add: prime_elem_dvd_mult_iff) next case False then have "n > 0" by simp from \prime_elem p\ have "p \ 0" by auto from Suc.prems have *: "p * p ^ n dvd a * b" by simp then have "p dvd a * b" by (rule dvd_mult_left) with Suc \prime_elem p\ \\ p dvd a\ have "p dvd b" by (simp add: prime_elem_dvd_mult_iff) moreover define c where "c = b div p" ultimately have b: "b = p * c" by simp with * have "p * p ^ n dvd p * (a * c)" by (simp add: ac_simps) with \p \ 0\ have "p ^ n dvd a * c" by simp with Suc.hyps \n > 0\ have "p ^ n dvd c" by blast with \p \ 0\ show ?thesis by (simp add: b) qed qed end subsection \Generalized primes: normalized prime elements\ context normalization_semidom begin lemma irreducible_normalized_divisors: assumes "irreducible x" "y dvd x" "normalize y = y" shows "y = 1 \ y = normalize x" proof - from assms have "is_unit y \ x dvd y" by (auto simp: irreducible_altdef) thus ?thesis proof (elim disjE) assume "is_unit y" hence "normalize y = 1" by (simp add: is_unit_normalize) with assms show ?thesis by simp next assume "x dvd y" with \y dvd x\ have "normalize y = normalize x" by (rule associatedI) with assms show ?thesis by simp qed qed lemma irreducible_normalize_iff [simp]: "irreducible (normalize x) = irreducible x" using irreducible_mult_unit_left[of "1 div unit_factor x" x] by (cases "x = 0") (simp_all add: unit_div_commute) lemma prime_elem_normalize_iff [simp]: "prime_elem (normalize x) = prime_elem x" using prime_elem_mult_unit_left[of "1 div unit_factor x" x] by (cases "x = 0") (simp_all add: unit_div_commute) lemma prime_elem_associated: assumes "prime_elem p" and "prime_elem q" and "q dvd p" shows "normalize q = normalize p" using \q dvd p\ proof (rule associatedI) from \prime_elem q\ have "\ is_unit q" by (auto simp add: prime_elem_not_unit) with \prime_elem p\ \q dvd p\ show "p dvd q" by (blast intro: prime_elemD2) qed definition prime :: "'a \ bool" where "prime p \ prime_elem p \ normalize p = p" lemma not_prime_0 [simp]: "\prime 0" by (simp add: prime_def) lemma not_prime_unit: "is_unit x \ \prime x" using prime_elem_not_unit[of x] by (auto simp add: prime_def) lemma not_prime_1 [simp]: "\prime 1" by (simp add: not_prime_unit) lemma primeI: "prime_elem x \ normalize x = x \ prime x" by (simp add: prime_def) lemma prime_imp_prime_elem [dest]: "prime p \ prime_elem p" by (simp add: prime_def) lemma normalize_prime: "prime p \ normalize p = p" by (simp add: prime_def) lemma prime_normalize_iff [simp]: "prime (normalize p) \ prime_elem p" by (auto simp add: prime_def) lemma prime_power_iff: "prime (p ^ n) \ prime p \ n = 1" by (auto simp: prime_def prime_elem_power_iff) lemma prime_imp_nonzero [simp]: "ASSUMPTION (prime x) \ x \ 0" unfolding ASSUMPTION_def prime_def by auto lemma prime_imp_not_one [simp]: "ASSUMPTION (prime x) \ x \ 1" unfolding ASSUMPTION_def by auto lemma prime_not_unit' [simp]: "ASSUMPTION (prime x) \ \is_unit x" unfolding ASSUMPTION_def prime_def by auto lemma prime_normalize' [simp]: "ASSUMPTION (prime x) \ normalize x = x" unfolding ASSUMPTION_def prime_def by simp lemma unit_factor_prime: "prime x \ unit_factor x = 1" using unit_factor_normalize[of x] unfolding prime_def by auto lemma unit_factor_prime' [simp]: "ASSUMPTION (prime x) \ unit_factor x = 1" unfolding ASSUMPTION_def by (rule unit_factor_prime) lemma prime_imp_prime_elem' [simp]: "ASSUMPTION (prime x) \ prime_elem x" by (simp add: prime_def ASSUMPTION_def) lemma prime_dvd_multD: "prime p \ p dvd a * b \ p dvd a \ p dvd b" by (intro prime_elem_dvd_multD) simp_all lemma prime_dvd_mult_iff: "prime p \ p dvd a * b \ p dvd a \ p dvd b" by (auto dest: prime_dvd_multD) lemma prime_dvd_power: "prime p \ p dvd x ^ n \ p dvd x" by (auto dest!: prime_elem_dvd_power simp: prime_def) lemma prime_dvd_power_iff: "prime p \ n > 0 \ p dvd x ^ n \ p dvd x" by (subst prime_elem_dvd_power_iff) simp_all lemma prime_dvd_prod_mset_iff: "prime p \ p dvd prod_mset A \ (\x. x \# A \ p dvd x)" by (induction A) (simp_all add: prime_elem_dvd_mult_iff prime_imp_prime_elem, blast+) lemma prime_dvd_prod_iff: "finite A \ prime p \ p dvd prod f A \ (\x\A. p dvd f x)" by (auto simp: prime_dvd_prod_mset_iff prod_unfold_prod_mset) lemma primes_dvd_imp_eq: assumes "prime p" "prime q" "p dvd q" shows "p = q" proof - from assms have "irreducible q" by (simp add: prime_elem_imp_irreducible prime_def) from irreducibleD'[OF this \p dvd q\] assms have "q dvd p" by simp with \p dvd q\ have "normalize p = normalize q" by (rule associatedI) with assms show "p = q" by simp qed lemma prime_dvd_prod_mset_primes_iff: assumes "prime p" "\q. q \# A \ prime q" shows "p dvd prod_mset A \ p \# A" proof - from assms(1) have "p dvd prod_mset A \ (\x. x \# A \ p dvd x)" by (rule prime_dvd_prod_mset_iff) also from assms have "\ \ p \# A" by (auto dest: primes_dvd_imp_eq) finally show ?thesis . qed lemma prod_mset_primes_dvd_imp_subset: assumes "prod_mset A dvd prod_mset B" "\p. p \# A \ prime p" "\p. p \# B \ prime p" shows "A \# B" using assms proof (induction A arbitrary: B) case empty thus ?case by simp next case (add p A B) hence p: "prime p" by simp define B' where "B' = B - {#p#}" from add.prems have "p dvd prod_mset B" by (simp add: dvd_mult_left) with add.prems have "p \# B" by (subst (asm) (2) prime_dvd_prod_mset_primes_iff) simp_all hence B: "B = B' + {#p#}" by (simp add: B'_def) from add.prems p have "A \# B'" by (intro add.IH) (simp_all add: B) thus ?case by (simp add: B) qed lemma prod_mset_dvd_prod_mset_primes_iff: assumes "\x. x \# A \ prime x" "\x. x \# B \ prime x" shows "prod_mset A dvd prod_mset B \ A \# B" using assms by (auto intro: prod_mset_subset_imp_dvd prod_mset_primes_dvd_imp_subset) lemma is_unit_prod_mset_primes_iff: assumes "\x. x \# A \ prime x" shows "is_unit (prod_mset A) \ A = {#}" by (auto simp add: is_unit_prod_mset_iff) (meson all_not_in_conv assms not_prime_unit set_mset_eq_empty_iff) lemma prod_mset_primes_irreducible_imp_prime: assumes irred: "irreducible (prod_mset A)" assumes A: "\x. x \# A \ prime x" assumes B: "\x. x \# B \ prime x" assumes C: "\x. x \# C \ prime x" assumes dvd: "prod_mset A dvd prod_mset B * prod_mset C" shows "prod_mset A dvd prod_mset B \ prod_mset A dvd prod_mset C" proof - from dvd have "prod_mset A dvd prod_mset (B + C)" by simp with A B C have subset: "A \# B + C" by (subst (asm) prod_mset_dvd_prod_mset_primes_iff) auto define A1 and A2 where "A1 = A \# B" and "A2 = A - A1" have "A = A1 + A2" unfolding A1_def A2_def by (rule sym, intro subset_mset.add_diff_inverse) simp_all from subset have "A1 \# B" "A2 \# C" by (auto simp: A1_def A2_def Multiset.subset_eq_diff_conv Multiset.union_commute) from \A = A1 + A2\ have "prod_mset A = prod_mset A1 * prod_mset A2" by simp from irred and this have "is_unit (prod_mset A1) \ is_unit (prod_mset A2)" by (rule irreducibleD) with A have "A1 = {#} \ A2 = {#}" unfolding A1_def A2_def by (subst (asm) (1 2) is_unit_prod_mset_primes_iff) (auto dest: Multiset.in_diffD) with dvd \A = A1 + A2\ \A1 \# B\ \A2 \# C\ show ?thesis by (auto intro: prod_mset_subset_imp_dvd) qed lemma prod_mset_primes_finite_divisor_powers: assumes A: "\x. x \# A \ prime x" assumes B: "\x. x \# B \ prime x" assumes "A \ {#}" shows "finite {n. prod_mset A ^ n dvd prod_mset B}" proof - from \A \ {#}\ obtain x where x: "x \# A" by blast define m where "m = count B x" have "{n. prod_mset A ^ n dvd prod_mset B} \ {..m}" proof safe fix n assume dvd: "prod_mset A ^ n dvd prod_mset B" from x have "x ^ n dvd prod_mset A ^ n" by (intro dvd_power_same dvd_prod_mset) also note dvd also have "x ^ n = prod_mset (replicate_mset n x)" by simp finally have "replicate_mset n x \# B" by (rule prod_mset_primes_dvd_imp_subset) (insert A B x, simp_all split: if_splits) thus "n \ m" by (simp add: count_le_replicate_mset_subset_eq m_def) qed moreover have "finite {..m}" by simp ultimately show ?thesis by (rule finite_subset) qed end subsection \In a semiring with GCD, each irreducible element is a prime element\ context semiring_gcd begin lemma irreducible_imp_prime_elem_gcd: assumes "irreducible x" shows "prime_elem x" proof (rule prime_elemI) fix a b assume "x dvd a * b" from dvd_productE[OF this] obtain y z where yz: "x = y * z" "y dvd a" "z dvd b" . from \irreducible x\ and \x = y * z\ have "is_unit y \ is_unit z" by (rule irreducibleD) with yz show "x dvd a \ x dvd b" by (auto simp: mult_unit_dvd_iff mult_unit_dvd_iff') qed (insert assms, auto simp: irreducible_not_unit) lemma prime_elem_imp_coprime: assumes "prime_elem p" "\p dvd n" shows "coprime p n" proof (rule coprimeI) fix d assume "d dvd p" "d dvd n" show "is_unit d" proof (rule ccontr) assume "\is_unit d" from \prime_elem p\ and \d dvd p\ and this have "p dvd d" by (rule prime_elemD2) from this and \d dvd n\ have "p dvd n" by (rule dvd_trans) with \\p dvd n\ show False by contradiction qed qed lemma prime_imp_coprime: assumes "prime p" "\p dvd n" shows "coprime p n" using assms by (simp add: prime_elem_imp_coprime) lemma prime_elem_imp_power_coprime: "prime_elem p \ \ p dvd a \ coprime a (p ^ m)" by (cases "m > 0") (auto dest: prime_elem_imp_coprime simp add: ac_simps) lemma prime_imp_power_coprime: "prime p \ \ p dvd a \ coprime a (p ^ m)" by (rule prime_elem_imp_power_coprime) simp_all lemma prime_elem_divprod_pow: assumes p: "prime_elem p" and ab: "coprime a b" and pab: "p^n dvd a * b" shows "p^n dvd a \ p^n dvd b" using assms proof - from p have "\ is_unit p" by simp with ab p have "\ p dvd a \ \ p dvd b" using not_coprimeI by blast with p have "coprime (p ^ n) a \ coprime (p ^ n) b" by (auto dest: prime_elem_imp_power_coprime simp add: ac_simps) with pab show ?thesis by (auto simp add: coprime_dvd_mult_left_iff coprime_dvd_mult_right_iff) qed lemma primes_coprime: "prime p \ prime q \ p \ q \ coprime p q" using prime_imp_coprime primes_dvd_imp_eq by blast end subsection \Factorial semirings: algebraic structures with unique prime factorizations\ class factorial_semiring = normalization_semidom + assumes prime_factorization_exists: "x \ 0 \ \A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize x" text \Alternative characterization\ lemma (in normalization_semidom) factorial_semiring_altI_aux: assumes finite_divisors: "\x. x \ 0 \ finite {y. y dvd x \ normalize y = y}" assumes irreducible_imp_prime_elem: "\x. irreducible x \ prime_elem x" assumes "x \ 0" shows "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize x" using \x \ 0\ proof (induction "card {b. b dvd x \ normalize b = b}" arbitrary: x rule: less_induct) case (less a) let ?fctrs = "\a. {b. b dvd a \ normalize b = b}" show ?case proof (cases "is_unit a") case True thus ?thesis by (intro exI[of _ "{#}"]) (auto simp: is_unit_normalize) next case False show ?thesis proof (cases "\b. b dvd a \ \is_unit b \ \a dvd b") case False with \\is_unit a\ less.prems have "irreducible a" by (auto simp: irreducible_altdef) hence "prime_elem a" by (rule irreducible_imp_prime_elem) thus ?thesis by (intro exI[of _ "{#normalize a#}"]) auto next case True - then guess b by (elim exE conjE) note b = this - + then obtain b where b: "b dvd a" "\ is_unit b" "\ a dvd b" by auto from b have "?fctrs b \ ?fctrs a" by (auto intro: dvd_trans) moreover from b have "normalize a \ ?fctrs b" "normalize a \ ?fctrs a" by simp_all hence "?fctrs b \ ?fctrs a" by blast ultimately have "?fctrs b \ ?fctrs a" by (subst subset_not_subset_eq) blast with finite_divisors[OF \a \ 0\] have "card (?fctrs b) < card (?fctrs a)" by (rule psubset_card_mono) moreover from \a \ 0\ b have "b \ 0" by auto ultimately have "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize b" by (intro less) auto - then guess A .. note A = this + then obtain A where A: "(\x. x \# A \ prime_elem x) \ normalize (\\<^sub># A) = normalize b" + by auto define c where "c = a div b" from b have c: "a = b * c" by (simp add: c_def) from less.prems c have "c \ 0" by auto from b c have "?fctrs c \ ?fctrs a" by (auto intro: dvd_trans) moreover have "normalize a \ ?fctrs c" proof safe assume "normalize a dvd c" hence "b * c dvd 1 * c" by (simp add: c) hence "b dvd 1" by (subst (asm) dvd_times_right_cancel_iff) fact+ with b show False by simp qed with \normalize a \ ?fctrs a\ have "?fctrs a \ ?fctrs c" by blast ultimately have "?fctrs c \ ?fctrs a" by (subst subset_not_subset_eq) blast with finite_divisors[OF \a \ 0\] have "card (?fctrs c) < card (?fctrs a)" by (rule psubset_card_mono) with \c \ 0\ have "\A. (\x. x \# A \ prime_elem x) \ normalize (prod_mset A) = normalize c" by (intro less) auto - then guess B .. note B = this + then obtain B where B: "(\x. x \# B \ prime_elem x) \ normalize (\\<^sub># B) = normalize c" + by auto show ?thesis proof (rule exI[of _ "A + B"]; safe) have "normalize (prod_mset (A + B)) = normalize (normalize (prod_mset A) * normalize (prod_mset B))" by simp also have "\ = normalize (b * c)" by (simp only: A B) auto also have "b * c = a" using c by simp finally show "normalize (prod_mset (A + B)) = normalize a" . next qed (use A B in auto) qed qed qed lemma factorial_semiring_altI: assumes finite_divisors: "\x::'a. x \ 0 \ finite {y. y dvd x \ normalize y = y}" assumes irreducible_imp_prime: "\x::'a. irreducible x \ prime_elem x" shows "OFCLASS('a :: normalization_semidom, factorial_semiring_class)" by intro_classes (rule factorial_semiring_altI_aux[OF assms]) text \Properties\ context factorial_semiring begin lemma prime_factorization_exists': assumes "x \ 0" obtains A where "\x. x \# A \ prime x" "normalize (prod_mset A) = normalize x" proof - from prime_factorization_exists[OF assms] obtain A where A: "\x. x \# A \ prime_elem x" "normalize (prod_mset A) = normalize x" by blast define A' where "A' = image_mset normalize A" have "normalize (prod_mset A') = normalize (prod_mset A)" by (simp add: A'_def normalize_prod_mset_normalize) also note A(2) finally have "normalize (prod_mset A') = normalize x" by simp moreover from A(1) have "\x. x \# A' \ prime x" by (auto simp: prime_def A'_def) ultimately show ?thesis by (intro that[of A']) blast qed lemma irreducible_imp_prime_elem: assumes "irreducible x" shows "prime_elem x" proof (rule prime_elemI) fix a b assume dvd: "x dvd a * b" from assms have "x \ 0" by auto show "x dvd a \ x dvd b" proof (cases "a = 0 \ b = 0") case False hence "a \ 0" "b \ 0" by blast+ note nz = \x \ 0\ this - from nz[THEN prime_factorization_exists'] guess A B C . note ABC = this + from nz[THEN prime_factorization_exists'] obtain A B C + where ABC: + "\z. z \# A \ prime z" + "normalize (\\<^sub># A) = normalize x" + "\z. z \# B \ prime z" + "normalize (\\<^sub># B) = normalize a" + "\z. z \# C \ prime z" + "normalize (\\<^sub># C) = normalize b" + by this blast have "irreducible (prod_mset A)" by (subst irreducible_cong[OF ABC(2)]) fact moreover have "normalize (prod_mset A) dvd normalize (normalize (prod_mset B) * normalize (prod_mset C))" unfolding ABC using dvd by simp hence "prod_mset A dvd prod_mset B * prod_mset C" unfolding normalize_mult_normalize_left normalize_mult_normalize_right by simp ultimately have "prod_mset A dvd prod_mset B \ prod_mset A dvd prod_mset C" by (intro prod_mset_primes_irreducible_imp_prime) (use ABC in auto) hence "normalize (prod_mset A) dvd normalize (prod_mset B) \ normalize (prod_mset A) dvd normalize (prod_mset C)" by simp thus ?thesis unfolding ABC by simp qed auto -qed (insert assms, simp_all add: irreducible_def) +qed (use assms in \simp_all add: irreducible_def\) lemma finite_divisor_powers: assumes "y \ 0" "\is_unit x" shows "finite {n. x ^ n dvd y}" proof (cases "x = 0") case True with assms have "{n. x ^ n dvd y} = {0}" by (auto simp: power_0_left) thus ?thesis by simp next case False note nz = this \y \ 0\ - from nz[THEN prime_factorization_exists'] guess A B . note AB = this + from nz[THEN prime_factorization_exists'] obtain A B + where AB: + "\z. z \# A \ prime z" + "normalize (\\<^sub># A) = normalize x" + "\z. z \# B \ prime z" + "normalize (\\<^sub># B) = normalize y" + by this blast + from AB assms have "A \ {#}" by (auto simp: normalize_1_iff) from AB(2,4) prod_mset_primes_finite_divisor_powers [of A B, OF AB(1,3) this] have "finite {n. prod_mset A ^ n dvd prod_mset B}" by simp also have "{n. prod_mset A ^ n dvd prod_mset B} = {n. normalize (normalize (prod_mset A) ^ n) dvd normalize (prod_mset B)}" unfolding normalize_power_normalize by simp also have "\ = {n. x ^ n dvd y}" unfolding AB unfolding normalize_power_normalize by simp finally show ?thesis . qed lemma finite_prime_divisors: assumes "x \ 0" shows "finite {p. prime p \ p dvd x}" proof - - from prime_factorization_exists'[OF assms] guess A . note A = this + from prime_factorization_exists'[OF assms] obtain A + where A: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize x" by this blast have "{p. prime p \ p dvd x} \ set_mset A" proof safe fix p assume p: "prime p" and dvd: "p dvd x" from dvd have "p dvd normalize x" by simp also from A have "normalize x = normalize (prod_mset A)" by simp finally have "p dvd prod_mset A" by simp thus "p \# A" using p A by (subst (asm) prime_dvd_prod_mset_primes_iff) qed moreover have "finite (set_mset A)" by simp ultimately show ?thesis by (rule finite_subset) qed lemma infinite_unit_divisor_powers: assumes "y \ 0" assumes "is_unit x" shows "infinite {n. x^n dvd y}" proof - from `is_unit x` have "is_unit (x^n)" for n using is_unit_power_iff by auto hence "x^n dvd y" for n by auto hence "{n. x^n dvd y} = UNIV" by auto thus ?thesis by auto qed corollary is_unit_iff_infinite_divisor_powers: assumes "y \ 0" shows "is_unit x \ infinite {n. x^n dvd y}" using infinite_unit_divisor_powers finite_divisor_powers assms by auto lemma prime_elem_iff_irreducible: "prime_elem x \ irreducible x" by (blast intro: irreducible_imp_prime_elem prime_elem_imp_irreducible) lemma prime_divisor_exists: assumes "a \ 0" "\is_unit a" shows "\b. b dvd a \ prime b" proof - - from prime_factorization_exists'[OF assms(1)] guess A . note A = this - moreover from A and assms have "A \ {#}" by auto + from prime_factorization_exists'[OF assms(1)] + obtain A where A: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize a" + by this blast + with assms have "A \ {#}" by auto then obtain x where "x \# A" by blast with A(1) have *: "x dvd normalize (prod_mset A)" "prime x" by (auto simp: dvd_prod_mset) - hence "x dvd a" unfolding A by simp + hence "x dvd a" by (simp add: A(2)) with * show ?thesis by blast qed lemma prime_divisors_induct [case_names zero unit factor]: assumes "P 0" "\x. is_unit x \ P x" "\p x. prime p \ P x \ P (p * x)" shows "P x" proof (cases "x = 0") case False - from prime_factorization_exists'[OF this] guess A . note A = this + from prime_factorization_exists'[OF this] + obtain A where A: "\z. z \# A \ prime z" "normalize (\\<^sub># A) = normalize x" + by this blast from A obtain u where u: "is_unit u" "x = u * prod_mset A" by (elim associatedE2) from A(1) have "P (u * prod_mset A)" proof (induction A) case (add p A) from add.prems have "prime p" by simp moreover from add.prems have "P (u * prod_mset A)" by (intro add.IH) simp_all ultimately have "P (p * (u * prod_mset A))" by (rule assms(3)) thus ?case by (simp add: mult_ac) qed (simp_all add: assms False u) with A u show ?thesis by simp qed (simp_all add: assms(1)) lemma no_prime_divisors_imp_unit: assumes "a \ 0" "\b. b dvd a \ normalize b = b \ \ prime_elem b" shows "is_unit a" proof (rule ccontr) assume "\is_unit a" - from prime_divisor_exists[OF assms(1) this] guess b by (elim exE conjE) + from prime_divisor_exists[OF assms(1) this] obtain b where "b dvd a" "prime b" by auto with assms(2)[of b] show False by (simp add: prime_def) qed lemma prime_divisorE: assumes "a \ 0" and "\ is_unit a" obtains p where "prime p" and "p dvd a" using assms no_prime_divisors_imp_unit unfolding prime_def by blast definition multiplicity :: "'a \ 'a \ nat" where "multiplicity p x = (if finite {n. p ^ n dvd x} then Max {n. p ^ n dvd x} else 0)" lemma multiplicity_dvd: "p ^ multiplicity p x dvd x" proof (cases "finite {n. p ^ n dvd x}") case True hence "multiplicity p x = Max {n. p ^ n dvd x}" by (simp add: multiplicity_def) also have "\ \ {n. p ^ n dvd x}" by (rule Max_in) (auto intro!: True exI[of _ "0::nat"]) finally show ?thesis by simp qed (simp add: multiplicity_def) lemma multiplicity_dvd': "n \ multiplicity p x \ p ^ n dvd x" by (rule dvd_trans[OF le_imp_power_dvd multiplicity_dvd]) context fixes x p :: 'a assumes xp: "x \ 0" "\is_unit p" begin lemma multiplicity_eq_Max: "multiplicity p x = Max {n. p ^ n dvd x}" using finite_divisor_powers[OF xp] by (simp add: multiplicity_def) lemma multiplicity_geI: assumes "p ^ n dvd x" shows "multiplicity p x \ n" proof - from assms have "n \ Max {n. p ^ n dvd x}" by (intro Max_ge finite_divisor_powers xp) simp_all thus ?thesis by (subst multiplicity_eq_Max) qed lemma multiplicity_lessI: assumes "\p ^ n dvd x" shows "multiplicity p x < n" proof (rule ccontr) assume "\(n > multiplicity p x)" hence "p ^ n dvd x" by (intro multiplicity_dvd') simp with assms show False by contradiction qed lemma power_dvd_iff_le_multiplicity: "p ^ n dvd x \ n \ multiplicity p x" using multiplicity_geI[of n] multiplicity_lessI[of n] by (cases "p ^ n dvd x") auto lemma multiplicity_eq_zero_iff: shows "multiplicity p x = 0 \ \p dvd x" using power_dvd_iff_le_multiplicity[of 1] by auto lemma multiplicity_gt_zero_iff: shows "multiplicity p x > 0 \ p dvd x" using power_dvd_iff_le_multiplicity[of 1] by auto lemma multiplicity_decompose: "\p dvd (x div p ^ multiplicity p x)" proof assume *: "p dvd x div p ^ multiplicity p x" have "x = x div p ^ multiplicity p x * (p ^ multiplicity p x)" using multiplicity_dvd[of p x] by simp also from * have "x div p ^ multiplicity p x = (x div p ^ multiplicity p x div p) * p" by simp also have "x div p ^ multiplicity p x div p * p * p ^ multiplicity p x = x div p ^ multiplicity p x div p * p ^ Suc (multiplicity p x)" by (simp add: mult_assoc) also have "p ^ Suc (multiplicity p x) dvd \" by (rule dvd_triv_right) finally show False by (subst (asm) power_dvd_iff_le_multiplicity) simp qed lemma multiplicity_decompose': obtains y where "x = p ^ multiplicity p x * y" "\p dvd y" using that[of "x div p ^ multiplicity p x"] by (simp add: multiplicity_decompose multiplicity_dvd) end lemma multiplicity_zero [simp]: "multiplicity p 0 = 0" by (simp add: multiplicity_def) lemma prime_elem_multiplicity_eq_zero_iff: "prime_elem p \ x \ 0 \ multiplicity p x = 0 \ \p dvd x" by (rule multiplicity_eq_zero_iff) simp_all lemma prime_multiplicity_other: assumes "prime p" "prime q" "p \ q" shows "multiplicity p q = 0" using assms by (subst prime_elem_multiplicity_eq_zero_iff) (auto dest: primes_dvd_imp_eq) lemma prime_multiplicity_gt_zero_iff: "prime_elem p \ x \ 0 \ multiplicity p x > 0 \ p dvd x" by (rule multiplicity_gt_zero_iff) simp_all lemma multiplicity_unit_left: "is_unit p \ multiplicity p x = 0" by (simp add: multiplicity_def is_unit_power_iff unit_imp_dvd) lemma multiplicity_unit_right: assumes "is_unit x" shows "multiplicity p x = 0" proof (cases "is_unit p \ x = 0") case False with multiplicity_lessI[of x p 1] this assms show ?thesis by (auto dest: dvd_unit_imp_unit) qed (auto simp: multiplicity_unit_left) lemma multiplicity_one [simp]: "multiplicity p 1 = 0" by (rule multiplicity_unit_right) simp_all lemma multiplicity_eqI: assumes "p ^ n dvd x" "\p ^ Suc n dvd x" shows "multiplicity p x = n" proof - consider "x = 0" | "is_unit p" | "x \ 0" "\is_unit p" by blast thus ?thesis proof cases assume xp: "x \ 0" "\is_unit p" from xp assms(1) have "multiplicity p x \ n" by (intro multiplicity_geI) moreover from assms(2) xp have "multiplicity p x < Suc n" by (intro multiplicity_lessI) ultimately show ?thesis by simp next assume "is_unit p" hence "is_unit (p ^ Suc n)" by (simp add: is_unit_power_iff del: power_Suc) hence "p ^ Suc n dvd x" by (rule unit_imp_dvd) with \\p ^ Suc n dvd x\ show ?thesis by contradiction qed (insert assms, simp_all) qed context fixes x p :: 'a assumes xp: "x \ 0" "\is_unit p" begin lemma multiplicity_times_same: assumes "p \ 0" shows "multiplicity p (p * x) = Suc (multiplicity p x)" proof (rule multiplicity_eqI) show "p ^ Suc (multiplicity p x) dvd p * x" by (auto intro!: mult_dvd_mono multiplicity_dvd) from xp assms show "\ p ^ Suc (Suc (multiplicity p x)) dvd p * x" using power_dvd_iff_le_multiplicity[OF xp, of "Suc (multiplicity p x)"] by simp qed end lemma multiplicity_same_power': "multiplicity p (p ^ n) = (if p = 0 \ is_unit p then 0 else n)" proof - consider "p = 0" | "is_unit p" |"p \ 0" "\is_unit p" by blast thus ?thesis proof cases assume "p \ 0" "\is_unit p" thus ?thesis by (induction n) (simp_all add: multiplicity_times_same) qed (simp_all add: power_0_left multiplicity_unit_left) qed lemma multiplicity_same_power: "p \ 0 \ \is_unit p \ multiplicity p (p ^ n) = n" by (simp add: multiplicity_same_power') lemma multiplicity_prime_elem_times_other: assumes "prime_elem p" "\p dvd q" shows "multiplicity p (q * x) = multiplicity p x" proof (cases "x = 0") case False show ?thesis proof (rule multiplicity_eqI) have "1 * p ^ multiplicity p x dvd q * x" by (intro mult_dvd_mono multiplicity_dvd) simp_all thus "p ^ multiplicity p x dvd q * x" by simp next define n where "n = multiplicity p x" from assms have "\is_unit p" by simp - from multiplicity_decompose'[OF False this] guess y . note y = this [folded n_def] + from multiplicity_decompose'[OF False this] + obtain y where y [folded n_def]: "x = p ^ multiplicity p x * y" "\ p dvd y" . from y have "p ^ Suc n dvd q * x \ p ^ n * p dvd p ^ n * (q * y)" by (simp add: mult_ac) also from assms have "\ \ p dvd q * y" by simp also have "\ \ p dvd q \ p dvd y" by (rule prime_elem_dvd_mult_iff) fact+ also from assms y have "\ \ False" by simp finally show "\(p ^ Suc n dvd q * x)" by blast qed qed simp_all lemma multiplicity_self: assumes "p \ 0" "\is_unit p" shows "multiplicity p p = 1" proof - from assms have "multiplicity p p = Max {n. p ^ n dvd p}" by (simp add: multiplicity_eq_Max) also from assms have "p ^ n dvd p \ n \ 1" for n using dvd_power_iff[of p n 1] by auto hence "{n. p ^ n dvd p} = {..1}" by auto also have "\ = {0,1}" by auto finally show ?thesis by simp qed lemma multiplicity_times_unit_left: assumes "is_unit c" shows "multiplicity (c * p) x = multiplicity p x" proof - from assms have "{n. (c * p) ^ n dvd x} = {n. p ^ n dvd x}" by (subst mult.commute) (simp add: mult_unit_dvd_iff power_mult_distrib is_unit_power_iff) thus ?thesis by (simp add: multiplicity_def) qed lemma multiplicity_times_unit_right: assumes "is_unit c" shows "multiplicity p (c * x) = multiplicity p x" proof - from assms have "{n. p ^ n dvd c * x} = {n. p ^ n dvd x}" by (subst mult.commute) (simp add: dvd_mult_unit_iff) thus ?thesis by (simp add: multiplicity_def) qed lemma multiplicity_normalize_left [simp]: "multiplicity (normalize p) x = multiplicity p x" proof (cases "p = 0") case [simp]: False have "normalize p = (1 div unit_factor p) * p" by (simp add: unit_div_commute is_unit_unit_factor) also have "multiplicity \ x = multiplicity p x" by (rule multiplicity_times_unit_left) (simp add: is_unit_unit_factor) finally show ?thesis . qed simp_all lemma multiplicity_normalize_right [simp]: "multiplicity p (normalize x) = multiplicity p x" proof (cases "x = 0") case [simp]: False have "normalize x = (1 div unit_factor x) * x" by (simp add: unit_div_commute is_unit_unit_factor) also have "multiplicity p \ = multiplicity p x" by (rule multiplicity_times_unit_right) (simp add: is_unit_unit_factor) finally show ?thesis . qed simp_all lemma multiplicity_prime [simp]: "prime_elem p \ multiplicity p p = 1" by (rule multiplicity_self) auto lemma multiplicity_prime_power [simp]: "prime_elem p \ multiplicity p (p ^ n) = n" by (subst multiplicity_same_power') auto lift_definition prime_factorization :: "'a \ 'a multiset" is "\x p. if prime p then multiplicity p x else 0" proof - fix x :: 'a show "finite {p. 0 < (if prime p then multiplicity p x else 0)}" (is "finite ?A") proof (cases "x = 0") case False from False have "?A \ {p. prime p \ p dvd x}" by (auto simp: multiplicity_gt_zero_iff) moreover from False have "finite {p. prime p \ p dvd x}" by (rule finite_prime_divisors) ultimately show ?thesis by (rule finite_subset) qed simp_all qed abbreviation prime_factors :: "'a \ 'a set" where "prime_factors a \ set_mset (prime_factorization a)" lemma count_prime_factorization_nonprime: "\prime p \ count (prime_factorization x) p = 0" by transfer simp lemma count_prime_factorization_prime: "prime p \ count (prime_factorization x) p = multiplicity p x" by transfer simp lemma count_prime_factorization: "count (prime_factorization x) p = (if prime p then multiplicity p x else 0)" by transfer simp lemma dvd_imp_multiplicity_le: assumes "a dvd b" "b \ 0" shows "multiplicity p a \ multiplicity p b" proof (cases "is_unit p") case False with assms show ?thesis by (intro multiplicity_geI ) (auto intro: dvd_trans[OF multiplicity_dvd' assms(1)]) qed (insert assms, auto simp: multiplicity_unit_left) lemma prime_power_inj: assumes "prime a" "a ^ m = a ^ n" shows "m = n" proof - have "multiplicity a (a ^ m) = multiplicity a (a ^ n)" by (simp only: assms) thus ?thesis using assms by (subst (asm) (1 2) multiplicity_prime_power) simp_all qed lemma prime_power_inj': assumes "prime p" "prime q" assumes "p ^ m = q ^ n" "m > 0" "n > 0" shows "p = q" "m = n" proof - from assms have "p ^ 1 dvd p ^ m" by (intro le_imp_power_dvd) simp also have "p ^ m = q ^ n" by fact finally have "p dvd q ^ n" by simp with assms have "p dvd q" using prime_dvd_power[of p q] by simp with assms show "p = q" by (simp add: primes_dvd_imp_eq) with assms show "m = n" by (simp add: prime_power_inj) qed lemma prime_power_eq_one_iff [simp]: "prime p \ p ^ n = 1 \ n = 0" using prime_power_inj[of p n 0] by auto lemma one_eq_prime_power_iff [simp]: "prime p \ 1 = p ^ n \ n = 0" using prime_power_inj[of p 0 n] by auto lemma prime_power_inj'': assumes "prime p" "prime q" shows "p ^ m = q ^ n \ (m = 0 \ n = 0) \ (p = q \ m = n)" using assms by (cases "m = 0"; cases "n = 0") (auto dest: prime_power_inj'[OF assms]) lemma prime_factorization_0 [simp]: "prime_factorization 0 = {#}" by (simp add: multiset_eq_iff count_prime_factorization) lemma prime_factorization_empty_iff: "prime_factorization x = {#} \ x = 0 \ is_unit x" proof assume *: "prime_factorization x = {#}" { assume x: "x \ 0" "\is_unit x" { fix p assume p: "prime p" have "count (prime_factorization x) p = 0" by (simp add: *) also from p have "count (prime_factorization x) p = multiplicity p x" by (rule count_prime_factorization_prime) also from x p have "\ = 0 \ \p dvd x" by (simp add: multiplicity_eq_zero_iff) finally have "\p dvd x" . } with prime_divisor_exists[OF x] have False by blast } thus "x = 0 \ is_unit x" by blast next assume "x = 0 \ is_unit x" thus "prime_factorization x = {#}" proof assume x: "is_unit x" { fix p assume p: "prime p" from p x have "multiplicity p x = 0" by (subst multiplicity_eq_zero_iff) (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors) } thus ?thesis by (simp add: multiset_eq_iff count_prime_factorization) qed simp_all qed lemma prime_factorization_unit: assumes "is_unit x" shows "prime_factorization x = {#}" proof (rule multiset_eqI) fix p :: 'a show "count (prime_factorization x) p = count {#} p" proof (cases "prime p") case True with assms have "multiplicity p x = 0" by (subst multiplicity_eq_zero_iff) (auto simp: multiplicity_eq_zero_iff dest: unit_imp_no_prime_divisors) with True show ?thesis by (simp add: count_prime_factorization_prime) qed (simp_all add: count_prime_factorization_nonprime) qed lemma prime_factorization_1 [simp]: "prime_factorization 1 = {#}" by (simp add: prime_factorization_unit) lemma prime_factorization_times_prime: assumes "x \ 0" "prime p" shows "prime_factorization (p * x) = {#p#} + prime_factorization x" proof (rule multiset_eqI) fix q :: 'a consider "\prime q" | "p = q" | "prime q" "p \ q" by blast thus "count (prime_factorization (p * x)) q = count ({#p#} + prime_factorization x) q" proof cases assume q: "prime q" "p \ q" with assms primes_dvd_imp_eq[of q p] have "\q dvd p" by auto with q assms show ?thesis by (simp add: multiplicity_prime_elem_times_other count_prime_factorization) qed (insert assms, auto simp: count_prime_factorization multiplicity_times_same) qed lemma prod_mset_prime_factorization_weak: assumes "x \ 0" shows "normalize (prod_mset (prime_factorization x)) = normalize x" using assms proof (induction x rule: prime_divisors_induct) case (factor p x) have "normalize (prod_mset (prime_factorization (p * x))) = normalize (p * normalize (prod_mset (prime_factorization x)))" using factor.prems factor.hyps by (simp add: prime_factorization_times_prime) also have "normalize (prod_mset (prime_factorization x)) = normalize x" by (rule factor.IH) (use factor in auto) finally show ?case by simp qed (auto simp: prime_factorization_unit is_unit_normalize) lemma in_prime_factors_iff: "p \ prime_factors x \ x \ 0 \ p dvd x \ prime p" proof - have "p \ prime_factors x \ count (prime_factorization x) p > 0" by simp also have "\ \ x \ 0 \ p dvd x \ prime p" by (subst count_prime_factorization, cases "x = 0") (auto simp: multiplicity_eq_zero_iff multiplicity_gt_zero_iff) finally show ?thesis . qed lemma in_prime_factors_imp_prime [intro]: "p \ prime_factors x \ prime p" by (simp add: in_prime_factors_iff) lemma in_prime_factors_imp_dvd [dest]: "p \ prime_factors x \ p dvd x" by (simp add: in_prime_factors_iff) lemma prime_factorsI: "x \ 0 \ prime p \ p dvd x \ p \ prime_factors x" by (auto simp: in_prime_factors_iff) lemma prime_factors_dvd: "x \ 0 \ prime_factors x = {p. prime p \ p dvd x}" by (auto intro: prime_factorsI) lemma prime_factors_multiplicity: "prime_factors n = {p. prime p \ multiplicity p n > 0}" by (cases "n = 0") (auto simp add: prime_factors_dvd prime_multiplicity_gt_zero_iff) lemma prime_factorization_prime: assumes "prime p" shows "prime_factorization p = {#p#}" proof (rule multiset_eqI) fix q :: 'a consider "\prime q" | "q = p" | "prime q" "q \ p" by blast thus "count (prime_factorization p) q = count {#p#} q" by cases (insert assms, auto dest: primes_dvd_imp_eq simp: count_prime_factorization multiplicity_self multiplicity_eq_zero_iff) qed lemma prime_factorization_prod_mset_primes: assumes "\p. p \# A \ prime p" shows "prime_factorization (prod_mset A) = A" using assms proof (induction A) case (add p A) from add.prems[of 0] have "0 \# A" by auto hence "prod_mset A \ 0" by auto with add show ?case by (simp_all add: mult_ac prime_factorization_times_prime Multiset.union_commute) qed simp_all lemma prime_factorization_cong: "normalize x = normalize y \ prime_factorization x = prime_factorization y" by (simp add: multiset_eq_iff count_prime_factorization multiplicity_normalize_right [of _ x, symmetric] multiplicity_normalize_right [of _ y, symmetric] del: multiplicity_normalize_right) lemma prime_factorization_unique: assumes "x \ 0" "y \ 0" shows "prime_factorization x = prime_factorization y \ normalize x = normalize y" proof assume "prime_factorization x = prime_factorization y" hence "prod_mset (prime_factorization x) = prod_mset (prime_factorization y)" by simp hence "normalize (prod_mset (prime_factorization x)) = normalize (prod_mset (prime_factorization y))" by (simp only: ) with assms show "normalize x = normalize y" by (simp add: prod_mset_prime_factorization_weak) qed (rule prime_factorization_cong) lemma prime_factorization_normalize [simp]: "prime_factorization (normalize x) = prime_factorization x" by (cases "x = 0", simp, subst prime_factorization_unique) auto lemma prime_factorization_eqI_strong: assumes "\p. p \# P \ prime p" "prod_mset P = n" shows "prime_factorization n = P" using prime_factorization_prod_mset_primes[of P] assms by simp lemma prime_factorization_eqI: assumes "\p. p \# P \ prime p" "normalize (prod_mset P) = normalize n" shows "prime_factorization n = P" proof - have "P = prime_factorization (normalize (prod_mset P))" using prime_factorization_prod_mset_primes[of P] assms(1) by simp with assms(2) show ?thesis by simp qed lemma prime_factorization_mult: assumes "x \ 0" "y \ 0" shows "prime_factorization (x * y) = prime_factorization x + prime_factorization y" proof - have "normalize (prod_mset (prime_factorization x) * prod_mset (prime_factorization y)) = normalize (normalize (prod_mset (prime_factorization x)) * normalize (prod_mset (prime_factorization y)))" by (simp only: normalize_mult_normalize_left normalize_mult_normalize_right) also have "\ = normalize (x * y)" by (subst (1 2) prod_mset_prime_factorization_weak) (use assms in auto) finally show ?thesis by (intro prime_factorization_eqI) auto qed lemma prime_factorization_prod: assumes "finite A" "\x. x \ A \ f x \ 0" shows "prime_factorization (prod f A) = (\n\A. prime_factorization (f n))" using assms by (induction A rule: finite_induct) (auto simp: Sup_multiset_empty prime_factorization_mult) lemma prime_elem_multiplicity_mult_distrib: assumes "prime_elem p" "x \ 0" "y \ 0" shows "multiplicity p (x * y) = multiplicity p x + multiplicity p y" proof - have "multiplicity p (x * y) = count (prime_factorization (x * y)) (normalize p)" by (subst count_prime_factorization_prime) (simp_all add: assms) also from assms have "prime_factorization (x * y) = prime_factorization x + prime_factorization y" by (intro prime_factorization_mult) also have "count \ (normalize p) = count (prime_factorization x) (normalize p) + count (prime_factorization y) (normalize p)" by simp also have "\ = multiplicity p x + multiplicity p y" by (subst (1 2) count_prime_factorization_prime) (simp_all add: assms) finally show ?thesis . qed lemma prime_elem_multiplicity_prod_mset_distrib: assumes "prime_elem p" "0 \# A" shows "multiplicity p (prod_mset A) = sum_mset (image_mset (multiplicity p) A)" using assms by (induction A) (auto simp: prime_elem_multiplicity_mult_distrib) lemma prime_elem_multiplicity_power_distrib: assumes "prime_elem p" "x \ 0" shows "multiplicity p (x ^ n) = n * multiplicity p x" using assms prime_elem_multiplicity_prod_mset_distrib [of p "replicate_mset n x"] by simp lemma prime_elem_multiplicity_prod_distrib: assumes "prime_elem p" "0 \ f ` A" "finite A" shows "multiplicity p (prod f A) = (\x\A. multiplicity p (f x))" proof - have "multiplicity p (prod f A) = (\x\#mset_set A. multiplicity p (f x))" using assms by (subst prod_unfold_prod_mset) (simp_all add: prime_elem_multiplicity_prod_mset_distrib sum_unfold_sum_mset multiset.map_comp o_def) also from \finite A\ have "\ = (\x\A. multiplicity p (f x))" by (induction A rule: finite_induct) simp_all finally show ?thesis . qed lemma multiplicity_distinct_prime_power: "prime p \ prime q \ p \ q \ multiplicity p (q ^ n) = 0" by (subst prime_elem_multiplicity_power_distrib) (auto simp: prime_multiplicity_other) lemma prime_factorization_prime_power: "prime p \ prime_factorization (p ^ n) = replicate_mset n p" by (induction n) (simp_all add: prime_factorization_mult prime_factorization_prime Multiset.union_commute) lemma prime_factorization_subset_iff_dvd: assumes [simp]: "x \ 0" "y \ 0" shows "prime_factorization x \# prime_factorization y \ x dvd y" proof - have "x dvd y \ normalize (prod_mset (prime_factorization x)) dvd normalize (prod_mset (prime_factorization y))" using assms by (subst (1 2) prod_mset_prime_factorization_weak) auto also have "\ \ prime_factorization x \# prime_factorization y" by (auto intro!: prod_mset_primes_dvd_imp_subset prod_mset_subset_imp_dvd) finally show ?thesis .. qed lemma prime_factorization_subset_imp_dvd: "x \ 0 \ (prime_factorization x \# prime_factorization y) \ x dvd y" by (cases "y = 0") (simp_all add: prime_factorization_subset_iff_dvd) lemma prime_factorization_divide: assumes "b dvd a" shows "prime_factorization (a div b) = prime_factorization a - prime_factorization b" proof (cases "a = 0") case [simp]: False from assms have [simp]: "b \ 0" by auto have "prime_factorization ((a div b) * b) = prime_factorization (a div b) + prime_factorization b" by (intro prime_factorization_mult) (insert assms, auto elim!: dvdE) with assms show ?thesis by simp qed simp_all lemma zero_not_in_prime_factors [simp]: "0 \ prime_factors x" by (auto dest: in_prime_factors_imp_prime) lemma prime_prime_factors: "prime p \ prime_factors p = {p}" by (drule prime_factorization_prime) simp lemma prime_factors_product: "x \ 0 \ y \ 0 \ prime_factors (x * y) = prime_factors x \ prime_factors y" by (simp add: prime_factorization_mult) lemma dvd_prime_factors [intro]: "y \ 0 \ x dvd y \ prime_factors x \ prime_factors y" by (intro set_mset_mono, subst prime_factorization_subset_iff_dvd) auto (* RENAMED multiplicity_dvd *) lemma multiplicity_le_imp_dvd: assumes "x \ 0" "\p. prime p \ multiplicity p x \ multiplicity p y" shows "x dvd y" proof (cases "y = 0") case False from assms this have "prime_factorization x \# prime_factorization y" by (intro mset_subset_eqI) (auto simp: count_prime_factorization) with assms False show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) qed auto lemma dvd_multiplicity_eq: "x \ 0 \ y \ 0 \ x dvd y \ (\p. multiplicity p x \ multiplicity p y)" by (auto intro: dvd_imp_multiplicity_le multiplicity_le_imp_dvd) lemma multiplicity_eq_imp_eq: assumes "x \ 0" "y \ 0" assumes "\p. prime p \ multiplicity p x = multiplicity p y" shows "normalize x = normalize y" using assms by (intro associatedI multiplicity_le_imp_dvd) simp_all lemma prime_factorization_unique': assumes "\p \# M. prime p" "\p \# N. prime p" "(\i \# M. i) = (\i \# N. i)" shows "M = N" proof - have "prime_factorization (\i \# M. i) = prime_factorization (\i \# N. i)" by (simp only: assms) also from assms have "prime_factorization (\i \# M. i) = M" by (subst prime_factorization_prod_mset_primes) simp_all also from assms have "prime_factorization (\i \# N. i) = N" by (subst prime_factorization_prod_mset_primes) simp_all finally show ?thesis . qed lemma prime_factorization_unique'': assumes "\p \# M. prime p" "\p \# N. prime p" "normalize (\i \# M. i) = normalize (\i \# N. i)" shows "M = N" proof - have "prime_factorization (normalize (\i \# M. i)) = prime_factorization (normalize (\i \# N. i))" by (simp only: assms) also from assms have "prime_factorization (normalize (\i \# M. i)) = M" by (subst prime_factorization_normalize, subst prime_factorization_prod_mset_primes) simp_all also from assms have "prime_factorization (normalize (\i \# N. i)) = N" by (subst prime_factorization_normalize, subst prime_factorization_prod_mset_primes) simp_all finally show ?thesis . qed lemma multiplicity_cong: "(\r. p ^ r dvd a \ p ^ r dvd b) \ multiplicity p a = multiplicity p b" by (simp add: multiplicity_def) lemma not_dvd_imp_multiplicity_0: assumes "\p dvd x" shows "multiplicity p x = 0" proof - from assms have "multiplicity p x < 1" by (intro multiplicity_lessI) auto thus ?thesis by simp qed lemma multiplicity_zero_left [simp]: "multiplicity 0 x = 0" by (cases "x = 0") (auto intro: not_dvd_imp_multiplicity_0) lemma inj_on_Prod_primes: assumes "\P p. P \ A \ p \ P \ prime p" assumes "\P. P \ A \ finite P" shows "inj_on Prod A" proof (rule inj_onI) fix P Q assume PQ: "P \ A" "Q \ A" "\P = \Q" with prime_factorization_unique'[of "mset_set P" "mset_set Q"] assms[of P] assms[of Q] have "mset_set P = mset_set Q" by (auto simp: prod_unfold_prod_mset) with assms[of P] assms[of Q] PQ show "P = Q" by simp qed lemma divides_primepow_weak: assumes "prime p" and "a dvd p ^ n" obtains m where "m \ n" and "normalize a = normalize (p ^ m)" proof - from assms have "a \ 0" by auto with assms have "normalize (prod_mset (prime_factorization a)) dvd normalize (prod_mset (prime_factorization (p ^ n)))" by (subst (1 2) prod_mset_prime_factorization_weak) auto then have "prime_factorization a \# prime_factorization (p ^ n)" by (simp add: in_prime_factors_imp_prime prod_mset_dvd_prod_mset_primes_iff) with assms have "prime_factorization a \# replicate_mset n p" by (simp add: prime_factorization_prime_power) then obtain m where "m \ n" and "prime_factorization a = replicate_mset m p" by (rule msubseteq_replicate_msetE) then have *: "normalize (prod_mset (prime_factorization a)) = normalize (prod_mset (replicate_mset m p))" by metis also have "normalize (prod_mset (prime_factorization a)) = normalize a" using \a \ 0\ by (simp add: prod_mset_prime_factorization_weak) also have "prod_mset (replicate_mset m p) = p ^ m" by simp finally show ?thesis using \m \ n\ by (intro that[of m]) qed lemma divide_out_primepow_ex: assumes "n \ 0" "\p\prime_factors n. P p" obtains p k n' where "P p" "prime p" "p dvd n" "\p dvd n'" "k > 0" "n = p ^ k * n'" proof - from assms obtain p where p: "P p" "prime p" "p dvd n" by auto define k where "k = multiplicity p n" define n' where "n' = n div p ^ k" have n': "n = p ^ k * n'" "\p dvd n'" using assms p multiplicity_decompose[of n p] by (auto simp: n'_def k_def multiplicity_dvd) from n' p have "k > 0" by (intro Nat.gr0I) auto with n' p that[of p n' k] show ?thesis by auto qed lemma divide_out_primepow: assumes "n \ 0" "\is_unit n" obtains p k n' where "prime p" "p dvd n" "\p dvd n'" "k > 0" "n = p ^ k * n'" using divide_out_primepow_ex[OF assms(1), of "\_. True"] prime_divisor_exists[OF assms] assms prime_factorsI by metis subsection \GCD and LCM computation with unique factorizations\ definition "gcd_factorial a b = (if a = 0 then normalize b else if b = 0 then normalize a else normalize (prod_mset (prime_factorization a \# prime_factorization b)))" definition "lcm_factorial a b = (if a = 0 \ b = 0 then 0 else normalize (prod_mset (prime_factorization a \# prime_factorization b)))" definition "Gcd_factorial A = (if A \ {0} then 0 else normalize (prod_mset (Inf (prime_factorization ` (A - {0})))))" definition "Lcm_factorial A = (if A = {} then 1 else if 0 \ A \ subset_mset.bdd_above (prime_factorization ` (A - {0})) then normalize (prod_mset (Sup (prime_factorization ` A))) else 0)" lemma prime_factorization_gcd_factorial: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (gcd_factorial a b) = prime_factorization a \# prime_factorization b" proof - have "prime_factorization (gcd_factorial a b) = prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" by (simp add: gcd_factorial_def) also have "\ = prime_factorization a \# prime_factorization b" by (subst prime_factorization_prod_mset_primes) auto finally show ?thesis . qed lemma prime_factorization_lcm_factorial: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (lcm_factorial a b) = prime_factorization a \# prime_factorization b" proof - have "prime_factorization (lcm_factorial a b) = prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" by (simp add: lcm_factorial_def) also have "\ = prime_factorization a \# prime_factorization b" by (subst prime_factorization_prod_mset_primes) auto finally show ?thesis . qed lemma prime_factorization_Gcd_factorial: assumes "\A \ {0}" shows "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))" proof - from assms obtain x where x: "x \ A - {0}" by auto hence "Inf (prime_factorization ` (A - {0})) \# prime_factorization x" by (intro subset_mset.cInf_lower) simp_all hence "\y. y \# Inf (prime_factorization ` (A - {0})) \ y \ prime_factors x" by (auto dest: mset_subset_eqD) with in_prime_factors_imp_prime[of _ x] have "\p. p \# Inf (prime_factorization ` (A - {0})) \ prime p" by blast with assms show ?thesis by (simp add: Gcd_factorial_def prime_factorization_prod_mset_primes) qed lemma prime_factorization_Lcm_factorial: assumes "0 \ A" "subset_mset.bdd_above (prime_factorization ` A)" shows "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)" proof (cases "A = {}") case True hence "prime_factorization ` A = {}" by auto also have "Sup \ = {#}" by (simp add: Sup_multiset_empty) finally show ?thesis by (simp add: Lcm_factorial_def) next case False have "\y. y \# Sup (prime_factorization ` A) \ prime y" by (auto simp: in_Sup_multiset_iff assms) with assms False show ?thesis by (simp add: Lcm_factorial_def prime_factorization_prod_mset_primes) qed lemma gcd_factorial_commute: "gcd_factorial a b = gcd_factorial b a" by (simp add: gcd_factorial_def multiset_inter_commute) lemma gcd_factorial_dvd1: "gcd_factorial a b dvd a" proof (cases "a = 0 \ b = 0") case False hence "gcd_factorial a b \ 0" by (auto simp: gcd_factorial_def) with False show ?thesis by (subst prime_factorization_subset_iff_dvd [symmetric]) (auto simp: prime_factorization_gcd_factorial) qed (auto simp: gcd_factorial_def) lemma gcd_factorial_dvd2: "gcd_factorial a b dvd b" by (subst gcd_factorial_commute) (rule gcd_factorial_dvd1) lemma normalize_gcd_factorial [simp]: "normalize (gcd_factorial a b) = gcd_factorial a b" by (simp add: gcd_factorial_def) lemma normalize_lcm_factorial [simp]: "normalize (lcm_factorial a b) = lcm_factorial a b" by (simp add: lcm_factorial_def) lemma gcd_factorial_greatest: "c dvd gcd_factorial a b" if "c dvd a" "c dvd b" for a b c proof (cases "a = 0 \ b = 0") case False with that have [simp]: "c \ 0" by auto let ?p = "prime_factorization" from that False have "?p c \# ?p a" "?p c \# ?p b" by (simp_all add: prime_factorization_subset_iff_dvd) hence "prime_factorization c \# prime_factorization (prod_mset (prime_factorization a \# prime_factorization b))" using False by (subst prime_factorization_prod_mset_primes) auto with False show ?thesis by (auto simp: gcd_factorial_def prime_factorization_subset_iff_dvd [symmetric]) qed (auto simp: gcd_factorial_def that) lemma lcm_factorial_gcd_factorial: "lcm_factorial a b = normalize (a * b div gcd_factorial a b)" for a b proof (cases "a = 0 \ b = 0") case False let ?p = "prime_factorization" have 1: "normalize x * normalize y dvd z \ x * y dvd z" for x y z :: 'a proof - have "normalize (normalize x * normalize y) dvd z \ x * y dvd z" unfolding normalize_mult_normalize_left normalize_mult_normalize_right by simp thus ?thesis unfolding normalize_dvd_iff by simp qed have "?p (a * b) = (?p a \# ?p b) + (?p a \# ?p b)" using False by (subst prime_factorization_mult) (auto intro!: multiset_eqI) hence "normalize (prod_mset (?p (a * b))) = normalize (prod_mset ((?p a \# ?p b) + (?p a \# ?p b)))" by (simp only:) hence *: "normalize (a * b) = normalize (lcm_factorial a b * gcd_factorial a b)" using False by (subst (asm) prod_mset_prime_factorization_weak) (auto simp: lcm_factorial_def gcd_factorial_def) have [simp]: "gcd_factorial a b dvd a * b" "lcm_factorial a b dvd a * b" using associatedD2[OF *] by auto from False have [simp]: "gcd_factorial a b \ 0" "lcm_factorial a b \ 0" by (auto simp: gcd_factorial_def lcm_factorial_def) show ?thesis by (rule associated_eqI) (use * in \auto simp: dvd_div_iff_mult div_dvd_iff_mult dest: associatedD1 associatedD2\) qed (auto simp: lcm_factorial_def) lemma normalize_Gcd_factorial: "normalize (Gcd_factorial A) = Gcd_factorial A" by (simp add: Gcd_factorial_def) lemma Gcd_factorial_eq_0_iff: "Gcd_factorial A = 0 \ A \ {0}" by (auto simp: Gcd_factorial_def in_Inf_multiset_iff split: if_splits) lemma Gcd_factorial_dvd: assumes "x \ A" shows "Gcd_factorial A dvd x" proof (cases "x = 0") case False with assms have "prime_factorization (Gcd_factorial A) = Inf (prime_factorization ` (A - {0}))" by (intro prime_factorization_Gcd_factorial) auto also from False assms have "\ \# prime_factorization x" by (intro subset_mset.cInf_lower) auto finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert assms False, auto simp: Gcd_factorial_eq_0_iff) qed simp_all lemma Gcd_factorial_greatest: assumes "\y. y \ A \ x dvd y" shows "x dvd Gcd_factorial A" proof (cases "A \ {0}") case False from False obtain y where "y \ A" "y \ 0" by auto with assms[of y] have nz: "x \ 0" by auto from nz assms have "prime_factorization x \# prime_factorization y" if "y \ A - {0}" for y using that by (subst prime_factorization_subset_iff_dvd) auto with False have "prime_factorization x \# Inf (prime_factorization ` (A - {0}))" by (intro subset_mset.cInf_greatest) auto also from False have "\ = prime_factorization (Gcd_factorial A)" by (rule prime_factorization_Gcd_factorial [symmetric]) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert nz False, auto simp: Gcd_factorial_eq_0_iff) qed (simp_all add: Gcd_factorial_def) lemma normalize_Lcm_factorial: "normalize (Lcm_factorial A) = Lcm_factorial A" by (simp add: Lcm_factorial_def) lemma Lcm_factorial_eq_0_iff: "Lcm_factorial A = 0 \ 0 \ A \ \subset_mset.bdd_above (prime_factorization ` A)" by (auto simp: Lcm_factorial_def in_Sup_multiset_iff) lemma dvd_Lcm_factorial: assumes "x \ A" shows "x dvd Lcm_factorial A" proof (cases "0 \ A \ subset_mset.bdd_above (prime_factorization ` A)") case True with assms have [simp]: "0 \ A" "x \ 0" "A \ {}" by auto from assms True have "prime_factorization x \# Sup (prime_factorization ` A)" by (intro subset_mset.cSup_upper) auto also have "\ = prime_factorization (Lcm_factorial A)" by (rule prime_factorization_Lcm_factorial [symmetric]) (insert True, simp_all) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert True, auto simp: Lcm_factorial_eq_0_iff) qed (insert assms, auto simp: Lcm_factorial_def) lemma Lcm_factorial_least: assumes "\y. y \ A \ y dvd x" shows "Lcm_factorial A dvd x" proof - consider "A = {}" | "0 \ A" | "x = 0" | "A \ {}" "0 \ A" "x \ 0" by blast thus ?thesis proof cases assume *: "A \ {}" "0 \ A" "x \ 0" hence nz: "x \ 0" if "x \ A" for x using that by auto from * have bdd: "subset_mset.bdd_above (prime_factorization ` A)" by (intro subset_mset.bdd_aboveI[of _ "prime_factorization x"]) (auto simp: prime_factorization_subset_iff_dvd nz dest: assms) have "prime_factorization (Lcm_factorial A) = Sup (prime_factorization ` A)" by (rule prime_factorization_Lcm_factorial) fact+ also from * have "\ \# prime_factorization x" by (intro subset_mset.cSup_least) (auto simp: prime_factorization_subset_iff_dvd nz dest: assms) finally show ?thesis by (subst (asm) prime_factorization_subset_iff_dvd) (insert * bdd, auto simp: Lcm_factorial_eq_0_iff) qed (auto simp: Lcm_factorial_def dest: assms) qed lemmas gcd_lcm_factorial = gcd_factorial_dvd1 gcd_factorial_dvd2 gcd_factorial_greatest normalize_gcd_factorial lcm_factorial_gcd_factorial normalize_Gcd_factorial Gcd_factorial_dvd Gcd_factorial_greatest normalize_Lcm_factorial dvd_Lcm_factorial Lcm_factorial_least end class factorial_semiring_gcd = factorial_semiring + gcd + Gcd + assumes gcd_eq_gcd_factorial: "gcd a b = gcd_factorial a b" and lcm_eq_lcm_factorial: "lcm a b = lcm_factorial a b" and Gcd_eq_Gcd_factorial: "Gcd A = Gcd_factorial A" and Lcm_eq_Lcm_factorial: "Lcm A = Lcm_factorial A" begin lemma prime_factorization_gcd: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (gcd a b) = prime_factorization a \# prime_factorization b" by (simp add: gcd_eq_gcd_factorial prime_factorization_gcd_factorial) lemma prime_factorization_lcm: assumes [simp]: "a \ 0" "b \ 0" shows "prime_factorization (lcm a b) = prime_factorization a \# prime_factorization b" by (simp add: lcm_eq_lcm_factorial prime_factorization_lcm_factorial) lemma prime_factorization_Gcd: assumes "Gcd A \ 0" shows "prime_factorization (Gcd A) = Inf (prime_factorization ` (A - {0}))" using assms by (simp add: prime_factorization_Gcd_factorial Gcd_eq_Gcd_factorial Gcd_factorial_eq_0_iff) lemma prime_factorization_Lcm: assumes "Lcm A \ 0" shows "prime_factorization (Lcm A) = Sup (prime_factorization ` A)" using assms by (simp add: prime_factorization_Lcm_factorial Lcm_eq_Lcm_factorial Lcm_factorial_eq_0_iff) lemma prime_factors_gcd [simp]: "a \ 0 \ b \ 0 \ prime_factors (gcd a b) = prime_factors a \ prime_factors b" by (subst prime_factorization_gcd) auto lemma prime_factors_lcm [simp]: "a \ 0 \ b \ 0 \ prime_factors (lcm a b) = prime_factors a \ prime_factors b" by (subst prime_factorization_lcm) auto subclass semiring_gcd by (standard, unfold gcd_eq_gcd_factorial lcm_eq_lcm_factorial) (rule gcd_lcm_factorial; assumption)+ subclass semiring_Gcd by (standard, unfold Gcd_eq_Gcd_factorial Lcm_eq_Lcm_factorial) (rule gcd_lcm_factorial; assumption)+ lemma assumes "x \ 0" "y \ 0" shows gcd_eq_factorial': "gcd x y = normalize (\p \ prime_factors x \ prime_factors y. p ^ min (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs1") and lcm_eq_factorial': "lcm x y = normalize (\p \ prime_factors x \ prime_factors y. p ^ max (multiplicity p x) (multiplicity p y))" (is "_ = ?rhs2") proof - have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial) also have "\ = ?rhs1" by (auto simp: gcd_factorial_def assms prod_mset_multiplicity count_prime_factorization_prime intro!: arg_cong[of _ _ normalize] dest: in_prime_factors_imp_prime intro!: prod.cong) finally show "gcd x y = ?rhs1" . have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial) also have "\ = ?rhs2" by (auto simp: lcm_factorial_def assms prod_mset_multiplicity count_prime_factorization_prime intro!: arg_cong[of _ _ normalize] dest: in_prime_factors_imp_prime intro!: prod.cong) finally show "lcm x y = ?rhs2" . qed lemma assumes "x \ 0" "y \ 0" "prime p" shows multiplicity_gcd: "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" and multiplicity_lcm: "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" proof - have "gcd x y = gcd_factorial x y" by (rule gcd_eq_gcd_factorial) also from assms have "multiplicity p \ = min (multiplicity p x) (multiplicity p y)" by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_gcd_factorial) finally show "multiplicity p (gcd x y) = min (multiplicity p x) (multiplicity p y)" . have "lcm x y = lcm_factorial x y" by (rule lcm_eq_lcm_factorial) also from assms have "multiplicity p \ = max (multiplicity p x) (multiplicity p y)" by (simp add: count_prime_factorization_prime [symmetric] prime_factorization_lcm_factorial) finally show "multiplicity p (lcm x y) = max (multiplicity p x) (multiplicity p y)" . qed lemma gcd_lcm_distrib: "gcd x (lcm y z) = lcm (gcd x y) (gcd x z)" proof (cases "x = 0 \ y = 0 \ z = 0") case True thus ?thesis by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd) next case False hence "normalize (gcd x (lcm y z)) = normalize (lcm (gcd x y) (gcd x z))" by (intro associatedI prime_factorization_subset_imp_dvd) (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm subset_mset.inf_sup_distrib1) thus ?thesis by simp qed lemma lcm_gcd_distrib: "lcm x (gcd y z) = gcd (lcm x y) (lcm x z)" proof (cases "x = 0 \ y = 0 \ z = 0") case True thus ?thesis by (auto simp: lcm_proj1_if_dvd lcm_proj2_if_dvd) next case False hence "normalize (lcm x (gcd y z)) = normalize (gcd (lcm x y) (lcm x z))" by (intro associatedI prime_factorization_subset_imp_dvd) (auto simp: lcm_eq_0_iff prime_factorization_gcd prime_factorization_lcm subset_mset.sup_inf_distrib1) thus ?thesis by simp qed end class factorial_ring_gcd = factorial_semiring_gcd + idom begin subclass ring_gcd .. subclass idom_divide .. end class factorial_semiring_multiplicative = factorial_semiring + normalization_semidom_multiplicative begin lemma normalize_prod_mset_primes: "(\p. p \# A \ prime p) \ normalize (prod_mset A) = prod_mset A" proof (induction A) case (add p A) hence "prime p" by simp hence "normalize p = p" by simp with add show ?case by (simp add: normalize_mult) qed simp_all lemma prod_mset_prime_factorization: assumes "x \ 0" shows "prod_mset (prime_factorization x) = normalize x" using assms by (induction x rule: prime_divisors_induct) (simp_all add: prime_factorization_unit prime_factorization_times_prime is_unit_normalize normalize_mult) lemma prime_decomposition: "unit_factor x * prod_mset (prime_factorization x) = x" by (cases "x = 0") (simp_all add: prod_mset_prime_factorization) lemma prod_prime_factors: assumes "x \ 0" shows "(\p \ prime_factors x. p ^ multiplicity p x) = normalize x" proof - have "normalize x = prod_mset (prime_factorization x)" by (simp add: prod_mset_prime_factorization assms) also have "\ = (\p \ prime_factors x. p ^ count (prime_factorization x) p)" by (subst prod_mset_multiplicity) simp_all also have "\ = (\p \ prime_factors x. p ^ multiplicity p x)" by (intro prod.cong) (simp_all add: assms count_prime_factorization_prime in_prime_factors_imp_prime) finally show ?thesis .. qed lemma prime_factorization_unique'': assumes S_eq: "S = {p. 0 < f p}" and "finite S" and S: "\p\S. prime p" "normalize n = (\p\S. p ^ f p)" shows "S = prime_factors n \ (\p. prime p \ f p = multiplicity p n)" proof define A where "A = Abs_multiset f" from \finite S\ S(1) have "(\p\S. p ^ f p) \ 0" by auto with S(2) have nz: "n \ 0" by auto from S_eq \finite S\ have count_A: "count A = f" unfolding A_def by (subst multiset.Abs_multiset_inverse) simp_all from S_eq count_A have set_mset_A: "set_mset A = S" by (simp only: set_mset_def) from S(2) have "normalize n = (\p\S. p ^ f p)" . also have "\ = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A) also from nz have "normalize n = prod_mset (prime_factorization n)" by (simp add: prod_mset_prime_factorization) finally have "prime_factorization (prod_mset A) = prime_factorization (prod_mset (prime_factorization n))" by simp also from S(1) have "prime_factorization (prod_mset A) = A" by (intro prime_factorization_prod_mset_primes) (auto simp: set_mset_A) also have "prime_factorization (prod_mset (prime_factorization n)) = prime_factorization n" by (intro prime_factorization_prod_mset_primes) auto finally show "S = prime_factors n" by (simp add: set_mset_A [symmetric]) show "(\p. prime p \ f p = multiplicity p n)" proof safe fix p :: 'a assume p: "prime p" have "multiplicity p n = multiplicity p (normalize n)" by simp also have "normalize n = prod_mset A" by (simp add: prod_mset_multiplicity S_eq set_mset_A count_A S) also from p set_mset_A S(1) have "multiplicity p \ = sum_mset (image_mset (multiplicity p) A)" by (intro prime_elem_multiplicity_prod_mset_distrib) auto also from S(1) p have "image_mset (multiplicity p) A = image_mset (\q. if p = q then 1 else 0) A" by (intro image_mset_cong) (auto simp: set_mset_A multiplicity_self prime_multiplicity_other) also have "sum_mset \ = f p" by (simp add: semiring_1_class.sum_mset_delta' count_A) finally show "f p = multiplicity p n" .. qed qed lemma divides_primepow: assumes "prime p" and "a dvd p ^ n" obtains m where "m \ n" and "normalize a = p ^ m" using divides_primepow_weak[OF assms] that assms by (auto simp add: normalize_power) lemma Ex_other_prime_factor: assumes "n \ 0" and "\(\k. normalize n = p ^ k)" "prime p" shows "\q\prime_factors n. q \ p" proof (rule ccontr) assume *: "\(\q\prime_factors n. q \ p)" have "normalize n = (\p\prime_factors n. p ^ multiplicity p n)" using assms(1) by (intro prod_prime_factors [symmetric]) auto also from * have "\ = (\p\{p}. p ^ multiplicity p n)" using assms(3) by (intro prod.mono_neutral_left) (auto simp: prime_factors_multiplicity) finally have "normalize n = p ^ multiplicity p n" by auto with assms show False by auto qed text \Now a string of results due to Jakub Kądziołka\ lemma multiplicity_dvd_iff_dvd: assumes "x \ 0" shows "p^k dvd x \ p^k dvd p^multiplicity p x" proof (cases "is_unit p") case True then have "is_unit (p^k)" using is_unit_power_iff by simp hence "p^k dvd x" by auto moreover from `is_unit p` have "p^k dvd p^multiplicity p x" using multiplicity_unit_left is_unit_power_iff by simp ultimately show ?thesis by simp next case False show ?thesis proof (cases "p = 0") case True then have "p^multiplicity p x = 1" by simp moreover have "p^k dvd x \ k = 0" proof (rule ccontr) assume "p^k dvd x" and "k \ 0" with `p = 0` have "p^k = 0" by auto with `p^k dvd x` have "0 dvd x" by auto hence "x = 0" by auto with `x \ 0` show False by auto qed ultimately show ?thesis by (auto simp add: is_unit_power_iff `\ is_unit p`) next case False with `x \ 0` `\ is_unit p` show ?thesis by (simp add: power_dvd_iff_le_multiplicity dvd_power_iff multiplicity_same_power) qed qed lemma multiplicity_decomposeI: assumes "x = p^k * x'" and "\ p dvd x'" and "p \ 0" shows "multiplicity p x = k" using assms local.multiplicity_eqI local.power_Suc2 by force lemma multiplicity_sum_lt: assumes "multiplicity p a < multiplicity p b" "a \ 0" "b \ 0" shows "multiplicity p (a + b) = multiplicity p a" proof - let ?vp = "multiplicity p" have unit: "\ is_unit p" proof assume "is_unit p" then have "?vp a = 0" and "?vp b = 0" using multiplicity_unit_left by auto with assms show False by auto qed from multiplicity_decompose' obtain a' where a': "a = p^?vp a * a'" "\ p dvd a'" using unit assms by metis from multiplicity_decompose' obtain b' where b': "b = p^?vp b * b'" using unit assms by metis show "?vp (a + b) = ?vp a" proof (rule multiplicity_decomposeI) let ?k = "?vp b - ?vp a" from assms have k: "?k > 0" by simp with b' have "b = p^?vp a * p^?k * b'" by (simp flip: power_add) with a' show *: "a + b = p^?vp a * (a' + p^?k * b')" by (simp add: ac_simps distrib_left) moreover show "\ p dvd a' + p^?k * b'" using a' k dvd_add_left_iff by auto show "p \ 0" using assms by auto qed qed corollary multiplicity_sum_min: assumes "multiplicity p a \ multiplicity p b" "a \ 0" "b \ 0" shows "multiplicity p (a + b) = min (multiplicity p a) (multiplicity p b)" proof - let ?vp = "multiplicity p" from assms have "?vp a < ?vp b \ ?vp a > ?vp b" by auto then show ?thesis by (metis assms multiplicity_sum_lt min.commute add_commute min.strict_order_iff) qed end end diff --git a/src/HOL/Computational_Algebra/Formal_Power_Series.thy b/src/HOL/Computational_Algebra/Formal_Power_Series.thy --- a/src/HOL/Computational_Algebra/Formal_Power_Series.thy +++ b/src/HOL/Computational_Algebra/Formal_Power_Series.thy @@ -1,6244 +1,6244 @@ (* Title: HOL/Computational_Algebra/Formal_Power_Series.thy Author: Amine Chaieb, University of Cambridge Author: Jeremy Sylvestre, University of Alberta (Augustana Campus) Author: Manuel Eberl, TU München *) section \A formalization of formal power series\ theory Formal_Power_Series imports Complex_Main Euclidean_Algorithm begin subsection \The type of formal power series\ typedef 'a fps = "{f :: nat \ 'a. True}" morphisms fps_nth Abs_fps by simp notation fps_nth (infixl "$" 75) lemma expand_fps_eq: "p = q \ (\n. p $ n = q $ n)" by (simp add: fps_nth_inject [symmetric] fun_eq_iff) lemmas fps_eq_iff = expand_fps_eq lemma fps_ext: "(\n. p $ n = q $ n) \ p = q" by (simp add: expand_fps_eq) lemma fps_nth_Abs_fps [simp]: "Abs_fps f $ n = f n" by (simp add: Abs_fps_inverse) text \Definition of the basic elements 0 and 1 and the basic operations of addition, negation and multiplication.\ instantiation fps :: (zero) zero begin definition fps_zero_def: "0 = Abs_fps (\n. 0)" instance .. end lemma fps_zero_nth [simp]: "0 $ n = 0" unfolding fps_zero_def by simp lemma fps_nonzero_nth: "f \ 0 \ (\ n. f $n \ 0)" by (simp add: expand_fps_eq) lemma fps_nonzero_nth_minimal: "f \ 0 \ (\n. f $ n \ 0 \ (\m < n. f $ m = 0))" (is "?lhs \ ?rhs") proof let ?n = "LEAST n. f $ n \ 0" show ?rhs if ?lhs proof - from that have "\n. f $ n \ 0" by (simp add: fps_nonzero_nth) then have "f $ ?n \ 0" by (rule LeastI_ex) moreover have "\m 0 \ (\m 0 \ f \ 0" by auto instantiation fps :: ("{one, zero}") one begin definition fps_one_def: "1 = Abs_fps (\n. if n = 0 then 1 else 0)" instance .. end lemma fps_one_nth [simp]: "1 $ n = (if n = 0 then 1 else 0)" unfolding fps_one_def by simp instantiation fps :: (plus) plus begin definition fps_plus_def: "(+) = (\f g. Abs_fps (\n. f $ n + g $ n))" instance .. end lemma fps_add_nth [simp]: "(f + g) $ n = f $ n + g $ n" unfolding fps_plus_def by simp instantiation fps :: (minus) minus begin definition fps_minus_def: "(-) = (\f g. Abs_fps (\n. f $ n - g $ n))" instance .. end lemma fps_sub_nth [simp]: "(f - g) $ n = f $ n - g $ n" unfolding fps_minus_def by simp instantiation fps :: (uminus) uminus begin definition fps_uminus_def: "uminus = (\f. Abs_fps (\n. - (f $ n)))" instance .. end lemma fps_neg_nth [simp]: "(- f) $ n = - (f $ n)" unfolding fps_uminus_def by simp lemma fps_neg_0 [simp]: "-(0::'a::group_add fps) = 0" by (rule iffD2, rule fps_eq_iff, auto) instantiation fps :: ("{comm_monoid_add, times}") times begin definition fps_times_def: "(*) = (\f g. Abs_fps (\n. \i=0..n. f $ i * g $ (n - i)))" instance .. end lemma fps_mult_nth: "(f * g) $ n = (\i=0..n. f$i * g$(n - i))" unfolding fps_times_def by simp lemma fps_mult_nth_0 [simp]: "(f * g) $ 0 = f $ 0 * g $ 0" unfolding fps_times_def by simp lemma fps_mult_nth_1: "(f * g) $ 1 = f$0 * g$1 + f$1 * g$0" by (simp add: fps_mult_nth) lemma fps_mult_nth_1' [simp]: "(f * g) $ Suc 0 = f$0 * g$Suc 0 + f$Suc 0 * g$0" by (simp add: fps_mult_nth) lemmas mult_nth_0 = fps_mult_nth_0 lemmas mult_nth_1 = fps_mult_nth_1 instance fps :: ("{comm_monoid_add, mult_zero}") mult_zero proof fix a :: "'a fps" show "0 * a = 0" by (simp add: fps_ext fps_mult_nth) show "a * 0 = 0" by (simp add: fps_ext fps_mult_nth) qed declare atLeastAtMost_iff [presburger] declare Bex_def [presburger] declare Ball_def [presburger] lemma mult_delta_left: fixes x y :: "'a::mult_zero" shows "(if b then x else 0) * y = (if b then x * y else 0)" by simp lemma mult_delta_right: fixes x y :: "'a::mult_zero" shows "x * (if b then y else 0) = (if b then x * y else 0)" by simp lemma fps_one_mult: fixes f :: "'a::{comm_monoid_add, mult_zero, monoid_mult} fps" shows "1 * f = f" and "f * 1 = f" by (simp_all add: fps_ext fps_mult_nth mult_delta_left mult_delta_right) subsection \Subdegrees\ definition subdegree :: "('a::zero) fps \ nat" where "subdegree f = (if f = 0 then 0 else LEAST n. f$n \ 0)" lemma subdegreeI: assumes "f $ d \ 0" and "\i. i < d \ f $ i = 0" shows "subdegree f = d" proof- from assms(1) have "f \ 0" by auto moreover from assms(1) have "(LEAST i. f $ i \ 0) = d" proof (rule Least_equality) fix e assume "f $ e \ 0" with assms(2) have "\(e < d)" by blast thus "e \ d" by simp qed ultimately show ?thesis unfolding subdegree_def by simp qed lemma nth_subdegree_nonzero [simp,intro]: "f \ 0 \ f $ subdegree f \ 0" proof- assume "f \ 0" hence "subdegree f = (LEAST n. f $ n \ 0)" by (simp add: subdegree_def) also from \f \ 0\ have "\n. f$n \ 0" using fps_nonzero_nth by blast from LeastI_ex[OF this] have "f $ (LEAST n. f $ n \ 0) \ 0" . finally show ?thesis . qed lemma nth_less_subdegree_zero [dest]: "n < subdegree f \ f $ n = 0" proof (cases "f = 0") assume "f \ 0" and less: "n < subdegree f" note less also from \f \ 0\ have "subdegree f = (LEAST n. f $ n \ 0)" by (simp add: subdegree_def) finally show "f $ n = 0" using not_less_Least by blast qed simp_all lemma subdegree_geI: assumes "f \ 0" "\i. i < n \ f$i = 0" shows "subdegree f \ n" proof (rule ccontr) assume "\(subdegree f \ n)" with assms(2) have "f $ subdegree f = 0" by simp moreover from assms(1) have "f $ subdegree f \ 0" by simp ultimately show False by contradiction qed lemma subdegree_greaterI: assumes "f \ 0" "\i. i \ n \ f$i = 0" shows "subdegree f > n" proof (rule ccontr) assume "\(subdegree f > n)" with assms(2) have "f $ subdegree f = 0" by simp moreover from assms(1) have "f $ subdegree f \ 0" by simp ultimately show False by contradiction qed lemma subdegree_leI: "f $ n \ 0 \ subdegree f \ n" by (rule leI) auto lemma subdegree_0 [simp]: "subdegree 0 = 0" by (simp add: subdegree_def) lemma subdegree_1 [simp]: "subdegree 1 = 0" by (cases "(1::'a) = 0") (auto intro: subdegreeI fps_ext simp: subdegree_def) lemma subdegree_eq_0_iff: "subdegree f = 0 \ f = 0 \ f $ 0 \ 0" proof (cases "f = 0") assume "f \ 0" thus ?thesis using nth_subdegree_nonzero[OF \f \ 0\] by (fastforce intro!: subdegreeI) qed simp_all lemma subdegree_eq_0 [simp]: "f $ 0 \ 0 \ subdegree f = 0" by (simp add: subdegree_eq_0_iff) lemma nth_subdegree_zero_iff [simp]: "f $ subdegree f = 0 \ f = 0" by (cases "f = 0") auto lemma fps_nonzero_subdegree_nonzeroI: "subdegree f > 0 \ f \ 0" by auto lemma subdegree_uminus [simp]: "subdegree (-(f::('a::group_add) fps)) = subdegree f" proof (cases "f=0") case False thus ?thesis by (force intro: subdegreeI) qed simp lemma subdegree_minus_commute [simp]: "subdegree (f-(g::('a::group_add) fps)) = subdegree (g - f)" proof (-, cases "g-f=0") case True have "\n. (f - g) $ n = -((g - f) $ n)" by simp with True have "f - g = 0" by (intro fps_ext) simp with True show ?thesis by simp next case False show ?thesis using nth_subdegree_nonzero[OF False] by (fastforce intro: subdegreeI) qed lemma subdegree_add_ge': fixes f g :: "'a::monoid_add fps" assumes "f + g \ 0" shows "subdegree (f + g) \ min (subdegree f) (subdegree g)" using assms by (force intro: subdegree_geI) lemma subdegree_add_ge: assumes "f \ -(g :: ('a :: group_add) fps)" shows "subdegree (f + g) \ min (subdegree f) (subdegree g)" proof (rule subdegree_add_ge') have "f + g = 0 \ False" proof- assume fg: "f + g = 0" have "\n. f $ n = - g $ n" proof- fix n from fg have "(f + g) $ n = 0" by simp hence "f $ n + g $ n - g $ n = - g $ n" by simp thus "f $ n = - g $ n" by simp qed with assms show False by (auto intro: fps_ext) qed thus "f + g \ 0" by fast qed lemma subdegree_add_eq1: assumes "f \ 0" and "subdegree f < subdegree (g :: 'a::monoid_add fps)" shows "subdegree (f + g) = subdegree f" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_add_eq2: assumes "g \ 0" and "subdegree g < subdegree (f :: 'a :: monoid_add fps)" shows "subdegree (f + g) = subdegree g" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_eq1: assumes "f \ 0" and "subdegree f < subdegree (g :: 'a :: group_add fps)" shows "subdegree (f - g) = subdegree f" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_eq1_cancel: assumes "f \ 0" and "subdegree f < subdegree (g :: 'a :: cancel_comm_monoid_add fps)" shows "subdegree (f - g) = subdegree f" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_eq2: assumes "g \ 0" and "subdegree g < subdegree (f :: 'a :: group_add fps)" shows "subdegree (f - g) = subdegree g" using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma subdegree_diff_ge [simp]: assumes "f \ (g :: 'a :: group_add fps)" shows "subdegree (f - g) \ min (subdegree f) (subdegree g)" proof- from assms have "f = - (- g) \ False" using expand_fps_eq by fastforce hence "f \ - (- g)" by fast moreover have "f + - g = f - g" by (simp add: fps_ext) ultimately show ?thesis using subdegree_add_ge[of f "-g"] by simp qed lemma subdegree_diff_ge': fixes f g :: "'a :: comm_monoid_diff fps" assumes "f - g \ 0" shows "subdegree (f - g) \ subdegree f" using assms by (auto intro: subdegree_geI simp: nth_less_subdegree_zero) lemma nth_subdegree_mult_left [simp]: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ (subdegree f) = f $ subdegree f * g $ 0" by (cases "subdegree f") (simp_all add: fps_mult_nth nth_less_subdegree_zero) lemma nth_subdegree_mult_right [simp]: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ (subdegree g) = f $ 0 * g $ subdegree g" by (cases "subdegree g") (simp_all add: fps_mult_nth nth_less_subdegree_zero sum.atLeast_Suc_atMost) lemma nth_subdegree_mult [simp]: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ (subdegree f + subdegree g) = f $ subdegree f * g $ subdegree g" proof- let ?n = "subdegree f + subdegree g" have "(f * g) $ ?n = (\i=0..?n. f$i * g$(?n-i))" by (simp add: fps_mult_nth) also have "... = (\i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)" proof (intro sum.cong) fix x assume x: "x \ {0..?n}" hence "x = subdegree f \ x < subdegree f \ ?n - x < subdegree g" by auto thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)" by (elim disjE conjE) auto qed auto also have "... = f $ subdegree f * g $ subdegree g" by simp finally show ?thesis . qed lemma fps_mult_nth_eq0: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "n < subdegree f + subdegree g" shows "(f*g) $ n = 0" proof- have "\i. i\{0..n} \ f$i * g$(n - i) = 0" proof- fix i assume i: "i\{0..n}" show "f$i * g$(n - i) = 0" proof (cases "i < subdegree f \ n - i < subdegree g") case False with assms i show ?thesis by auto qed (auto simp: nth_less_subdegree_zero) qed thus "(f * g) $ n = 0" by (simp add: fps_mult_nth) qed lemma fps_mult_subdegree_ge: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "f*g \ 0" shows "subdegree (f*g) \ subdegree f + subdegree g" using assms fps_mult_nth_eq0 by (intro subdegree_geI) simp lemma subdegree_mult': fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "f $ subdegree f * g $ subdegree g \ 0" shows "subdegree (f*g) = subdegree f + subdegree g" proof- from assms have "(f * g) $ (subdegree f + subdegree g) \ 0" by simp hence "f*g \ 0" by fastforce hence "subdegree (f*g) \ subdegree f + subdegree g" using fps_mult_subdegree_ge by fast moreover from assms have "subdegree (f*g) \ subdegree f + subdegree g" by (intro subdegree_leI) simp ultimately show ?thesis by simp qed lemma subdegree_mult [simp]: fixes f g :: "'a :: {semiring_no_zero_divisors} fps" assumes "f \ 0" "g \ 0" shows "subdegree (f * g) = subdegree f + subdegree g" using assms by (intro subdegree_mult') simp lemma fps_mult_nth_conv_upto_subdegree_left: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ n = (\i=subdegree f..n. f $ i * g $ (n - i))" proof (cases "subdegree f \ n") case True hence "{0..n} = {0.. {subdegree f..n}" by auto moreover have "{0.. {subdegree f..n} = {}" by auto ultimately show ?thesis using nth_less_subdegree_zero[of _ f] by (simp add: fps_mult_nth sum.union_disjoint) qed (simp add: fps_mult_nth nth_less_subdegree_zero) lemma fps_mult_nth_conv_upto_subdegree_right: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ n = (\i=0..n - subdegree g. f $ i * g $ (n - i))" proof- have "{0..n} = {0..n - subdegree g} \ {n - subdegree g<..n}" by auto moreover have "{0..n - subdegree g} \ {n - subdegree g<..n} = {}" by auto moreover have "\i\{n - subdegree g<..n}. g $ (n - i) = 0" using nth_less_subdegree_zero[of _ g] by auto ultimately show ?thesis by (simp add: fps_mult_nth sum.union_disjoint) qed lemma fps_mult_nth_conv_inside_subdegrees: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "(f * g) $ n = (\i=subdegree f..n - subdegree g. f $ i * g $ (n - i))" proof (cases "subdegree f \ n - subdegree g") case True hence "{subdegree f..n} = {subdegree f..n - subdegree g} \ {n - subdegree g<..n}" by auto moreover have "{subdegree f..n - subdegree g} \ {n - subdegree g<..n} = {}" by auto moreover have "\i\{n - subdegree g<..n}. f $ i * g $ (n - i) = 0" using nth_less_subdegree_zero[of _ g] by auto ultimately show ?thesis using fps_mult_nth_conv_upto_subdegree_left[of f g n] by (simp add: sum.union_disjoint) next case False hence 1: "subdegree f > n - subdegree g" by simp show ?thesis proof (cases "f*g = 0") case False with 1 have "n < subdegree (f*g)" using fps_mult_subdegree_ge[of f g] by simp with 1 show ?thesis by auto qed (simp add: 1) qed lemma fps_mult_nth_outside_subdegrees: fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps" shows "n < subdegree f \ (f * g) $ n = 0" and "n < subdegree g \ (f * g) $ n = 0" by (auto simp: fps_mult_nth_conv_inside_subdegrees) subsection \Formal power series form a commutative ring with unity, if the range of sequences they represent is a commutative ring with unity\ instance fps :: (semigroup_add) semigroup_add proof fix a b c :: "'a fps" show "a + b + c = a + (b + c)" by (simp add: fps_ext add.assoc) qed instance fps :: (ab_semigroup_add) ab_semigroup_add proof fix a b :: "'a fps" show "a + b = b + a" by (simp add: fps_ext add.commute) qed instance fps :: (monoid_add) monoid_add proof fix a :: "'a fps" show "0 + a = a" by (simp add: fps_ext) show "a + 0 = a" by (simp add: fps_ext) qed instance fps :: (comm_monoid_add) comm_monoid_add proof fix a :: "'a fps" show "0 + a = a" by (simp add: fps_ext) qed instance fps :: (cancel_semigroup_add) cancel_semigroup_add proof fix a b c :: "'a fps" show "b = c" if "a + b = a + c" using that by (simp add: expand_fps_eq) show "b = c" if "b + a = c + a" using that by (simp add: expand_fps_eq) qed instance fps :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add proof fix a b c :: "'a fps" show "a + b - a = b" by (simp add: expand_fps_eq) show "a - b - c = a - (b + c)" by (simp add: expand_fps_eq diff_diff_eq) qed instance fps :: (cancel_comm_monoid_add) cancel_comm_monoid_add .. instance fps :: (group_add) group_add proof fix a b :: "'a fps" show "- a + a = 0" by (simp add: fps_ext) show "a + - b = a - b" by (simp add: fps_ext) qed instance fps :: (ab_group_add) ab_group_add proof fix a b :: "'a fps" show "- a + a = 0" by (simp add: fps_ext) show "a - b = a + - b" by (simp add: fps_ext) qed instance fps :: (zero_neq_one) zero_neq_one by standard (simp add: expand_fps_eq) lemma fps_mult_assoc_lemma: fixes k :: nat and f :: "nat \ nat \ nat \ 'a::comm_monoid_add" shows "(\j=0..k. \i=0..j. f i (j - i) (n - j)) = (\j=0..k. \i=0..k - j. f j i (n - j - i))" by (induct k) (simp_all add: Suc_diff_le sum.distrib add.assoc) instance fps :: (semiring_0) semiring_0 proof fix a b c :: "'a fps" show "(a + b) * c = a * c + b * c" by (simp add: expand_fps_eq fps_mult_nth distrib_right sum.distrib) show "a * (b + c) = a * b + a * c" by (simp add: expand_fps_eq fps_mult_nth distrib_left sum.distrib) show "(a * b) * c = a * (b * c)" proof (rule fps_ext) fix n :: nat have "(\j=0..n. \i=0..j. a$i * b$(j - i) * c$(n - j)) = (\j=0..n. \i=0..n - j. a$j * b$i * c$(n - j - i))" by (rule fps_mult_assoc_lemma) then show "((a * b) * c) $ n = (a * (b * c)) $ n" by (simp add: fps_mult_nth sum_distrib_left sum_distrib_right mult.assoc) qed qed instance fps :: (semiring_0_cancel) semiring_0_cancel .. lemma fps_mult_commute_lemma: fixes n :: nat and f :: "nat \ nat \ 'a::comm_monoid_add" shows "(\i=0..n. f i (n - i)) = (\i=0..n. f (n - i) i)" by (rule sum.reindex_bij_witness[where i="(-) n" and j="(-) n"]) auto instance fps :: (comm_semiring_0) comm_semiring_0 proof fix a b c :: "'a fps" show "a * b = b * a" proof (rule fps_ext) fix n :: nat have "(\i=0..n. a$i * b$(n - i)) = (\i=0..n. a$(n - i) * b$i)" by (rule fps_mult_commute_lemma) then show "(a * b) $ n = (b * a) $ n" by (simp add: fps_mult_nth mult.commute) qed qed (simp add: distrib_right) instance fps :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. instance fps :: (semiring_1) semiring_1 proof fix a :: "'a fps" show "1 * a = a" "a * 1 = a" by (simp_all add: fps_one_mult) qed instance fps :: (comm_semiring_1) comm_semiring_1 by standard simp instance fps :: (semiring_1_cancel) semiring_1_cancel .. subsection \Selection of the nth power of the implicit variable in the infinite sum\ lemma fps_square_nth: "(f^2) $ n = (\k\n. f $ k * f $ (n - k))" by (simp add: power2_eq_square fps_mult_nth atLeast0AtMost) lemma fps_sum_nth: "sum f S $ n = sum (\k. (f k) $ n) S" proof (cases "finite S") case True then show ?thesis by (induct set: finite) auto next case False then show ?thesis by simp qed subsection \Injection of the basic ring elements and multiplication by scalars\ definition "fps_const c = Abs_fps (\n. if n = 0 then c else 0)" lemma fps_nth_fps_const [simp]: "fps_const c $ n = (if n = 0 then c else 0)" unfolding fps_const_def by simp lemma fps_const_0_eq_0 [simp]: "fps_const 0 = 0" by (simp add: fps_ext) lemma fps_const_nonzero_eq_nonzero: "c \ 0 \ fps_const c \ 0" using fps_nonzeroI[of "fps_const c" 0] by simp lemma fps_const_1_eq_1 [simp]: "fps_const 1 = 1" by (simp add: fps_ext) lemma subdegree_fps_const [simp]: "subdegree (fps_const c) = 0" by (cases "c = 0") (auto intro!: subdegreeI) lemma fps_const_neg [simp]: "- (fps_const (c::'a::group_add)) = fps_const (- c)" by (simp add: fps_ext) lemma fps_const_add [simp]: "fps_const (c::'a::monoid_add) + fps_const d = fps_const (c + d)" by (simp add: fps_ext) lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f = Abs_fps (\n. if n = 0 then c + f$0 else f$n)" by (simp add: fps_ext) lemma fps_const_add_right: "f + fps_const (c::'a::monoid_add) = Abs_fps (\n. if n = 0 then f$0 + c else f$n)" by (simp add: fps_ext) lemma fps_const_sub [simp]: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)" by (simp add: fps_ext) lemmas fps_const_minus = fps_const_sub lemma fps_const_mult[simp]: fixes c d :: "'a::{comm_monoid_add,mult_zero}" shows "fps_const c * fps_const d = fps_const (c * d)" by (simp add: fps_eq_iff fps_mult_nth sum.neutral) lemma fps_const_mult_left: "fps_const (c::'a::{comm_monoid_add,mult_zero}) * f = Abs_fps (\n. c * f$n)" unfolding fps_eq_iff fps_mult_nth by (simp add: fps_const_def mult_delta_left) lemma fps_const_mult_right: "f * fps_const (c::'a::{comm_monoid_add,mult_zero}) = Abs_fps (\n. f$n * c)" unfolding fps_eq_iff fps_mult_nth by (simp add: fps_const_def mult_delta_right) lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::{comm_monoid_add,mult_zero}) * f)$n = c* f$n" by (simp add: fps_mult_nth mult_delta_left) lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::{comm_monoid_add,mult_zero}))$n = f$n * c" by (simp add: fps_mult_nth mult_delta_right) lemma fps_const_power [simp]: "fps_const c ^ n = fps_const (c^n)" by (induct n) auto subsection \Formal power series form an integral domain\ instance fps :: (ring) ring .. instance fps :: (comm_ring) comm_ring .. instance fps :: (ring_1) ring_1 .. instance fps :: (comm_ring_1) comm_ring_1 .. instance fps :: (semiring_no_zero_divisors) semiring_no_zero_divisors proof fix a b :: "'a fps" assume "a \ 0" and "b \ 0" hence "(a * b) $ (subdegree a + subdegree b) \ 0" by simp thus "a * b \ 0" using fps_nonzero_nth by fast qed instance fps :: (semiring_1_no_zero_divisors) semiring_1_no_zero_divisors .. instance fps :: ("{cancel_semigroup_add,semiring_no_zero_divisors_cancel}") semiring_no_zero_divisors_cancel proof fix a b c :: "'a fps" show "(a * c = b * c) = (c = 0 \ a = b)" proof assume ab: "a * c = b * c" have "c \ 0 \ a = b" proof (rule fps_ext) fix n assume c: "c \ 0" show "a $ n = b $ n" proof (induct n rule: nat_less_induct) case (1 n) with ab c show ?case using fps_mult_nth_conv_upto_subdegree_right[of a c "subdegree c + n"] fps_mult_nth_conv_upto_subdegree_right[of b c "subdegree c + n"] by (cases n) auto qed qed thus "c = 0 \ a = b" by fast qed auto show "(c * a = c * b) = (c = 0 \ a = b)" proof assume ab: "c * a = c * b" have "c \ 0 \ a = b" proof (rule fps_ext) fix n assume c: "c \ 0" show "a $ n = b $ n" proof (induct n rule: nat_less_induct) case (1 n) moreover have "\i\{Suc (subdegree c)..subdegree c + n}. subdegree c + n - i < n" by auto ultimately show ?case using ab c fps_mult_nth_conv_upto_subdegree_left[of c a "subdegree c + n"] fps_mult_nth_conv_upto_subdegree_left[of c b "subdegree c + n"] by (simp add: sum.atLeast_Suc_atMost) qed qed thus "c = 0 \ a = b" by fast qed auto qed instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors .. instance fps :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors .. instance fps :: (idom) idom .. lemma fps_numeral_fps_const: "numeral k = fps_const (numeral k)" by (induct k) (simp_all only: numeral.simps fps_const_1_eq_1 fps_const_add [symmetric]) lemmas numeral_fps_const = fps_numeral_fps_const lemma neg_numeral_fps_const: "(- numeral k :: 'a :: ring_1 fps) = fps_const (- numeral k)" by (simp add: numeral_fps_const) lemma fps_numeral_nth: "numeral n $ i = (if i = 0 then numeral n else 0)" by (simp add: numeral_fps_const) lemma fps_numeral_nth_0 [simp]: "numeral n $ 0 = numeral n" by (simp add: numeral_fps_const) lemma subdegree_numeral [simp]: "subdegree (numeral n) = 0" by (simp add: numeral_fps_const) lemma fps_of_nat: "fps_const (of_nat c) = of_nat c" by (induction c) (simp_all add: fps_const_add [symmetric] del: fps_const_add) lemma fps_of_int: "fps_const (of_int c) = of_int c" by (induction c) (simp_all add: fps_const_minus [symmetric] fps_of_nat fps_const_neg [symmetric] del: fps_const_minus fps_const_neg) lemma fps_nth_of_nat [simp]: "(of_nat c) $ n = (if n=0 then of_nat c else 0)" by (simp add: fps_of_nat[symmetric]) lemma fps_nth_of_int [simp]: "(of_int c) $ n = (if n=0 then of_int c else 0)" by (simp add: fps_of_int[symmetric]) lemma fps_mult_of_nat_nth [simp]: shows "(of_nat k * f) $ n = of_nat k * f$n" and "(f * of_nat k ) $ n = f$n * of_nat k" by (simp_all add: fps_of_nat[symmetric]) lemma fps_mult_of_int_nth [simp]: shows "(of_int k * f) $ n = of_int k * f$n" and "(f * of_int k ) $ n = f$n * of_int k" by (simp_all add: fps_of_int[symmetric]) lemma numeral_neq_fps_zero [simp]: "(numeral f :: 'a :: field_char_0 fps) \ 0" proof assume "numeral f = (0 :: 'a fps)" from arg_cong[of _ _ "\F. F $ 0", OF this] show False by simp qed lemma subdegree_power_ge: "f^n \ 0 \ subdegree (f^n) \ n * subdegree f" proof (induct n) case (Suc n) thus ?case using fps_mult_subdegree_ge by fastforce qed simp lemma fps_pow_nth_below_subdegree: "k < n * subdegree f \ (f^n) $ k = 0" proof (cases "f^n = 0") case False assume "k < n * subdegree f" with False have "k < subdegree (f^n)" using subdegree_power_ge[of f n] by simp thus "(f^n) $ k = 0" by auto qed simp lemma fps_pow_base [simp]: "(f ^ n) $ (n * subdegree f) = (f $ subdegree f) ^ n" proof (induct n) case (Suc n) show ?case proof (cases "Suc n * subdegree f < subdegree f + subdegree (f^n)") case True with Suc show ?thesis by (auto simp: fps_mult_nth_eq0 distrib_right) next case False hence "\i\{Suc (subdegree f)..Suc n * subdegree f - subdegree (f ^ n)}. f ^ n $ (Suc n * subdegree f - i) = 0" by (auto simp: fps_pow_nth_below_subdegree) with False Suc show ?thesis using fps_mult_nth_conv_inside_subdegrees[of f "f^n" "Suc n * subdegree f"] sum.atLeast_Suc_atMost[of "subdegree f" "Suc n * subdegree f - subdegree (f ^ n)" "\i. f $ i * f ^ n $ (Suc n * subdegree f - i)" ] by simp qed qed simp lemma subdegree_power_eqI: fixes f :: "'a::semiring_1 fps" shows "(f $ subdegree f) ^ n \ 0 \ subdegree (f ^ n) = n * subdegree f" proof (induct n) case (Suc n) from Suc have 1: "subdegree (f ^ n) = n * subdegree f" by fastforce with Suc(2) have "f $ subdegree f * f ^ n $ subdegree (f ^ n) \ 0" by simp with 1 show ?case using subdegree_mult'[of f "f^n"] by simp qed simp lemma subdegree_power [simp]: "subdegree ((f :: ('a :: semiring_1_no_zero_divisors) fps) ^ n) = n * subdegree f" by (cases "f = 0"; induction n) simp_all subsection \The efps_Xtractor series fps_X\ lemma minus_one_power_iff: "(- (1::'a::ring_1)) ^ n = (if even n then 1 else - 1)" by (induct n) auto definition "fps_X = Abs_fps (\n. if n = 1 then 1 else 0)" lemma subdegree_fps_X [simp]: "subdegree (fps_X :: ('a :: zero_neq_one) fps) = 1" by (auto intro!: subdegreeI simp: fps_X_def) lemma fps_X_mult_nth [simp]: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "(fps_X * f) $ n = (if n = 0 then 0 else f $ (n - 1))" proof (cases n) case (Suc m) moreover have "(fps_X * f) $ Suc m = f $ (Suc m - 1)" proof (cases m) case 0 thus ?thesis using fps_mult_nth_1[of "fps_X" f] by (simp add: fps_X_def) next case (Suc k) thus ?thesis by (simp add: fps_mult_nth fps_X_def sum.atLeast_Suc_atMost) qed ultimately show ?thesis by simp qed (simp add: fps_X_def) lemma fps_X_mult_right_nth [simp]: fixes a :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "(a * fps_X) $ n = (if n = 0 then 0 else a $ (n - 1))" proof (cases n) case (Suc m) moreover have "(a * fps_X) $ Suc m = a $ (Suc m - 1)" proof (cases m) case 0 thus ?thesis using fps_mult_nth_1[of a "fps_X"] by (simp add: fps_X_def) next case (Suc k) hence "(a * fps_X) $ Suc m = (\i=0..k. a$i * fps_X$(Suc m - i)) + a$(Suc k)" by (simp add: fps_mult_nth fps_X_def) moreover have "\i\{0..k}. a$i * fps_X$(Suc m - i) = 0" by (auto simp: Suc fps_X_def) ultimately show ?thesis by (simp add: Suc) qed ultimately show ?thesis by simp qed (simp add: fps_X_def) lemma fps_mult_fps_X_commute: fixes a :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "fps_X * a = a * fps_X" by (simp add: fps_eq_iff) lemma fps_mult_fps_X_power_commute: "fps_X ^ k * a = a * fps_X ^ k" proof (induct k) case (Suc k) hence "fps_X ^ Suc k * a = a * fps_X * fps_X ^ k" by (simp add: mult.assoc fps_mult_fps_X_commute[symmetric]) thus ?case by (simp add: mult.assoc) qed simp lemma fps_subdegree_mult_fps_X: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" assumes "f \ 0" shows "subdegree (fps_X * f) = subdegree f + 1" and "subdegree (f * fps_X) = subdegree f + 1" proof- show "subdegree (fps_X * f) = subdegree f + 1" proof (intro subdegreeI) fix i :: nat assume i: "i < subdegree f + 1" show "(fps_X * f) $ i = 0" proof (cases "i=0") case False with i show ?thesis by (simp add: nth_less_subdegree_zero) next case True thus ?thesis using fps_X_mult_nth[of f i] by simp qed qed (simp add: assms) thus "subdegree (f * fps_X) = subdegree f + 1" by (simp add: fps_mult_fps_X_commute) qed lemma fps_mult_fps_X_nonzero: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" assumes "f \ 0" shows "fps_X * f \ 0" and "f * fps_X \ 0" using assms fps_subdegree_mult_fps_X[of f] fps_nonzero_subdegree_nonzeroI[of "fps_X * f"] fps_nonzero_subdegree_nonzeroI[of "f * fps_X"] by auto lemma fps_mult_fps_X_power_nonzero: assumes "f \ 0" shows "fps_X ^ n * f \ 0" and "f * fps_X ^ n \ 0" proof - show "fps_X ^ n * f \ 0" by (induct n) (simp_all add: assms mult.assoc fps_mult_fps_X_nonzero(1)) thus "f * fps_X ^ n \ 0" by (simp add: fps_mult_fps_X_power_commute) qed lemma fps_X_power_iff: "fps_X ^ n = Abs_fps (\m. if m = n then 1 else 0)" by (induction n) (auto simp: fps_eq_iff) lemma fps_X_nth[simp]: "fps_X$n = (if n = 1 then 1 else 0)" by (simp add: fps_X_def) lemma fps_X_power_nth[simp]: "(fps_X^k) $n = (if n = k then 1 else 0)" by (simp add: fps_X_power_iff) lemma fps_X_power_subdegree: "subdegree (fps_X^n) = n" by (auto intro: subdegreeI) lemma fps_X_power_mult_nth: "(fps_X^k * f) $ n = (if n < k then 0 else f $ (n - k))" by (cases "n 0" shows "subdegree (fps_X ^ n * f) = subdegree f + n" and "subdegree (f * fps_X ^ n) = subdegree f + n" proof - from assms show "subdegree (fps_X ^ n * f) = subdegree f + n" by (induct n) (simp_all add: algebra_simps fps_subdegree_mult_fps_X(1) fps_mult_fps_X_power_nonzero(1)) thus "subdegree (f * fps_X ^ n) = subdegree f + n" by (simp add: fps_mult_fps_X_power_commute) qed lemma fps_mult_fps_X_plus_1_nth: "((1+fps_X)*a) $n = (if n = 0 then (a$n :: 'a::semiring_1) else a$n + a$(n - 1))" proof (cases n) case 0 then show ?thesis by (simp add: fps_mult_nth) next case (Suc m) have "((1 + fps_X)*a) $ n = sum (\i. (1 + fps_X) $ i * a $ (n - i)) {0..n}" by (simp add: fps_mult_nth) also have "\ = sum (\i. (1+fps_X)$i * a$(n-i)) {0.. 1}" unfolding Suc by (rule sum.mono_neutral_right) auto also have "\ = (if n = 0 then a$n else a$n + a$(n - 1))" by (simp add: Suc) finally show ?thesis . qed lemma fps_mult_right_fps_X_plus_1_nth: fixes a :: "'a :: semiring_1 fps" shows "(a*(1+fps_X)) $ n = (if n = 0 then a$n else a$n + a$(n - 1))" using fps_mult_fps_X_plus_1_nth by (simp add: distrib_left fps_mult_fps_X_commute distrib_right) lemma fps_X_neq_fps_const [simp]: "(fps_X :: 'a :: zero_neq_one fps) \ fps_const c" proof assume "(fps_X::'a fps) = fps_const (c::'a)" hence "fps_X$1 = (fps_const (c::'a))$1" by (simp only:) thus False by auto qed lemma fps_X_neq_zero [simp]: "(fps_X :: 'a :: zero_neq_one fps) \ 0" by (simp only: fps_const_0_eq_0[symmetric] fps_X_neq_fps_const) simp lemma fps_X_neq_one [simp]: "(fps_X :: 'a :: zero_neq_one fps) \ 1" by (simp only: fps_const_1_eq_1[symmetric] fps_X_neq_fps_const) simp lemma fps_X_neq_numeral [simp]: "fps_X \ numeral c" by (simp only: numeral_fps_const fps_X_neq_fps_const) simp lemma fps_X_pow_eq_fps_X_pow_iff [simp]: "fps_X ^ m = fps_X ^ n \ m = n" proof assume "(fps_X :: 'a fps) ^ m = fps_X ^ n" hence "(fps_X :: 'a fps) ^ m $ m = fps_X ^ n $ m" by (simp only:) thus "m = n" by (simp split: if_split_asm) qed simp_all subsection \Shifting and slicing\ definition fps_shift :: "nat \ 'a fps \ 'a fps" where "fps_shift n f = Abs_fps (\i. f $ (i + n))" lemma fps_shift_nth [simp]: "fps_shift n f $ i = f $ (i + n)" by (simp add: fps_shift_def) lemma fps_shift_0 [simp]: "fps_shift 0 f = f" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_zero [simp]: "fps_shift n 0 = 0" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_one: "fps_shift n 1 = (if n = 0 then 1 else 0)" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_fps_const: "fps_shift n (fps_const c) = (if n = 0 then fps_const c else 0)" by (intro fps_ext) (simp add: fps_shift_def) lemma fps_shift_numeral: "fps_shift n (numeral c) = (if n = 0 then numeral c else 0)" by (simp add: numeral_fps_const fps_shift_fps_const) lemma fps_shift_fps_X [simp]: "n \ 1 \ fps_shift n fps_X = (if n = 1 then 1 else 0)" by (intro fps_ext) (auto simp: fps_X_def) lemma fps_shift_fps_X_power [simp]: "n \ m \ fps_shift n (fps_X ^ m) = fps_X ^ (m - n)" by (intro fps_ext) auto lemma fps_shift_subdegree [simp]: "n \ subdegree f \ subdegree (fps_shift n f) = subdegree f - n" by (cases "f=0") (auto intro: subdegreeI simp: nth_less_subdegree_zero) lemma fps_shift_fps_shift: "fps_shift (m + n) f = fps_shift m (fps_shift n f)" by (rule fps_ext) (simp add: add_ac) lemma fps_shift_fps_shift_reorder: "fps_shift m (fps_shift n f) = fps_shift n (fps_shift m f)" using fps_shift_fps_shift[of m n f] fps_shift_fps_shift[of n m f] by (simp add: add.commute) lemma fps_shift_rev_shift: "m \ n \ fps_shift n (Abs_fps (\k. if k n \ fps_shift n (Abs_fps (\k. if kk. if k n" thus "fps_shift n (Abs_fps (\k. if k n" hence "\k. k \ m-n \ k+n-m = k - (m-n)" by auto thus "fps_shift n (Abs_fps (\k. if kk. if k subdegree (g :: 'b :: {comm_monoid_add, mult_zero} fps)" shows "fps_shift n (h*g) = h * fps_shift n g" proof- have case1: "\a b::'b fps. 1 \ subdegree b \ fps_shift 1 (a*b) = a * fps_shift 1 b" proof (rule fps_ext) fix a b :: "'b fps" and n :: nat assume b: "1 \ subdegree b" have "\i. i \ n \ n + 1 - i = (n-i) + 1" by (simp add: algebra_simps) with b show "fps_shift 1 (a*b) $ n = (a * fps_shift 1 b) $ n" by (simp add: fps_mult_nth nth_less_subdegree_zero) qed have "n \ subdegree g \ fps_shift n (h*g) = h * fps_shift n g" proof (induct n) case (Suc n) have "fps_shift (Suc n) (h*g) = fps_shift 1 (fps_shift n (h*g))" by (simp add: fps_shift_fps_shift[symmetric]) also have "\ = h * (fps_shift 1 (fps_shift n g))" using Suc case1 by force finally show ?case by (simp add: fps_shift_fps_shift[symmetric]) qed simp with assms show ?thesis by fast qed lemma fps_shift_mult_right_noncomm: assumes "n \ subdegree (g :: 'b :: {comm_monoid_add, mult_zero} fps)" shows "fps_shift n (g*h) = fps_shift n g * h" proof- have case1: "\a b::'b fps. 1 \ subdegree a \ fps_shift 1 (a*b) = fps_shift 1 a * b" proof (rule fps_ext) fix a b :: "'b fps" and n :: nat assume "1 \ subdegree a" hence "fps_shift 1 (a*b) $ n = (\i=Suc 0..Suc n. a$i * b$(n+1-i))" using sum.atLeast_Suc_atMost[of 0 "n+1" "\i. a$i * b$(n+1-i)"] by (simp add: fps_mult_nth nth_less_subdegree_zero) thus "fps_shift 1 (a*b) $ n = (fps_shift 1 a * b) $ n" using sum.shift_bounds_cl_Suc_ivl[of "\i. a$i * b$(n+1-i)" 0 n] by (simp add: fps_mult_nth) qed have "n \ subdegree g \ fps_shift n (g*h) = fps_shift n g * h" proof (induct n) case (Suc n) have "fps_shift (Suc n) (g*h) = fps_shift 1 (fps_shift n (g*h))" by (simp add: fps_shift_fps_shift[symmetric]) also have "\ = (fps_shift 1 (fps_shift n g)) * h" using Suc case1 by force finally show ?case by (simp add: fps_shift_fps_shift[symmetric]) qed simp with assms show ?thesis by fast qed lemma fps_shift_mult_right: assumes "n \ subdegree (g :: 'b :: comm_semiring_0 fps)" shows "fps_shift n (g*h) = h * fps_shift n g" by (simp add: assms fps_shift_mult_right_noncomm mult.commute) lemma fps_shift_mult_both: fixes f g :: "'a::{comm_monoid_add, mult_zero} fps" assumes "m \ subdegree f" "n \ subdegree g" shows "fps_shift m f * fps_shift n g = fps_shift (m+n) (f*g)" using assms by (simp add: fps_shift_mult fps_shift_mult_right_noncomm fps_shift_fps_shift) lemma fps_shift_subdegree_zero_iff [simp]: "fps_shift (subdegree f) f = 0 \ f = 0" by (subst (1) nth_subdegree_zero_iff[symmetric], cases "f = 0") (simp_all del: nth_subdegree_zero_iff) lemma fps_shift_times_fps_X: fixes f g :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "1 \ subdegree f \ fps_shift 1 f * fps_X = f" by (intro fps_ext) (simp add: nth_less_subdegree_zero) lemma fps_shift_times_fps_X' [simp]: fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "fps_shift 1 (f * fps_X) = f" by (intro fps_ext) (simp add: nth_less_subdegree_zero) lemma fps_shift_times_fps_X'': fixes f :: "'a::{comm_monoid_add,mult_zero,monoid_mult} fps" shows "1 \ n \ fps_shift n (f * fps_X) = fps_shift (n - 1) f" by (intro fps_ext) (simp add: nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power: "n \ subdegree f \ fps_shift n f * fps_X ^ n = f" by (intro fps_ext) (simp add: fps_X_power_mult_right_nth nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power' [simp]: "fps_shift n (f * fps_X^n) = f" by (intro fps_ext) (simp add: fps_X_power_mult_right_nth nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power'': "m \ n \ fps_shift n (f * fps_X^m) = fps_shift (n - m) f" by (intro fps_ext) (simp add: fps_X_power_mult_right_nth nth_less_subdegree_zero) lemma fps_shift_times_fps_X_power''': "m > n \ fps_shift n (f * fps_X^m) = f * fps_X^(m - n)" proof (cases "f=0") case False assume m: "m>n" hence "m = n + (m-n)" by auto with False m show ?thesis using power_add[of "fps_X::'a fps" n "m-n"] fps_shift_mult_right_noncomm[of n "f * fps_X^n" "fps_X^(m-n)"] by (simp add: mult.assoc fps_subdegree_mult_fps_X_power(2)) qed simp lemma subdegree_decompose: "f = fps_shift (subdegree f) f * fps_X ^ subdegree f" by (rule fps_ext) (auto simp: fps_X_power_mult_right_nth) lemma subdegree_decompose': "n \ subdegree f \ f = fps_shift n f * fps_X^n" by (rule fps_ext) (auto simp: fps_X_power_mult_right_nth intro!: nth_less_subdegree_zero) instantiation fps :: (zero) unit_factor begin definition fps_unit_factor_def [simp]: "unit_factor f = fps_shift (subdegree f) f" instance .. end lemma fps_unit_factor_zero_iff: "unit_factor (f::'a::zero fps) = 0 \ f = 0" by simp lemma fps_unit_factor_nth_0: "f \ 0 \ unit_factor f $ 0 \ 0" by simp lemma fps_X_unit_factor: "unit_factor (fps_X :: 'a :: zero_neq_one fps) = 1" by (intro fps_ext) auto lemma fps_X_power_unit_factor: "unit_factor (fps_X ^ n) = 1" proof- define X :: "'a fps" where "X \ fps_X" hence "unit_factor (X^n) = fps_shift n (X^n)" by (simp add: fps_X_power_subdegree) moreover have "fps_shift n (X^n) = 1" by (auto intro: fps_ext simp: fps_X_power_iff X_def) ultimately show ?thesis by (simp add: X_def) qed lemma fps_unit_factor_decompose: "f = unit_factor f * fps_X ^ subdegree f" by (simp add: subdegree_decompose) lemma fps_unit_factor_decompose': "f = fps_X ^ subdegree f * unit_factor f" using fps_unit_factor_decompose by (simp add: fps_mult_fps_X_power_commute) lemma fps_unit_factor_uminus: "unit_factor (-f) = - unit_factor (f::'a::group_add fps)" by (simp add: fps_shift_uminus) lemma fps_unit_factor_shift: assumes "n \ subdegree f" shows "unit_factor (fps_shift n f) = unit_factor f" by (simp add: assms fps_shift_fps_shift[symmetric]) lemma fps_unit_factor_mult_fps_X: fixes f :: "'a::{comm_monoid_add,monoid_mult,mult_zero} fps" shows "unit_factor (fps_X * f) = unit_factor f" and "unit_factor (f * fps_X) = unit_factor f" proof - show "unit_factor (fps_X * f) = unit_factor f" by (cases "f=0") (auto intro: fps_ext simp: fps_subdegree_mult_fps_X(1)) thus "unit_factor (f * fps_X) = unit_factor f" by (simp add: fps_mult_fps_X_commute) qed lemma fps_unit_factor_mult_fps_X_power: shows "unit_factor (fps_X ^ n * f) = unit_factor f" and "unit_factor (f * fps_X ^ n) = unit_factor f" proof - show "unit_factor (fps_X ^ n * f) = unit_factor f" proof (induct n) case (Suc m) thus ?case using fps_unit_factor_mult_fps_X(1)[of "fps_X ^ m * f"] by (simp add: mult.assoc) qed simp thus "unit_factor (f * fps_X ^ n) = unit_factor f" by (simp add: fps_mult_fps_X_power_commute) qed lemma fps_unit_factor_mult_unit_factor: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" shows "unit_factor (f * unit_factor g) = unit_factor (f * g)" and "unit_factor (unit_factor f * g) = unit_factor (f * g)" proof - show "unit_factor (f * unit_factor g) = unit_factor (f * g)" proof (cases "f*g = 0") case False thus ?thesis using fps_mult_subdegree_ge[of f g] fps_unit_factor_shift[of "subdegree g" "f*g"] by (simp add: fps_shift_mult) next case True moreover have "f * unit_factor g = fps_shift (subdegree g) (f*g)" by (simp add: fps_shift_mult) ultimately show ?thesis by simp qed show "unit_factor (unit_factor f * g) = unit_factor (f * g)" proof (cases "f*g = 0") case False thus ?thesis using fps_mult_subdegree_ge[of f g] fps_unit_factor_shift[of "subdegree f" "f*g"] by (simp add: fps_shift_mult_right_noncomm) next case True moreover have "unit_factor f * g = fps_shift (subdegree f) (f*g)" by (simp add: fps_shift_mult_right_noncomm) ultimately show ?thesis by simp qed qed lemma fps_unit_factor_mult_both_unit_factor: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" shows "unit_factor (unit_factor f * unit_factor g) = unit_factor (f * g)" using fps_unit_factor_mult_unit_factor(1)[of "unit_factor f" g] fps_unit_factor_mult_unit_factor(2)[of f g] by simp lemma fps_unit_factor_mult': fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "f $ subdegree f * g $ subdegree g \ 0" shows "unit_factor (f * g) = unit_factor f * unit_factor g" using assms by (simp add: subdegree_mult' fps_shift_mult_both) lemma fps_unit_factor_mult: fixes f g :: "'a::semiring_no_zero_divisors fps" shows "unit_factor (f * g) = unit_factor f * unit_factor g" using fps_unit_factor_mult'[of f g] by (cases "f=0 \ g=0") auto definition "fps_cutoff n f = Abs_fps (\i. if i < n then f$i else 0)" lemma fps_cutoff_nth [simp]: "fps_cutoff n f $ i = (if i < n then f$i else 0)" unfolding fps_cutoff_def by simp lemma fps_cutoff_zero_iff: "fps_cutoff n f = 0 \ (f = 0 \ n \ subdegree f)" proof assume A: "fps_cutoff n f = 0" thus "f = 0 \ n \ subdegree f" proof (cases "f = 0") assume "f \ 0" with A have "n \ subdegree f" by (intro subdegree_geI) (simp_all add: fps_eq_iff split: if_split_asm) thus ?thesis .. qed simp qed (auto simp: fps_eq_iff intro: nth_less_subdegree_zero) lemma fps_cutoff_0 [simp]: "fps_cutoff 0 f = 0" by (simp add: fps_eq_iff) lemma fps_cutoff_zero [simp]: "fps_cutoff n 0 = 0" by (simp add: fps_eq_iff) lemma fps_cutoff_one: "fps_cutoff n 1 = (if n = 0 then 0 else 1)" by (simp add: fps_eq_iff) lemma fps_cutoff_fps_const: "fps_cutoff n (fps_const c) = (if n = 0 then 0 else fps_const c)" by (simp add: fps_eq_iff) lemma fps_cutoff_numeral: "fps_cutoff n (numeral c) = (if n = 0 then 0 else numeral c)" by (simp add: numeral_fps_const fps_cutoff_fps_const) lemma fps_shift_cutoff: "fps_shift n f * fps_X^n + fps_cutoff n f = f" by (simp add: fps_eq_iff fps_X_power_mult_right_nth) lemma fps_shift_cutoff': "fps_X^n * fps_shift n f + fps_cutoff n f = f" by (simp add: fps_eq_iff fps_X_power_mult_nth) lemma fps_cutoff_left_mult_nth: "k < n \ (fps_cutoff n f * g) $ k = (f * g) $ k" by (simp add: fps_mult_nth) lemma fps_cutoff_right_mult_nth: assumes "k < n" shows "(f * fps_cutoff n g) $ k = (f * g) $ k" proof- from assms have "\i\{0..k}. fps_cutoff n g $ (k - i) = g $ (k - i)" by auto thus ?thesis by (simp add: fps_mult_nth) qed subsection \Formal Power series form a metric space\ instantiation fps :: ("{minus,zero}") dist begin definition dist_fps_def: "dist (a :: 'a fps) b = (if a = b then 0 else inverse (2 ^ subdegree (a - b)))" lemma dist_fps_ge0: "dist (a :: 'a fps) b \ 0" by (simp add: dist_fps_def) instance .. end instantiation fps :: (group_add) metric_space begin definition uniformity_fps_def [code del]: "(uniformity :: ('a fps \ 'a fps) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_fps_def' [code del]: "open (U :: 'a fps set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" lemma dist_fps_sym: "dist (a :: 'a fps) b = dist b a" by (simp add: dist_fps_def) instance proof show th: "dist a b = 0 \ a = b" for a b :: "'a fps" by (simp add: dist_fps_def split: if_split_asm) then have th'[simp]: "dist a a = 0" for a :: "'a fps" by simp fix a b c :: "'a fps" consider "a = b" | "c = a \ c = b" | "a \ b" "a \ c" "b \ c" by blast then show "dist a b \ dist a c + dist b c" proof cases case 1 then show ?thesis by (simp add: dist_fps_def) next case 2 then show ?thesis by (cases "c = a") (simp_all add: th dist_fps_sym) next case neq: 3 have False if "dist a b > dist a c + dist b c" proof - let ?n = "subdegree (a - b)" from neq have "dist a b > 0" "dist b c > 0" and "dist a c > 0" by (simp_all add: dist_fps_def) with that have "dist a b > dist a c" and "dist a b > dist b c" by simp_all with neq have "?n < subdegree (a - c)" and "?n < subdegree (b - c)" by (simp_all add: dist_fps_def field_simps) hence "(a - c) $ ?n = 0" and "(b - c) $ ?n = 0" by (simp_all only: nth_less_subdegree_zero) hence "(a - b) $ ?n = 0" by simp moreover from neq have "(a - b) $ ?n \ 0" by (intro nth_subdegree_nonzero) simp_all ultimately show False by contradiction qed thus ?thesis by (auto simp add: not_le[symmetric]) qed qed (rule open_fps_def' uniformity_fps_def)+ end declare uniformity_Abort[where 'a="'a :: group_add fps", code] lemma open_fps_def: "open (S :: 'a::group_add fps set) = (\a \ S. \r. r >0 \ {y. dist y a < r} \ S)" unfolding open_dist subset_eq by simp text \The infinite sums and justification of the notation in textbooks.\ lemma reals_power_lt_ex: fixes x y :: real assumes xp: "x > 0" and y1: "y > 1" shows "\k>0. (1/y)^k < x" proof - have yp: "y > 0" using y1 by simp from reals_Archimedean2[of "max 0 (- log y x) + 1"] obtain k :: nat where k: "real k > max 0 (- log y x) + 1" by blast from k have kp: "k > 0" by simp from k have "real k > - log y x" by simp then have "ln y * real k > - ln x" unfolding log_def using ln_gt_zero_iff[OF yp] y1 by (simp add: minus_divide_left field_simps del: minus_divide_left[symmetric]) then have "ln y * real k + ln x > 0" by simp then have "exp (real k * ln y + ln x) > exp 0" by (simp add: ac_simps) then have "y ^ k * x > 1" unfolding exp_zero exp_add exp_of_nat_mult exp_ln [OF xp] exp_ln [OF yp] by simp then have "x > (1 / y)^k" using yp by (simp add: field_simps) then show ?thesis using kp by blast qed lemma fps_sum_rep_nth: "(sum (\i. fps_const(a$i)*fps_X^i) {0..m})$n = (if n \ m then a$n else 0)" by (simp add: fps_sum_nth if_distrib cong del: if_weak_cong) lemma fps_notation: "(\n. sum (\i. fps_const(a$i) * fps_X^i) {0..n}) \ a" (is "?s \ a") proof - have "\n0. \n \ n0. dist (?s n) a < r" if "r > 0" for r proof - obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0" using reals_power_lt_ex[OF \r > 0\, of 2] by auto show ?thesis proof - have "dist (?s n) a < r" if nn0: "n \ n0" for n proof - from that have thnn0: "(1/2)^n \ (1/2 :: real)^n0" by (simp add: field_split_simps) show ?thesis proof (cases "?s n = a") case True then show ?thesis unfolding dist_eq_0_iff[of "?s n" a, symmetric] using \r > 0\ by (simp del: dist_eq_0_iff) next case False from False have dth: "dist (?s n) a = (1/2)^subdegree (?s n - a)" by (simp add: dist_fps_def field_simps) from False have kn: "subdegree (?s n - a) > n" by (intro subdegree_greaterI) (simp_all add: fps_sum_rep_nth) then have "dist (?s n) a < (1/2)^n" by (simp add: field_simps dist_fps_def) also have "\ \ (1/2)^n0" using nn0 by (simp add: field_split_simps) also have "\ < r" using n0 by simp finally show ?thesis . qed qed then show ?thesis by blast qed qed then show ?thesis unfolding lim_sequentially by blast qed subsection \Inverses and division of formal power series\ declare sum.cong[fundef_cong] fun fps_left_inverse_constructor :: "'a::{comm_monoid_add,times,uminus} fps \ 'a \ nat \ 'a" where "fps_left_inverse_constructor f a 0 = a" | "fps_left_inverse_constructor f a (Suc n) = - sum (\i. fps_left_inverse_constructor f a i * f$(Suc n - i)) {0..n} * a" \ \This will construct a left inverse for f in case that x * f$0 = 1\ abbreviation "fps_left_inverse \ (\f x. Abs_fps (fps_left_inverse_constructor f x))" fun fps_right_inverse_constructor :: "'a::{comm_monoid_add,times,uminus} fps \ 'a \ nat \ 'a" where "fps_right_inverse_constructor f a 0 = a" | "fps_right_inverse_constructor f a n = - a * sum (\i. f$i * fps_right_inverse_constructor f a (n - i)) {1..n}" \ \This will construct a right inverse for f in case that f$0 * y = 1\ abbreviation "fps_right_inverse \ (\f y. Abs_fps (fps_right_inverse_constructor f y))" instantiation fps :: ("{comm_monoid_add,inverse,times,uminus}") inverse begin \ \For backwards compatibility.\ abbreviation natfun_inverse:: "'a fps \ nat \ 'a" where "natfun_inverse f \ fps_right_inverse_constructor f (inverse (f$0))" definition fps_inverse_def: "inverse f = Abs_fps (natfun_inverse f)" \ \ With scalars from a (possibly non-commutative) ring, this defines a right inverse. Furthermore, if scalars are of class @{class mult_zero} and satisfy condition @{term "inverse 0 = 0"}, then this will evaluate to zero when the zeroth term is zero. \ definition fps_divide_def: "f div g = fps_shift (subdegree g) (f * inverse (unit_factor g))" \ \ If scalars are of class @{class mult_zero} and satisfy condition @{term "inverse 0 = 0"}, then div by zero will equal zero. \ instance .. end lemma fps_lr_inverse_0_iff: "(fps_left_inverse f x) $ 0 = 0 \ x = 0" "(fps_right_inverse f x) $ 0 = 0 \ x = 0" by auto lemma fps_inverse_0_iff': "(inverse f) $ 0 = 0 \ inverse (f $ 0) = 0" by (simp add: fps_inverse_def fps_lr_inverse_0_iff(2)) lemma fps_inverse_0_iff[simp]: "(inverse f) $ 0 = (0::'a::division_ring) \ f $ 0 = 0" by (simp add: fps_inverse_0_iff') lemma fps_lr_inverse_nth_0: "(fps_left_inverse f x) $ 0 = x" "(fps_right_inverse f x) $ 0 = x" by auto lemma fps_inverse_nth_0 [simp]: "(inverse f) $ 0 = inverse (f $ 0)" by (simp add: fps_inverse_def) lemma fps_lr_inverse_starting0: fixes f :: "'a::{comm_monoid_add,mult_zero,uminus} fps" and g :: "'b::{ab_group_add,mult_zero} fps" shows "fps_left_inverse f 0 = 0" and "fps_right_inverse g 0 = 0" proof- show "fps_left_inverse f 0 = 0" proof (rule fps_ext) fix n show "fps_left_inverse f 0 $ n = 0 $ n" by (cases n) (simp_all add: fps_inverse_def) qed show "fps_right_inverse g 0 = 0" proof (rule fps_ext) fix n show "fps_right_inverse g 0 $ n = 0 $ n" by (cases n) (simp_all add: fps_inverse_def) qed qed lemma fps_lr_inverse_eq0_imp_starting0: "fps_left_inverse f x = 0 \ x = 0" "fps_right_inverse f x = 0 \ x = 0" proof- assume A: "fps_left_inverse f x = 0" have "0 = fps_left_inverse f x $ 0" by (subst A) simp thus "x = 0" by simp next assume A: "fps_right_inverse f x = 0" have "0 = fps_right_inverse f x $ 0" by (subst A) simp thus "x = 0" by simp qed lemma fps_lr_inverse_eq_0_iff: fixes x :: "'a::{comm_monoid_add,mult_zero,uminus}" and y :: "'b::{ab_group_add,mult_zero}" shows "fps_left_inverse f x = 0 \ x = 0" and "fps_right_inverse g y = 0 \ y = 0" using fps_lr_inverse_starting0 fps_lr_inverse_eq0_imp_starting0 by auto lemma fps_inverse_eq_0_iff': fixes f :: "'a::{ab_group_add,inverse,mult_zero} fps" shows "inverse f = 0 \ inverse (f $ 0) = 0" by (simp add: fps_inverse_def fps_lr_inverse_eq_0_iff(2)) lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::division_ring) fps) \ f $ 0 = 0" using fps_inverse_eq_0_iff'[of f] by simp lemmas fps_inverse_eq_0' = iffD2[OF fps_inverse_eq_0_iff'] lemmas fps_inverse_eq_0 = iffD2[OF fps_inverse_eq_0_iff] lemma fps_const_lr_inverse: fixes a :: "'a::{ab_group_add,mult_zero}" and b :: "'b::{comm_monoid_add,mult_zero,uminus}" shows "fps_left_inverse (fps_const a) x = fps_const x" and "fps_right_inverse (fps_const b) y = fps_const y" proof- show "fps_left_inverse (fps_const a) x = fps_const x" proof (rule fps_ext) fix n show "fps_left_inverse (fps_const a) x $ n = fps_const x $ n" by (cases n) auto qed show "fps_right_inverse (fps_const b) y = fps_const y" proof (rule fps_ext) fix n show "fps_right_inverse (fps_const b) y $ n = fps_const y $ n" by (cases n) auto qed qed lemma fps_const_inverse: fixes a :: "'a::{comm_monoid_add,inverse,mult_zero,uminus}" shows "inverse (fps_const a) = fps_const (inverse a)" unfolding fps_inverse_def by (simp add: fps_const_lr_inverse(2)) lemma fps_lr_inverse_zero: fixes x :: "'a::{ab_group_add,mult_zero}" and y :: "'b::{comm_monoid_add,mult_zero,uminus}" shows "fps_left_inverse 0 x = fps_const x" and "fps_right_inverse 0 y = fps_const y" using fps_const_lr_inverse[of 0] by simp_all lemma fps_inverse_zero_conv_fps_const: "inverse (0::'a::{comm_monoid_add,mult_zero,uminus,inverse} fps) = fps_const (inverse 0)" using fps_lr_inverse_zero(2)[of "inverse (0::'a)"] by (simp add: fps_inverse_def) lemma fps_inverse_zero': assumes "inverse (0::'a::{comm_monoid_add,inverse,mult_zero,uminus}) = 0" shows "inverse (0::'a fps) = 0" by (simp add: assms fps_inverse_zero_conv_fps_const) lemma fps_inverse_zero [simp]: "inverse (0::'a::division_ring fps) = 0" by (rule fps_inverse_zero'[OF inverse_zero]) lemma fps_lr_inverse_one: fixes x :: "'a::{ab_group_add,mult_zero,one}" and y :: "'b::{comm_monoid_add,mult_zero,uminus,one}" shows "fps_left_inverse 1 x = fps_const x" and "fps_right_inverse 1 y = fps_const y" using fps_const_lr_inverse[of 1] by simp_all lemma fps_lr_inverse_one_one: "fps_left_inverse 1 1 = (1::'a::{ab_group_add,mult_zero,one} fps)" "fps_right_inverse 1 1 = (1::'b::{comm_monoid_add,mult_zero,uminus,one} fps)" by (simp_all add: fps_lr_inverse_one) lemma fps_inverse_one': assumes "inverse (1::'a::{comm_monoid_add,inverse,mult_zero,uminus,one}) = 1" shows "inverse (1 :: 'a fps) = 1" using assms fps_lr_inverse_one_one(2) by (simp add: fps_inverse_def) lemma fps_inverse_one [simp]: "inverse (1 :: 'a :: division_ring fps) = 1" by (rule fps_inverse_one'[OF inverse_1]) lemma fps_lr_inverse_minus: fixes f :: "'a::ring_1 fps" shows "fps_left_inverse (-f) (-x) = - fps_left_inverse f x" and "fps_right_inverse (-f) (-x) = - fps_right_inverse f x" proof- show "fps_left_inverse (-f) (-x) = - fps_left_inverse f x" proof (intro fps_ext) fix n show "fps_left_inverse (-f) (-x) $ n = - fps_left_inverse f x $ n" proof (induct n rule: nat_less_induct) case (1 n) thus ?case by (cases n) (simp_all add: sum_negf algebra_simps) qed qed show "fps_right_inverse (-f) (-x) = - fps_right_inverse f x" proof (intro fps_ext) fix n show "fps_right_inverse (-f) (-x) $ n = - fps_right_inverse f x $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc m) with 1 have "\i\{1..Suc m}. fps_right_inverse (-f) (-x) $ (Suc m - i) = - fps_right_inverse f x $ (Suc m - i)" by auto with Suc show ?thesis by (simp add: sum_negf algebra_simps) qed simp qed qed qed lemma fps_inverse_minus [simp]: "inverse (-f) = -inverse (f :: 'a :: division_ring fps)" by (simp add: fps_inverse_def fps_lr_inverse_minus(2)) lemma fps_left_inverse: fixes f :: "'a::ring_1 fps" assumes f0: "x * f$0 = 1" shows "fps_left_inverse f x * f = 1" proof (rule fps_ext) fix n show "(fps_left_inverse f x * f) $ n = 1 $ n" by (cases n) (simp_all add: f0 fps_mult_nth mult.assoc) qed lemma fps_right_inverse: fixes f :: "'a::ring_1 fps" assumes f0: "f$0 * y = 1" shows "f * fps_right_inverse f y = 1" proof (rule fps_ext) fix n show "(f * fps_right_inverse f y) $ n = 1 $ n" proof (cases n) case (Suc k) moreover from Suc have "fps_right_inverse f y $ n = - y * sum (\i. f$i * fps_right_inverse_constructor f y (n - i)) {1..n}" by simp hence "(f * fps_right_inverse f y) $ n = - 1 * sum (\i. f$i * fps_right_inverse_constructor f y (n - i)) {1..n} + sum (\i. f$i * (fps_right_inverse_constructor f y (n - i))) {1..n}" by (simp add: fps_mult_nth sum.atLeast_Suc_atMost mult.assoc f0[symmetric]) thus "(f * fps_right_inverse f y) $ n = 1 $ n" by (simp add: Suc) qed (simp add: f0 fps_inverse_def) qed \ \ It is possible in a ring for an element to have a left inverse but not a right inverse, or vice versa. But when an element has both, they must be the same. \ lemma fps_left_inverse_eq_fps_right_inverse: fixes f :: "'a::ring_1 fps" assumes f0: "x * f$0 = 1" "f $ 0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "fps_left_inverse f x = fps_right_inverse f y" proof- from f0(2) have "f * fps_right_inverse f y = 1" by (simp add: fps_right_inverse) hence "fps_left_inverse f x * f * fps_right_inverse f y = fps_left_inverse f x" by (simp add: mult.assoc) moreover from f0(1) have "fps_left_inverse f x * f * fps_right_inverse f y = fps_right_inverse f y" by (simp add: fps_left_inverse) ultimately show ?thesis by simp qed lemma fps_left_inverse_eq_fps_right_inverse_comm: fixes f :: "'a::comm_ring_1 fps" assumes f0: "x * f$0 = 1" shows "fps_left_inverse f x = fps_right_inverse f x" using assms fps_left_inverse_eq_fps_right_inverse[of x f x] by (simp add: mult.commute) lemma fps_left_inverse': fixes f :: "'a::ring_1 fps" assumes "x * f$0 = 1" "f$0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "fps_right_inverse f y * f = 1" using assms fps_left_inverse_eq_fps_right_inverse[of x f y] fps_left_inverse[of x f] by simp lemma fps_right_inverse': fixes f :: "'a::ring_1 fps" assumes "x * f$0 = 1" "f$0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "f * fps_left_inverse f x = 1" using assms fps_left_inverse_eq_fps_right_inverse[of x f y] fps_right_inverse[of f y] by simp lemma inverse_mult_eq_1 [intro]: assumes "f$0 \ (0::'a::division_ring)" shows "inverse f * f = 1" using fps_left_inverse'[of "inverse (f$0)"] by (simp add: assms fps_inverse_def) lemma inverse_mult_eq_1': assumes "f$0 \ (0::'a::division_ring)" shows "f * inverse f = 1" using assms fps_right_inverse by (force simp: fps_inverse_def) lemma fps_mult_left_inverse_unit_factor: fixes f :: "'a::ring_1 fps" assumes "x * f $ subdegree f = 1" shows "fps_left_inverse (unit_factor f) x * f = fps_X ^ subdegree f" proof- have "fps_left_inverse (unit_factor f) x * f = fps_left_inverse (unit_factor f) x * unit_factor f * fps_X ^ subdegree f" using fps_unit_factor_decompose[of f] by (simp add: mult.assoc) with assms show ?thesis by (simp add: fps_left_inverse) qed lemma fps_mult_right_inverse_unit_factor: fixes f :: "'a::ring_1 fps" assumes "f $ subdegree f * y = 1" shows "f * fps_right_inverse (unit_factor f) y = fps_X ^ subdegree f" proof- have "f * fps_right_inverse (unit_factor f) y = fps_X ^ subdegree f * (unit_factor f * fps_right_inverse (unit_factor f) y)" using fps_unit_factor_decompose'[of f] by (simp add: mult.assoc[symmetric]) with assms show ?thesis by (simp add: fps_right_inverse) qed lemma fps_mult_right_inverse_unit_factor_divring: "(f :: 'a::division_ring fps) \ 0 \ f * inverse (unit_factor f) = fps_X ^ subdegree f" using fps_mult_right_inverse_unit_factor[of f] by (simp add: fps_inverse_def) lemma fps_left_inverse_idempotent_ring1: fixes f :: "'a::ring_1 fps" assumes "x * f$0 = 1" "y * x = 1" \ \These assumptions imply y equals f$0, but no need to assume that.\ shows "fps_left_inverse (fps_left_inverse f x) y = f" proof- from assms(1) have "fps_left_inverse (fps_left_inverse f x) y * fps_left_inverse f x * f = fps_left_inverse (fps_left_inverse f x) y" by (simp add: fps_left_inverse mult.assoc) moreover from assms(2) have "fps_left_inverse (fps_left_inverse f x) y * fps_left_inverse f x = 1" by (simp add: fps_left_inverse) ultimately show ?thesis by simp qed lemma fps_left_inverse_idempotent_comm_ring1: fixes f :: "'a::comm_ring_1 fps" assumes "x * f$0 = 1" shows "fps_left_inverse (fps_left_inverse f x) (f$0) = f" using assms fps_left_inverse_idempotent_ring1[of x f "f$0"] by (simp add: mult.commute) lemma fps_right_inverse_idempotent_ring1: fixes f :: "'a::ring_1 fps" assumes "f$0 * x = 1" "x * y = 1" \ \These assumptions imply y equals f$0, but no need to assume that.\ shows "fps_right_inverse (fps_right_inverse f x) y = f" proof- from assms(1) have "f * (fps_right_inverse f x * fps_right_inverse (fps_right_inverse f x) y) = fps_right_inverse (fps_right_inverse f x) y" by (simp add: fps_right_inverse mult.assoc[symmetric]) moreover from assms(2) have "fps_right_inverse f x * fps_right_inverse (fps_right_inverse f x) y = 1" by (simp add: fps_right_inverse) ultimately show ?thesis by simp qed lemma fps_right_inverse_idempotent_comm_ring1: fixes f :: "'a::comm_ring_1 fps" assumes "f$0 * x = 1" shows "fps_right_inverse (fps_right_inverse f x) (f$0) = f" using assms fps_right_inverse_idempotent_ring1[of f x "f$0"] by (simp add: mult.commute) lemma fps_inverse_idempotent[intro, simp]: "f$0 \ (0::'a::division_ring) \ inverse (inverse f) = f" using fps_right_inverse_idempotent_ring1[of f] by (simp add: fps_inverse_def) lemma fps_lr_inverse_unique_ring1: fixes f g :: "'a :: ring_1 fps" assumes fg: "f * g = 1" "g$0 * f$0 = 1" shows "fps_left_inverse g (f$0) = f" and "fps_right_inverse f (g$0) = g" proof- show "fps_left_inverse g (f$0) = f" proof (intro fps_ext) fix n show "fps_left_inverse g (f$0) $ n = f $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc k) hence "\i\{0..k}. fps_left_inverse g (f$0) $ i = f $ i" using 1 by simp hence "fps_left_inverse g (f$0) $ Suc k = f $ Suc k - 1 $ Suc k * f$0" by (simp add: fps_mult_nth fg(1)[symmetric] distrib_right mult.assoc fg(2)) with Suc show ?thesis by simp qed simp qed qed show "fps_right_inverse f (g$0) = g" proof (intro fps_ext) fix n show "fps_right_inverse f (g$0) $ n = g $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc k) hence "\i\{1..Suc k}. fps_right_inverse f (g$0) $ (Suc k - i) = g $ (Suc k - i)" using 1 by auto hence "fps_right_inverse f (g$0) $ Suc k = 1 * g $ Suc k - g$0 * 1 $ Suc k" by (simp add: fps_mult_nth fg(1)[symmetric] algebra_simps fg(2)[symmetric] sum.atLeast_Suc_atMost) with Suc show ?thesis by simp qed simp qed qed qed lemma fps_lr_inverse_unique_divring: fixes f g :: "'a ::division_ring fps" assumes fg: "f * g = 1" shows "fps_left_inverse g (f$0) = f" and "fps_right_inverse f (g$0) = g" proof- from fg have "f$0 * g$0 = 1" using fps_mult_nth_0[of f g] by simp hence "g$0 * f$0 = 1" using inverse_unique[of "f$0"] left_inverse[of "f$0"] by force thus "fps_left_inverse g (f$0) = f" "fps_right_inverse f (g$0) = g" using fg fps_lr_inverse_unique_ring1 by auto qed lemma fps_inverse_unique: fixes f g :: "'a :: division_ring fps" assumes fg: "f * g = 1" shows "inverse f = g" proof - from fg have if0: "inverse (f$0) = g$0" "f$0 \ 0" using inverse_unique[of "f$0"] fps_mult_nth_0[of f g] by auto with fg have "fps_right_inverse f (g$0) = g" using left_inverse[of "f$0"] by (intro fps_lr_inverse_unique_ring1(2)) simp_all with if0(1) show ?thesis by (simp add: fps_inverse_def) qed lemma inverse_fps_numeral: "inverse (numeral n :: ('a :: field_char_0) fps) = fps_const (inverse (numeral n))" by (intro fps_inverse_unique fps_ext) (simp_all add: fps_numeral_nth) lemma inverse_fps_of_nat: "inverse (of_nat n :: 'a :: {semiring_1,times,uminus,inverse} fps) = fps_const (inverse (of_nat n))" by (simp add: fps_of_nat fps_const_inverse[symmetric]) lemma fps_lr_inverse_mult_ring1: fixes f g :: "'a::ring_1 fps" assumes x: "x * f$0 = 1" "f$0 * x = 1" and y: "y * g$0 = 1" "g$0 * y = 1" shows "fps_left_inverse (f * g) (y*x) = fps_left_inverse g y * fps_left_inverse f x" and "fps_right_inverse (f * g) (y*x) = fps_right_inverse g y * fps_right_inverse f x" proof - define h where "h \ fps_left_inverse g y * fps_left_inverse f x" hence h0: "h$0 = y*x" by simp have "fps_left_inverse (f*g) (h$0) = h" proof (intro fps_lr_inverse_unique_ring1(1)) from h_def have "h * (f * g) = fps_left_inverse g y * (fps_left_inverse f x * f) * g" by (simp add: mult.assoc) thus "h * (f * g) = 1" using fps_left_inverse[OF x(1)] fps_left_inverse[OF y(1)] by simp from h_def have "(f*g)$0 * h$0 = f$0 * 1 * x" by (simp add: mult.assoc y(2)[symmetric]) with x(2) show "(f * g) $ 0 * h $ 0 = 1" by simp qed with h_def show "fps_left_inverse (f * g) (y*x) = fps_left_inverse g y * fps_left_inverse f x" by simp next define h where "h \ fps_right_inverse g y * fps_right_inverse f x" hence h0: "h$0 = y*x" by simp have "fps_right_inverse (f*g) (h$0) = h" proof (intro fps_lr_inverse_unique_ring1(2)) from h_def have "f * g * h = f * (g * fps_right_inverse g y) * fps_right_inverse f x" by (simp add: mult.assoc) thus "f * g * h = 1" using fps_right_inverse[OF x(2)] fps_right_inverse[OF y(2)] by simp from h_def have "h$0 * (f*g)$0 = y * 1 * g$0" by (simp add: mult.assoc x(1)[symmetric]) with y(1) show "h$0 * (f*g)$0 = 1" by simp qed with h_def show "fps_right_inverse (f * g) (y*x) = fps_right_inverse g y * fps_right_inverse f x" by simp qed lemma fps_lr_inverse_mult_divring: fixes f g :: "'a::division_ring fps" shows "fps_left_inverse (f * g) (inverse ((f*g)$0)) = fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0))" and "fps_right_inverse (f * g) (inverse ((f*g)$0)) = fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0))" proof- show "fps_left_inverse (f * g) (inverse ((f*g)$0)) = fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0))" proof (cases "f$0 = 0 \ g$0 = 0") case True hence "fps_left_inverse (f * g) (inverse ((f*g)$0)) = 0" by (simp add: fps_lr_inverse_eq_0_iff(1)) moreover from True have "fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0)) = 0" by (auto simp: fps_lr_inverse_eq_0_iff(1)) ultimately show ?thesis by simp next case False hence "fps_left_inverse (f * g) (inverse (g$0) * inverse (f$0)) = fps_left_inverse g (inverse (g$0)) * fps_left_inverse f (inverse (f$0))" by (intro fps_lr_inverse_mult_ring1(1)) simp_all with False show ?thesis by (simp add: nonzero_inverse_mult_distrib) qed show "fps_right_inverse (f * g) (inverse ((f*g)$0)) = fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0))" proof (cases "f$0 = 0 \ g$0 = 0") case True from True have "fps_right_inverse (f * g) (inverse ((f*g)$0)) = 0" by (simp add: fps_lr_inverse_eq_0_iff(2)) moreover from True have "fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0)) = 0" by (auto simp: fps_lr_inverse_eq_0_iff(2)) ultimately show ?thesis by simp next case False hence "fps_right_inverse (f * g) (inverse (g$0) * inverse (f$0)) = fps_right_inverse g (inverse (g$0)) * fps_right_inverse f (inverse (f$0))" by (intro fps_lr_inverse_mult_ring1(2)) simp_all with False show ?thesis by (simp add: nonzero_inverse_mult_distrib) qed qed lemma fps_inverse_mult_divring: "inverse (f * g) = inverse g * inverse (f :: 'a::division_ring fps)" using fps_lr_inverse_mult_divring(2) by (simp add: fps_inverse_def) lemma fps_inverse_mult: "inverse (f * g :: 'a::field fps) = inverse f * inverse g" by (simp add: fps_inverse_mult_divring) lemma fps_lr_inverse_gp_ring1: fixes ones ones_inv :: "'a :: ring_1 fps" defines "ones \ Abs_fps (\n. 1)" and "ones_inv \ Abs_fps (\n. if n=0 then 1 else if n=1 then - 1 else 0)" shows "fps_left_inverse ones 1 = ones_inv" and "fps_right_inverse ones 1 = ones_inv" proof- show "fps_left_inverse ones 1 = ones_inv" proof (rule fps_ext) fix n show "fps_left_inverse ones 1 $ n = ones_inv $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc m) have m: "n = Suc m" by fact moreover have "fps_left_inverse ones 1 $ Suc m = ones_inv $ Suc m" proof (cases m) case (Suc k) thus ?thesis using Suc m 1 by (simp add: ones_def ones_inv_def sum.atLeast_Suc_atMost) qed (simp add: ones_def ones_inv_def) ultimately show ?thesis by simp qed (simp add: ones_inv_def) qed qed moreover have "fps_right_inverse ones 1 = fps_left_inverse ones 1" by (auto intro: fps_left_inverse_eq_fps_right_inverse[symmetric] simp: ones_def) ultimately show "fps_right_inverse ones 1 = ones_inv" by simp qed lemma fps_lr_inverse_gp_ring1': fixes ones :: "'a :: ring_1 fps" defines "ones \ Abs_fps (\n. 1)" shows "fps_left_inverse ones 1 = 1 - fps_X" and "fps_right_inverse ones 1 = 1 - fps_X" proof- define ones_inv :: "'a :: ring_1 fps" where "ones_inv \ Abs_fps (\n. if n=0 then 1 else if n=1 then - 1 else 0)" hence "fps_left_inverse ones 1 = ones_inv" and "fps_right_inverse ones 1 = ones_inv" using ones_def fps_lr_inverse_gp_ring1 by auto thus "fps_left_inverse ones 1 = 1 - fps_X" and "fps_right_inverse ones 1 = 1 - fps_X" by (auto intro: fps_ext simp: ones_inv_def) qed lemma fps_inverse_gp: "inverse (Abs_fps(\n. (1::'a::division_ring))) = Abs_fps (\n. if n= 0 then 1 else if n=1 then - 1 else 0)" using fps_lr_inverse_gp_ring1(2) by (simp add: fps_inverse_def) lemma fps_inverse_gp': "inverse (Abs_fps (\n. 1::'a::division_ring)) = 1 - fps_X" by (simp add: fps_inverse_def fps_lr_inverse_gp_ring1'(2)) lemma fps_lr_inverse_one_minus_fps_X: fixes ones :: "'a :: ring_1 fps" defines "ones \ Abs_fps (\n. 1)" shows "fps_left_inverse (1 - fps_X) 1 = ones" and "fps_right_inverse (1 - fps_X) 1 = ones" proof- have "fps_left_inverse ones 1 = 1 - fps_X" using fps_lr_inverse_gp_ring1'(1) by (simp add: ones_def) thus "fps_left_inverse (1 - fps_X) 1 = ones" using fps_left_inverse_idempotent_ring1[of 1 ones 1] by (simp add: ones_def) have "fps_right_inverse ones 1 = 1 - fps_X" using fps_lr_inverse_gp_ring1'(2) by (simp add: ones_def) thus "fps_right_inverse (1 - fps_X) 1 = ones" using fps_right_inverse_idempotent_ring1[of ones 1 1] by (simp add: ones_def) qed lemma fps_inverse_one_minus_fps_X: fixes ones :: "'a :: division_ring fps" defines "ones \ Abs_fps (\n. 1)" shows "inverse (1 - fps_X) = ones" by (simp add: fps_inverse_def assms fps_lr_inverse_one_minus_fps_X(2)) lemma fps_lr_one_over_one_minus_fps_X_squared: shows "fps_left_inverse ((1 - fps_X)^2) (1::'a::ring_1) = Abs_fps (\n. of_nat (n+1))" "fps_right_inverse ((1 - fps_X)^2) (1::'a) = Abs_fps (\n. of_nat (n+1))" proof- define f invf2 :: "'a fps" where "f \ (1 - fps_X)" and "invf2 \ Abs_fps (\n. of_nat (n+1))" have f2_nth_simps: "f^2 $ 1 = - of_nat 2" "f^2 $ 2 = 1" "\n. n>2 \ f^2 $ n = 0" by (simp_all add: power2_eq_square f_def fps_mult_nth sum.atLeast_Suc_atMost) show "fps_left_inverse (f^2) 1 = invf2" proof (intro fps_ext) fix n show "fps_left_inverse (f^2) 1 $ n = invf2 $ n" proof (induct n rule: nat_less_induct) case (1 t) hence induct_assm: "\m. m < t \ fps_left_inverse (f\<^sup>2) 1 $ m = invf2 $ m" by fast show ?case proof (cases t) case (Suc m) have m: "t = Suc m" by fact moreover have "fps_left_inverse (f^2) 1 $ Suc m = invf2 $ Suc m" proof (cases m) case 0 thus ?thesis using f2_nth_simps(1) by (simp add: invf2_def) next case (Suc l) have l: "m = Suc l" by fact moreover have "fps_left_inverse (f^2) 1 $ Suc (Suc l) = invf2 $ Suc (Suc l)" proof (cases l) case 0 thus ?thesis using f2_nth_simps(1,2) by (simp add: Suc_1[symmetric] invf2_def) next case (Suc k) from Suc l m have A: "fps_left_inverse (f\<^sup>2) 1 $ Suc (Suc k) = invf2 $ Suc (Suc k)" and B: "fps_left_inverse (f\<^sup>2) 1 $ Suc k = invf2 $ Suc k" using induct_assm[of "Suc k"] induct_assm[of "Suc (Suc k)"] by auto have times2: "\a::nat. 2*a = a + a" by simp have "\i\{0..k}. (f^2)$(Suc (Suc (Suc k)) - i) = 0" using f2_nth_simps(3) by auto hence "fps_left_inverse (f^2) 1 $ Suc (Suc (Suc k)) = fps_left_inverse (f\<^sup>2) 1 $ Suc (Suc k) * of_nat 2 - fps_left_inverse (f\<^sup>2) 1 $ Suc k" using sum.ub_add_nat f2_nth_simps(1,2) by simp also have "\ = of_nat (2 * Suc (Suc (Suc k))) - of_nat (Suc (Suc k))" by (subst A, subst B) (simp add: invf2_def mult.commute) also have "\ = of_nat (Suc (Suc (Suc k)) + 1)" by (subst times2[of "Suc (Suc (Suc k))"]) simp finally have "fps_left_inverse (f^2) 1 $ Suc (Suc (Suc k)) = invf2 $ Suc (Suc (Suc k))" by (simp add: invf2_def) with Suc show ?thesis by simp qed ultimately show ?thesis by simp qed ultimately show ?thesis by simp qed (simp add: invf2_def) qed qed moreover have "fps_right_inverse (f^2) 1 = fps_left_inverse (f^2) 1" by (auto intro: fps_left_inverse_eq_fps_right_inverse[symmetric] simp: f_def power2_eq_square ) ultimately show "fps_right_inverse (f^2) 1 = invf2" by simp qed lemma fps_one_over_one_minus_fps_X_squared': assumes "inverse (1::'a::{ring_1,inverse}) = 1" shows "inverse ((1 - fps_X)^2 :: 'a fps) = Abs_fps (\n. of_nat (n+1))" using assms fps_lr_one_over_one_minus_fps_X_squared(2) by (simp add: fps_inverse_def power2_eq_square) lemma fps_one_over_one_minus_fps_X_squared: "inverse ((1 - fps_X)^2 :: 'a :: division_ring fps) = Abs_fps (\n. of_nat (n+1))" by (rule fps_one_over_one_minus_fps_X_squared'[OF inverse_1]) lemma fps_lr_inverse_fps_X_plus1: "fps_left_inverse (1 + fps_X) (1::'a::ring_1) = Abs_fps (\n. (-1)^n)" "fps_right_inverse (1 + fps_X) (1::'a) = Abs_fps (\n. (-1)^n)" proof- show "fps_left_inverse (1 + fps_X) (1::'a) = Abs_fps (\n. (-1)^n)" proof (rule fps_ext) fix n show "fps_left_inverse (1 + fps_X) (1::'a) $ n = Abs_fps (\n. (-1)^n) $ n" proof (induct n rule: nat_less_induct) case (1 n) show ?case proof (cases n) case (Suc m) have m: "n = Suc m" by fact from Suc 1 have A: "fps_left_inverse (1 + fps_X) (1::'a) $ n = - (\i=0..m. (- 1)^i * (1 + fps_X) $ (Suc m - i))" by simp show ?thesis proof (cases m) case (Suc l) have "\i\{0..l}. ((1::'a fps) + fps_X) $ (Suc (Suc l) - i) = 0" by auto with Suc A m show ?thesis by simp qed (simp add: m A) qed simp qed qed moreover have "fps_right_inverse (1 + fps_X) (1::'a) = fps_left_inverse (1 + fps_X) 1" by (intro fps_left_inverse_eq_fps_right_inverse[symmetric]) simp_all ultimately show "fps_right_inverse (1 + fps_X) (1::'a) = Abs_fps (\n. (-1)^n)" by simp qed lemma fps_inverse_fps_X_plus1': assumes "inverse (1::'a::{ring_1,inverse}) = 1" shows "inverse (1 + fps_X) = Abs_fps (\n. (- (1::'a)) ^ n)" using assms fps_lr_inverse_fps_X_plus1(2) by (simp add: fps_inverse_def) lemma fps_inverse_fps_X_plus1: "inverse (1 + fps_X) = Abs_fps (\n. (- (1::'a::division_ring)) ^ n)" by (rule fps_inverse_fps_X_plus1'[OF inverse_1]) lemma subdegree_lr_inverse: fixes x :: "'a::{comm_monoid_add,mult_zero,uminus}" and y :: "'b::{ab_group_add,mult_zero}" shows "subdegree (fps_left_inverse f x) = 0" and "subdegree (fps_right_inverse g y) = 0" proof- show "subdegree (fps_left_inverse f x) = 0" using fps_lr_inverse_eq_0_iff(1) subdegree_eq_0_iff by fastforce show "subdegree (fps_right_inverse g y) = 0" using fps_lr_inverse_eq_0_iff(2) subdegree_eq_0_iff by fastforce qed lemma subdegree_inverse [simp]: fixes f :: "'a::{ab_group_add,inverse,mult_zero} fps" shows "subdegree (inverse f) = 0" using subdegree_lr_inverse(2) by (simp add: fps_inverse_def) lemma fps_div_zero [simp]: "0 div (g :: 'a :: {comm_monoid_add,inverse,mult_zero,uminus} fps) = 0" by (simp add: fps_divide_def) lemma fps_div_by_zero': fixes g :: "'a::{comm_monoid_add,inverse,mult_zero,uminus} fps" assumes "inverse (0::'a) = 0" shows "g div 0 = 0" by (simp add: fps_divide_def assms fps_inverse_zero') lemma fps_div_by_zero [simp]: "(g::'a::division_ring fps) div 0 = 0" by (rule fps_div_by_zero'[OF inverse_zero]) lemma fps_divide_unit': "subdegree g = 0 \ f div g = f * inverse g" by (simp add: fps_divide_def) lemma fps_divide_unit: "g$0 \ 0 \ f div g = f * inverse g" by (intro fps_divide_unit') (simp add: subdegree_eq_0_iff) lemma fps_divide_nth_0': "subdegree (g::'a::division_ring fps) = 0 \ (f div g) $ 0 = f $ 0 / (g $ 0)" by (simp add: fps_divide_unit' divide_inverse) lemma fps_divide_nth_0 [simp]: "g $ 0 \ 0 \ (f div g) $ 0 = f $ 0 / (g $ 0 :: _ :: division_ring)" by (simp add: fps_divide_nth_0') lemma fps_divide_nth_below: fixes f g :: "'a::{comm_monoid_add,uminus,mult_zero,inverse} fps" shows "n < subdegree f - subdegree g \ (f div g) $ n = 0" by (simp add: fps_divide_def fps_mult_nth_eq0) lemma fps_divide_nth_base: fixes f g :: "'a::division_ring fps" assumes "subdegree g \ subdegree f" shows "(f div g) $ (subdegree f - subdegree g) = f $ subdegree f * inverse (g $ subdegree g)" by (simp add: assms fps_divide_def fps_divide_unit') lemma fps_divide_subdegree_ge: fixes f g :: "'a::{comm_monoid_add,uminus,mult_zero,inverse} fps" assumes "f / g \ 0" shows "subdegree (f / g) \ subdegree f - subdegree g" by (intro subdegree_geI) (simp_all add: assms fps_divide_nth_below) lemma fps_divide_subdegree: fixes f g :: "'a::division_ring fps" assumes "f \ 0" "g \ 0" "subdegree g \ subdegree f" shows "subdegree (f / g) = subdegree f - subdegree g" proof (intro antisym) from assms have 1: "(f div g) $ (subdegree f - subdegree g) \ 0" using fps_divide_nth_base[of g f] by simp thus "subdegree (f / g) \ subdegree f - subdegree g" by (intro subdegree_leI) simp from 1 have "f / g \ 0" by (auto intro: fps_nonzeroI) thus "subdegree f - subdegree g \ subdegree (f / g)" by (rule fps_divide_subdegree_ge) qed lemma fps_divide_shift_numer: fixes f g :: "'a::{inverse,comm_monoid_add,uminus,mult_zero} fps" assumes "n \ subdegree f" shows "fps_shift n f / g = fps_shift n (f/g)" using assms fps_shift_mult_right_noncomm[of n f "inverse (unit_factor g)"] fps_shift_fps_shift_reorder[of "subdegree g" n "f * inverse (unit_factor g)"] by (simp add: fps_divide_def) lemma fps_divide_shift_denom: fixes f g :: "'a::{inverse,comm_monoid_add,uminus,mult_zero} fps" assumes "n \ subdegree g" "subdegree g \ subdegree f" shows "f / fps_shift n g = Abs_fps (\k. if kk. if k subdegree f" shows "f / unit_factor g = Abs_fps (\k. if k subdegree f" shows "unit_factor f / unit_factor g = fps_shift (subdegree f - subdegree g) (f / g)" using assms fps_divide_unit_factor_numer[of f "unit_factor g"] fps_divide_unit_factor_denom[of g f] fps_shift_rev_shift(1)[of "subdegree g" "subdegree f" "f/g"] by simp lemma fps_divide_unit_factor_both: fixes f g :: "'a::division_ring fps" assumes "subdegree g \ subdegree f" shows "unit_factor f / unit_factor g = unit_factor (f / g)" using assms fps_divide_unit_factor_both'[of g f] fps_divide_subdegree[of f g] by (cases "f=0 \ g=0") auto lemma fps_divide_self: "(f::'a::division_ring fps) \ 0 \ f / f = 1" using fps_mult_right_inverse_unit_factor_divring[of f] by (simp add: fps_divide_def) lemma fps_divide_add: fixes f g h :: "'a::{semiring_0,inverse,uminus} fps" shows "(f + g) / h = f / h + g / h" by (simp add: fps_divide_def algebra_simps fps_shift_add) lemma fps_divide_diff: fixes f g h :: "'a::{ring,inverse} fps" shows "(f - g) / h = f / h - g / h" by (simp add: fps_divide_def algebra_simps fps_shift_diff) lemma fps_divide_uminus: fixes f g h :: "'a::{ring,inverse} fps" shows "(- f) / g = - (f / g)" by (simp add: fps_divide_def algebra_simps fps_shift_uminus) lemma fps_divide_uminus': fixes f g h :: "'a::division_ring fps" shows "f / (- g) = - (f / g)" by (simp add: fps_divide_def fps_unit_factor_uminus fps_shift_uminus) lemma fps_divide_times: fixes f g h :: "'a::{semiring_0,inverse,uminus} fps" assumes "subdegree h \ subdegree g" shows "(f * g) / h = f * (g / h)" using assms fps_mult_subdegree_ge[of g "inverse (unit_factor h)"] fps_shift_mult[of "subdegree h" "g * inverse (unit_factor h)" f] by (fastforce simp add: fps_divide_def mult.assoc) lemma fps_divide_times2: fixes f g h :: "'a::{comm_semiring_0,inverse,uminus} fps" assumes "subdegree h \ subdegree f" shows "(f * g) / h = (f / h) * g" using assms fps_divide_times[of h f g] by (simp add: mult.commute) lemma fps_times_divide_eq: fixes f g :: "'a::field fps" assumes "g \ 0" and "subdegree f \ subdegree g" shows "f div g * g = f" using assms fps_divide_times2[of g f g] by (simp add: fps_divide_times fps_divide_self) lemma fps_divide_times_eq: "(g :: 'a::division_ring fps) \ 0 \ (f * g) div g = f" by (simp add: fps_divide_times fps_divide_self) lemma fps_divide_by_mult': fixes f g h :: "'a :: division_ring fps" assumes "subdegree h \ subdegree f" shows "f / (g * h) = f / h / g" proof (cases "f=0 \ g=0 \ h=0") case False with assms show ?thesis using fps_unit_factor_mult[of g h] by (auto simp: fps_divide_def fps_shift_fps_shift fps_inverse_mult_divring mult.assoc fps_shift_mult_right_noncomm ) qed auto lemma fps_divide_by_mult: fixes f g h :: "'a :: field fps" assumes "subdegree g \ subdegree f" shows "f / (g * h) = f / g / h" proof- have "f / (g * h) = f / (h * g)" by (simp add: mult.commute) also have "\ = f / g / h" using fps_divide_by_mult'[OF assms] by simp finally show ?thesis by simp qed lemma fps_divide_cancel: fixes f g h :: "'a :: division_ring fps" shows "h \ 0 \ (f * h) div (g * h) = f div g" by (cases "f=0") (auto simp: fps_divide_by_mult' fps_divide_times_eq) lemma fps_divide_1': fixes a :: "'a::{comm_monoid_add,inverse,mult_zero,uminus,zero_neq_one,monoid_mult} fps" assumes "inverse (1::'a) = 1" shows "a / 1 = a" using assms fps_inverse_one' fps_one_mult(2)[of a] by (force simp: fps_divide_def) lemma fps_divide_1 [simp]: "(a :: 'a::division_ring fps) / 1 = a" by (rule fps_divide_1'[OF inverse_1]) lemma fps_divide_X': fixes f :: "'a::{comm_monoid_add,inverse,mult_zero,uminus,zero_neq_one,monoid_mult} fps" assumes "inverse (1::'a) = 1" shows "f / fps_X = fps_shift 1 f" using assms fps_one_mult(2)[of f] by (simp add: fps_divide_def fps_X_unit_factor fps_inverse_one') lemma fps_divide_X [simp]: "a / fps_X = fps_shift 1 (a::'a::division_ring fps)" by (rule fps_divide_X'[OF inverse_1]) lemma fps_divide_X_power': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "f / (fps_X ^ n) = fps_shift n f" using fps_inverse_one'[OF assms] fps_one_mult(2)[of f] by (simp add: fps_divide_def fps_X_power_subdegree) lemma fps_divide_X_power [simp]: "a / (fps_X ^ n) = fps_shift n (a::'a::division_ring fps)" by (rule fps_divide_X_power'[OF inverse_1]) lemma fps_divide_shift_denom_conv_times_fps_X_power: fixes f g :: "'a::{semiring_1,inverse,uminus} fps" assumes "n \ subdegree g" "subdegree g \ subdegree f" shows "f / fps_shift n g = f / g * fps_X ^ n" using assms by (intro fps_ext) (simp_all add: fps_divide_shift_denom fps_X_power_mult_right_nth) lemma fps_divide_unit_factor_denom_conv_times_fps_X_power: fixes f g :: "'a::{semiring_1,inverse,uminus} fps" assumes "subdegree g \ subdegree f" shows "f / unit_factor g = f / g * fps_X ^ subdegree g" by (simp add: assms fps_divide_shift_denom_conv_times_fps_X_power) lemma fps_shift_altdef': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "fps_shift n f = f div fps_X^n" using assms by (simp add: fps_divide_def fps_X_power_subdegree fps_X_power_unit_factor fps_inverse_one' ) lemma fps_shift_altdef: "fps_shift n f = (f :: 'a :: division_ring fps) div fps_X^n" by (rule fps_shift_altdef'[OF inverse_1]) lemma fps_div_fps_X_power_nth': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "(f div fps_X^n) $ k = f $ (k + n)" using assms by (simp add: fps_shift_altdef' [symmetric]) lemma fps_div_fps_X_power_nth: "((f :: 'a :: division_ring fps) div fps_X^n) $ k = f $ (k + n)" by (rule fps_div_fps_X_power_nth'[OF inverse_1]) lemma fps_div_fps_X_nth': fixes f :: "'a::{semiring_1,inverse,uminus} fps" assumes "inverse (1::'a) = 1" shows "(f div fps_X) $ k = f $ Suc k" using assms fps_div_fps_X_power_nth'[of f 1] by simp lemma fps_div_fps_X_nth: "((f :: 'a :: division_ring fps) div fps_X) $ k = f $ Suc k" by (rule fps_div_fps_X_nth'[OF inverse_1]) lemma divide_fps_const': fixes c :: "'a :: {inverse,comm_monoid_add,uminus,mult_zero}" shows "f / fps_const c = f * fps_const (inverse c)" by (simp add: fps_divide_def fps_const_inverse) lemma divide_fps_const [simp]: fixes c :: "'a :: {comm_semiring_0,inverse,uminus}" shows "f / fps_const c = fps_const (inverse c) * f" by (simp add: divide_fps_const' mult.commute) lemma fps_const_divide: "fps_const (x :: _ :: division_ring) / fps_const y = fps_const (x / y)" by (simp add: fps_divide_def fps_const_inverse divide_inverse) lemma fps_numeral_divide_divide: "x / numeral b / numeral c = (x / numeral (b * c) :: 'a :: field fps)" by (simp add: fps_divide_by_mult[symmetric]) lemma fps_numeral_mult_divide: "numeral b * x / numeral c = (numeral b / numeral c * x :: 'a :: field fps)" by (simp add: fps_divide_times2) lemmas fps_numeral_simps = fps_numeral_divide_divide fps_numeral_mult_divide inverse_fps_numeral neg_numeral_fps_const lemma fps_is_left_unit_iff_zeroth_is_left_unit: fixes f :: "'a :: ring_1 fps" shows "(\g. 1 = f * g) \ (\k. 1 = f$0 * k)" proof assume "\g. 1 = f * g" then obtain g where "1 = f * g" by fast hence "1 = f$0 * g$0" using fps_mult_nth_0[of f g] by simp thus "\k. 1 = f$0 * k" by auto next assume "\k. 1 = f$0 * k" then obtain k where "1 = f$0 * k" by fast hence "1 = f * fps_right_inverse f k" using fps_right_inverse by simp thus "\g. 1 = f * g" by fast qed lemma fps_is_right_unit_iff_zeroth_is_right_unit: fixes f :: "'a :: ring_1 fps" shows "(\g. 1 = g * f) \ (\k. 1 = k * f$0)" proof assume "\g. 1 = g * f" then obtain g where "1 = g * f" by fast hence "1 = g$0 * f$0" using fps_mult_nth_0[of g f] by simp thus "\k. 1 = k * f$0" by auto next assume "\k. 1 = k * f$0" then obtain k where "1 = k * f$0" by fast hence "1 = fps_left_inverse f k * f" using fps_left_inverse by simp thus "\g. 1 = g * f" by fast qed lemma fps_is_unit_iff [simp]: "(f :: 'a :: field fps) dvd 1 \ f $ 0 \ 0" proof assume "f dvd 1" then obtain g where "1 = f * g" by (elim dvdE) from this[symmetric] have "(f*g) $ 0 = 1" by simp thus "f $ 0 \ 0" by auto next assume A: "f $ 0 \ 0" thus "f dvd 1" by (simp add: inverse_mult_eq_1[OF A, symmetric]) qed lemma subdegree_eq_0_left: fixes f :: "'a::{comm_monoid_add,zero_neq_one,mult_zero} fps" assumes "\g. 1 = f * g" shows "subdegree f = 0" proof (intro subdegree_eq_0) from assms obtain g where "1 = f * g" by fast hence "f$0 * g$0 = 1" using fps_mult_nth_0[of f g] by simp thus "f$0 \ 0" by auto qed lemma subdegree_eq_0_right: fixes f :: "'a::{comm_monoid_add,zero_neq_one,mult_zero} fps" assumes "\g. 1 = g * f" shows "subdegree f = 0" proof (intro subdegree_eq_0) from assms obtain g where "1 = g * f" by fast hence "g$0 * f$0 = 1" using fps_mult_nth_0[of g f] by simp thus "f$0 \ 0" by auto qed lemma subdegree_eq_0' [simp]: "(f :: 'a :: field fps) dvd 1 \ subdegree f = 0" by simp lemma fps_dvd1_left_trivial_unit_factor: fixes f :: "'a::{comm_monoid_add, zero_neq_one, mult_zero} fps" assumes "\g. 1 = f * g" shows "unit_factor f = f" using assms subdegree_eq_0_left by fastforce lemma fps_dvd1_right_trivial_unit_factor: fixes f :: "'a::{comm_monoid_add, zero_neq_one, mult_zero} fps" assumes "\g. 1 = g * f" shows "unit_factor f = f" using assms subdegree_eq_0_right by fastforce lemma fps_dvd1_trivial_unit_factor: "(f :: 'a::comm_semiring_1 fps) dvd 1 \ unit_factor f = f" unfolding dvd_def by (rule fps_dvd1_left_trivial_unit_factor) simp lemma fps_unit_dvd_left: fixes f :: "'a :: division_ring fps" assumes "f $ 0 \ 0" shows "\g. 1 = f * g" using assms fps_is_left_unit_iff_zeroth_is_left_unit right_inverse by fastforce lemma fps_unit_dvd_right: fixes f :: "'a :: division_ring fps" assumes "f $ 0 \ 0" shows "\g. 1 = g * f" using assms fps_is_right_unit_iff_zeroth_is_right_unit left_inverse by fastforce lemma fps_unit_dvd [simp]: "(f $ 0 :: 'a :: field) \ 0 \ f dvd g" using fps_unit_dvd_left dvd_trans[of f 1] by simp lemma dvd_left_imp_subdegree_le: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "\k. g = f * k" "g \ 0" shows "subdegree f \ subdegree g" using assms fps_mult_subdegree_ge by fastforce lemma dvd_right_imp_subdegree_le: fixes f g :: "'a::{comm_monoid_add,mult_zero} fps" assumes "\k. g = k * f" "g \ 0" shows "subdegree f \ subdegree g" using assms fps_mult_subdegree_ge by fastforce lemma dvd_imp_subdegree_le: "f dvd g \ g \ 0 \ subdegree f \ subdegree g" using dvd_left_imp_subdegree_le by fast lemma subdegree_le_imp_dvd_left_ring1: fixes f g :: "'a :: ring_1 fps" assumes "\y. f $ subdegree f * y = 1" "subdegree f \ subdegree g" shows "\k. g = f * k" proof- define h :: "'a fps" where "h \ fps_X ^ (subdegree g - subdegree f)" from assms(1) obtain y where "f $ subdegree f * y = 1" by fast hence "unit_factor f $ 0 * y = 1" by simp from this obtain k where "1 = unit_factor f * k" using fps_is_left_unit_iff_zeroth_is_left_unit[of "unit_factor f"] by auto hence "fps_X ^ subdegree f = fps_X ^ subdegree f * unit_factor f * k" by (simp add: mult.assoc) moreover have "fps_X ^ subdegree f * unit_factor f = f" by (rule fps_unit_factor_decompose'[symmetric]) ultimately have "fps_X ^ (subdegree f + (subdegree g - subdegree f)) = f * k * h" by (simp add: power_add h_def) hence "g = f * (k * h * unit_factor g)" using fps_unit_factor_decompose'[of g] by (simp add: assms(2) mult.assoc) thus ?thesis by fast qed lemma subdegree_le_imp_dvd_left_divring: fixes f g :: "'a :: division_ring fps" assumes "f \ 0" "subdegree f \ subdegree g" shows "\k. g = f * k" proof (intro subdegree_le_imp_dvd_left_ring1) from assms(1) have "f $ subdegree f \ 0" by simp thus "\y. f $ subdegree f * y = 1" using right_inverse by blast qed (rule assms(2)) lemma subdegree_le_imp_dvd_right_ring1: fixes f g :: "'a :: ring_1 fps" assumes "\x. x * f $ subdegree f = 1" "subdegree f \ subdegree g" shows "\k. g = k * f" proof- define h :: "'a fps" where "h \ fps_X ^ (subdegree g - subdegree f)" from assms(1) obtain x where "x * f $ subdegree f = 1" by fast hence "x * unit_factor f $ 0 = 1" by simp from this obtain k where "1 = k * unit_factor f" using fps_is_right_unit_iff_zeroth_is_right_unit[of "unit_factor f"] by auto hence "fps_X ^ subdegree f = k * (unit_factor f * fps_X ^ subdegree f)" by (simp add: mult.assoc[symmetric]) moreover have "unit_factor f * fps_X ^ subdegree f = f" by (rule fps_unit_factor_decompose[symmetric]) ultimately have "fps_X ^ (subdegree g - subdegree f + subdegree f) = h * k * f" by (simp add: power_add h_def mult.assoc) hence "g = unit_factor g * h * k * f" using fps_unit_factor_decompose[of g] by (simp add: assms(2) mult.assoc) thus ?thesis by fast qed lemma subdegree_le_imp_dvd_right_divring: fixes f g :: "'a :: division_ring fps" assumes "f \ 0" "subdegree f \ subdegree g" shows "\k. g = k * f" proof (intro subdegree_le_imp_dvd_right_ring1) from assms(1) have "f $ subdegree f \ 0" by simp thus "\x. x * f $ subdegree f = 1" using left_inverse by blast qed (rule assms(2)) lemma fps_dvd_iff: assumes "(f :: 'a :: field fps) \ 0" "g \ 0" shows "f dvd g \ subdegree f \ subdegree g" proof assume "subdegree f \ subdegree g" with assms show "f dvd g" using subdegree_le_imp_dvd_left_divring by (auto intro: dvdI) qed (simp add: assms dvd_imp_subdegree_le) lemma subdegree_div': fixes p q :: "'a::division_ring fps" assumes "\k. p = k * q" shows "subdegree (p div q) = subdegree p - subdegree q" proof (cases "p = 0") case False from assms(1) obtain k where k: "p = k * q" by blast with False have "subdegree (p div q) = subdegree k" by (simp add: fps_divide_times_eq) moreover have "k $ subdegree k * q $ subdegree q \ 0" proof assume "k $ subdegree k * q $ subdegree q = 0" hence "k $ subdegree k * q $ subdegree q * inverse (q $ subdegree q) = 0" by simp with False k show False by (simp add: mult.assoc) qed ultimately show ?thesis by (simp add: k subdegree_mult') qed simp lemma subdegree_div: fixes p q :: "'a :: field fps" assumes "q dvd p" shows "subdegree (p div q) = subdegree p - subdegree q" using assms unfolding dvd_def by (auto intro: subdegree_div') lemma subdegree_div_unit': fixes p q :: "'a :: {ab_group_add,mult_zero,inverse} fps" assumes "q $ 0 \ 0" "p $ subdegree p * inverse (q $ 0) \ 0" shows "subdegree (p div q) = subdegree p" using assms subdegree_mult'[of p "inverse q"] by (auto simp add: fps_divide_unit) lemma subdegree_div_unit'': fixes p q :: "'a :: {ring_no_zero_divisors,inverse} fps" assumes "q $ 0 \ 0" "inverse (q $ 0) \ 0" shows "subdegree (p div q) = subdegree p" by (cases "p = 0") (auto intro: subdegree_div_unit' simp: assms) lemma subdegree_div_unit: fixes p q :: "'a :: division_ring fps" assumes "q $ 0 \ 0" shows "subdegree (p div q) = subdegree p" by (intro subdegree_div_unit'') (simp_all add: assms) instantiation fps :: ("{comm_semiring_1,inverse,uminus}") modulo begin definition fps_mod_def: "f mod g = (if g = 0 then f else let h = unit_factor g in fps_cutoff (subdegree g) (f * inverse h) * h)" instance .. end lemma fps_mod_zero [simp]: "(f::'a::{comm_semiring_1,inverse,uminus} fps) mod 0 = f" by (simp add: fps_mod_def) lemma fps_mod_eq_zero: assumes "g \ 0" and "subdegree f \ subdegree g" shows "f mod g = 0" proof (cases "f * inverse (unit_factor g) = 0") case False have "fps_cutoff (subdegree g) (f * inverse (unit_factor g)) = 0" using False assms(2) fps_mult_subdegree_ge fps_cutoff_zero_iff by force with assms(1) show ?thesis by (simp add: fps_mod_def Let_def) qed (simp add: assms fps_mod_def) lemma fps_mod_unit [simp]: "g$0 \ 0 \ f mod g = 0" by (intro fps_mod_eq_zero) auto lemma subdegree_mod: assumes "subdegree (f::'a::field fps) < subdegree g" shows "subdegree (f mod g) = subdegree f" proof (cases "f = 0") case False with assms show ?thesis by (intro subdegreeI) (auto simp: inverse_mult_eq_1 fps_mod_def Let_def fps_cutoff_left_mult_nth mult.assoc) qed (simp add: fps_mod_def) instance fps :: (field) idom_modulo proof fix f g :: "'a fps" define n where "n = subdegree g" define h where "h = f * inverse (unit_factor g)" show "f div g * g + f mod g = f" proof (cases "g = 0") case False with n_def h_def have "f div g * g + f mod g = (fps_shift n h * fps_X ^ n + fps_cutoff n h) * unit_factor g" by (simp add: fps_divide_def fps_mod_def Let_def subdegree_decompose algebra_simps) with False show ?thesis by (simp add: fps_shift_cutoff h_def inverse_mult_eq_1) qed auto qed (rule fps_divide_times_eq, simp_all add: fps_divide_def) instantiation fps :: (field) normalization_semidom_multiplicative begin definition fps_normalize_def [simp]: "normalize f = (if f = 0 then 0 else fps_X ^ subdegree f)" instance proof fix f g :: "'a fps" assume "is_unit f" thus "unit_factor (f * g) = f * unit_factor g" using fps_unit_factor_mult[of f g] by simp next fix f g :: "'a fps" show "unit_factor f * normalize f = f" by (simp add: fps_shift_times_fps_X_power) next fix f g :: "'a fps" show "unit_factor (f * g) = unit_factor f * unit_factor g" using fps_unit_factor_mult[of f g] by simp qed (simp_all add: fps_divide_def Let_def) end subsection \Formal power series form a Euclidean ring\ instantiation fps :: (field) euclidean_ring_cancel begin definition fps_euclidean_size_def: "euclidean_size f = (if f = 0 then 0 else 2 ^ subdegree f)" instance proof fix f g :: "'a fps" assume [simp]: "g \ 0" show "euclidean_size f \ euclidean_size (f * g)" by (cases "f = 0") (simp_all add: fps_euclidean_size_def) show "euclidean_size (f mod g) < euclidean_size g" proof (cases "f = 0") case True then show ?thesis by (simp add: fps_euclidean_size_def) next case False then show ?thesis using le_less_linear[of "subdegree g" "subdegree f"] by (force simp add: fps_mod_eq_zero fps_euclidean_size_def subdegree_mod) qed next fix f g h :: "'a fps" assume [simp]: "h \ 0" show "(h * f) div (h * g) = f div g" by (simp add: fps_divide_cancel mult.commute) show "(f + g * h) div h = g + f div h" by (simp add: fps_divide_add fps_divide_times_eq) qed (simp add: fps_euclidean_size_def) end instance fps :: (field) normalization_euclidean_semiring .. instantiation fps :: (field) euclidean_ring_gcd begin definition fps_gcd_def: "(gcd :: 'a fps \ _) = Euclidean_Algorithm.gcd" definition fps_lcm_def: "(lcm :: 'a fps \ _) = Euclidean_Algorithm.lcm" definition fps_Gcd_def: "(Gcd :: 'a fps set \ _) = Euclidean_Algorithm.Gcd" definition fps_Lcm_def: "(Lcm :: 'a fps set \ _) = Euclidean_Algorithm.Lcm" instance by standard (simp_all add: fps_gcd_def fps_lcm_def fps_Gcd_def fps_Lcm_def) end lemma fps_gcd: assumes [simp]: "f \ 0" "g \ 0" shows "gcd f g = fps_X ^ min (subdegree f) (subdegree g)" proof - let ?m = "min (subdegree f) (subdegree g)" show "gcd f g = fps_X ^ ?m" proof (rule sym, rule gcdI) fix d assume "d dvd f" "d dvd g" thus "d dvd fps_X ^ ?m" by (cases "d = 0") (simp_all add: fps_dvd_iff) qed (simp_all add: fps_dvd_iff) qed lemma fps_gcd_altdef: "gcd f g = (if f = 0 \ g = 0 then 0 else if f = 0 then fps_X ^ subdegree g else if g = 0 then fps_X ^ subdegree f else fps_X ^ min (subdegree f) (subdegree g))" by (simp add: fps_gcd) lemma fps_lcm: assumes [simp]: "f \ 0" "g \ 0" shows "lcm f g = fps_X ^ max (subdegree f) (subdegree g)" proof - let ?m = "max (subdegree f) (subdegree g)" show "lcm f g = fps_X ^ ?m" proof (rule sym, rule lcmI) fix d assume "f dvd d" "g dvd d" thus "fps_X ^ ?m dvd d" by (cases "d = 0") (simp_all add: fps_dvd_iff) qed (simp_all add: fps_dvd_iff) qed lemma fps_lcm_altdef: "lcm f g = (if f = 0 \ g = 0 then 0 else fps_X ^ max (subdegree f) (subdegree g))" by (simp add: fps_lcm) lemma fps_Gcd: assumes "A - {0} \ {}" shows "Gcd A = fps_X ^ (INF f\A-{0}. subdegree f)" proof (rule sym, rule GcdI) fix f assume "f \ A" thus "fps_X ^ (INF f\A - {0}. subdegree f) dvd f" by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cINF_lower) next fix d assume d: "\f. f \ A \ d dvd f" from assms obtain f where "f \ A - {0}" by auto with d[of f] have [simp]: "d \ 0" by auto from d assms have "subdegree d \ (INF f\A-{0}. subdegree f)" by (intro cINF_greatest) (simp_all add: fps_dvd_iff[symmetric]) with d assms show "d dvd fps_X ^ (INF f\A-{0}. subdegree f)" by (simp add: fps_dvd_iff) qed simp_all lemma fps_Gcd_altdef: "Gcd A = (if A \ {0} then 0 else fps_X ^ (INF f\A-{0}. subdegree f))" using fps_Gcd by auto lemma fps_Lcm: assumes "A \ {}" "0 \ A" "bdd_above (subdegree`A)" shows "Lcm A = fps_X ^ (SUP f\A. subdegree f)" proof (rule sym, rule LcmI) fix f assume "f \ A" moreover from assms(3) have "bdd_above (subdegree ` A)" by auto ultimately show "f dvd fps_X ^ (SUP f\A. subdegree f)" using assms(2) by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cSUP_upper) next fix d assume d: "\f. f \ A \ f dvd d" from assms obtain f where f: "f \ A" "f \ 0" by auto show "fps_X ^ (SUP f\A. subdegree f) dvd d" proof (cases "d = 0") assume "d \ 0" moreover from d have "\f. f \ A \ f \ 0 \ f dvd d" by blast ultimately have "subdegree d \ (SUP f\A. subdegree f)" using assms by (intro cSUP_least) (auto simp: fps_dvd_iff) with \d \ 0\ show ?thesis by (simp add: fps_dvd_iff) qed simp_all qed simp_all lemma fps_Lcm_altdef: "Lcm A = (if 0 \ A \ \bdd_above (subdegree`A) then 0 else if A = {} then 1 else fps_X ^ (SUP f\A. subdegree f))" proof (cases "bdd_above (subdegree`A)") assume unbounded: "\bdd_above (subdegree`A)" have "Lcm A = 0" proof (rule ccontr) assume "Lcm A \ 0" from unbounded obtain f where f: "f \ A" "subdegree (Lcm A) < subdegree f" unfolding bdd_above_def by (auto simp: not_le) moreover from f and \Lcm A \ 0\ have "subdegree f \ subdegree (Lcm A)" by (intro dvd_imp_subdegree_le dvd_Lcm) simp_all ultimately show False by simp qed with unbounded show ?thesis by simp qed (simp_all add: fps_Lcm Lcm_eq_0_I) subsection \Formal Derivatives, and the MacLaurin theorem around 0\ definition "fps_deriv f = Abs_fps (\n. of_nat (n + 1) * f $ (n + 1))" lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n + 1) * f $ (n + 1)" by (simp add: fps_deriv_def) lemma fps_0th_higher_deriv: "(fps_deriv ^^ n) f $ 0 = fact n * f $ n" by (induction n arbitrary: f) (simp_all add: funpow_Suc_right mult_of_nat_commute algebra_simps del: funpow.simps) lemma fps_deriv_mult[simp]: "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g" proof (intro fps_ext) fix n have LHS: "fps_deriv (f * g) $ n = (\i=0..Suc n. of_nat (n+1) * f$i * g$(Suc n - i))" by (simp add: fps_mult_nth sum_distrib_left algebra_simps) have "\i\{1..n}. n - (i - 1) = n - i + 1" by auto moreover have "(\i=0..n. of_nat (i+1) * f$(i+1) * g$(n - i)) = (\i=1..Suc n. of_nat i * f$i * g$(n - (i - 1)))" by (intro sum.reindex_bij_witness[where i="\x. x-1" and j="\x. x+1"]) auto ultimately have "(f * fps_deriv g + fps_deriv f * g) $ n = of_nat (Suc n) * f$0 * g$(Suc n) + (\i=1..n. (of_nat (n - i + 1) + of_nat i) * f $ i * g $ (n - i + 1)) + of_nat (Suc n) * f$(Suc n) * g$0" by (simp add: fps_mult_nth algebra_simps mult_of_nat_commute sum.atLeast_Suc_atMost sum.distrib) moreover have "\i\{1..n}. (of_nat (n - i + 1) + of_nat i) * f $ i * g $ (n - i + 1) = of_nat (n + 1) * f $ i * g $ (Suc n - i)" proof fix i assume i: "i \ {1..n}" from i have "of_nat (n - i + 1) + (of_nat i :: 'a) = of_nat (n + 1)" using of_nat_add[of "n-i+1" i,symmetric] by simp moreover from i have "Suc n - i = n - i + 1" by auto ultimately show "(of_nat (n - i + 1) + of_nat i) * f $ i * g $ (n - i + 1) = of_nat (n + 1) * f $ i * g $ (Suc n - i)" by simp qed ultimately have "(f * fps_deriv g + fps_deriv f * g) $ n = (\i=0..Suc n. of_nat (Suc n) * f $ i * g $ (Suc n - i))" by (simp add: sum.atLeast_Suc_atMost) with LHS show "fps_deriv (f * g) $ n = (f * fps_deriv g + fps_deriv f * g) $ n" by simp qed lemma fps_deriv_fps_X[simp]: "fps_deriv fps_X = 1" by (simp add: fps_deriv_def fps_X_def fps_eq_iff) lemma fps_deriv_neg[simp]: "fps_deriv (- (f:: 'a::ring_1 fps)) = - (fps_deriv f)" by (simp add: fps_eq_iff fps_deriv_def) lemma fps_deriv_add[simp]: "fps_deriv (f + g) = fps_deriv f + fps_deriv g" by (auto intro: fps_ext simp: algebra_simps) lemma fps_deriv_sub[simp]: "fps_deriv ((f:: 'a::ring_1 fps) - g) = fps_deriv f - fps_deriv g" using fps_deriv_add [of f "- g"] by simp lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0" by (simp add: fps_ext fps_deriv_def fps_const_def) lemma fps_deriv_of_nat [simp]: "fps_deriv (of_nat n) = 0" by (simp add: fps_of_nat [symmetric]) lemma fps_deriv_of_int [simp]: "fps_deriv (of_int n) = 0" by (simp add: fps_of_int [symmetric]) lemma fps_deriv_numeral [simp]: "fps_deriv (numeral n) = 0" by (simp add: numeral_fps_const) lemma fps_deriv_mult_const_left[simp]: "fps_deriv (fps_const c * f) = fps_const c * fps_deriv f" by simp lemma fps_deriv_linear[simp]: "fps_deriv (fps_const a * f + fps_const b * g) = fps_const a * fps_deriv f + fps_const b * fps_deriv g" by simp lemma fps_deriv_0[simp]: "fps_deriv 0 = 0" by (simp add: fps_deriv_def fps_eq_iff) lemma fps_deriv_1[simp]: "fps_deriv 1 = 0" by (simp add: fps_deriv_def fps_eq_iff) lemma fps_deriv_mult_const_right[simp]: "fps_deriv (f * fps_const c) = fps_deriv f * fps_const c" by simp lemma fps_deriv_sum: "fps_deriv (sum f S) = sum (\i. fps_deriv (f i)) S" proof (cases "finite S") case False then show ?thesis by simp next case True show ?thesis by (induct rule: finite_induct [OF True]) simp_all qed lemma fps_deriv_eq_0_iff [simp]: "fps_deriv f = 0 \ f = fps_const (f$0 :: 'a::{semiring_no_zero_divisors,semiring_char_0})" proof assume f: "fps_deriv f = 0" show "f = fps_const (f$0)" proof (intro fps_ext) fix n show "f $ n = fps_const (f$0) $ n" proof (cases n) case (Suc m) have "(of_nat (Suc m) :: 'a) \ 0" by (rule of_nat_neq_0) with f Suc show ?thesis using fps_deriv_nth[of f] by auto qed simp qed next show "f = fps_const (f$0) \ fps_deriv f = 0" using fps_deriv_const[of "f$0"] by simp qed lemma fps_deriv_eq_iff: fixes f g :: "'a::{ring_1_no_zero_divisors,semiring_char_0} fps" shows "fps_deriv f = fps_deriv g \ (f = fps_const(f$0 - g$0) + g)" proof - have "fps_deriv f = fps_deriv g \ fps_deriv (f - g) = 0" using fps_deriv_sub[of f g] by simp also have "\ \ f - g = fps_const ((f - g) $ 0)" unfolding fps_deriv_eq_0_iff .. finally show ?thesis by (simp add: field_simps) qed lemma fps_deriv_eq_iff_ex: fixes f g :: "'a::{ring_1_no_zero_divisors,semiring_char_0} fps" shows "(fps_deriv f = fps_deriv g) \ (\c. f = fps_const c + g)" by (auto simp: fps_deriv_eq_iff) fun fps_nth_deriv :: "nat \ 'a::semiring_1 fps \ 'a fps" where "fps_nth_deriv 0 f = f" | "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)" lemma fps_nth_deriv_commute: "fps_nth_deriv (Suc n) f = fps_deriv (fps_nth_deriv n f)" by (induct n arbitrary: f) auto lemma fps_nth_deriv_linear[simp]: "fps_nth_deriv n (fps_const a * f + fps_const b * g) = fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g" by (induct n arbitrary: f g) auto lemma fps_nth_deriv_neg[simp]: "fps_nth_deriv n (- (f :: 'a::ring_1 fps)) = - (fps_nth_deriv n f)" by (induct n arbitrary: f) simp_all lemma fps_nth_deriv_add[simp]: "fps_nth_deriv n ((f :: 'a::ring_1 fps) + g) = fps_nth_deriv n f + fps_nth_deriv n g" using fps_nth_deriv_linear[of n 1 f 1 g] by simp lemma fps_nth_deriv_sub[simp]: "fps_nth_deriv n ((f :: 'a::ring_1 fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g" using fps_nth_deriv_add [of n f "- g"] by simp lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0" by (induct n) simp_all lemma fps_nth_deriv_1[simp]: "fps_nth_deriv n 1 = (if n = 0 then 1 else 0)" by (induct n) simp_all lemma fps_nth_deriv_const[simp]: "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)" by (cases n) simp_all lemma fps_nth_deriv_mult_const_left[simp]: "fps_nth_deriv n (fps_const c * f) = fps_const c * fps_nth_deriv n f" using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp lemma fps_nth_deriv_mult_const_right[simp]: "fps_nth_deriv n (f * fps_const c) = fps_nth_deriv n f * fps_const c" by (induct n arbitrary: f) auto lemma fps_nth_deriv_sum: "fps_nth_deriv n (sum f S) = sum (\i. fps_nth_deriv n (f i :: 'a::ring_1 fps)) S" proof (cases "finite S") case True show ?thesis by (induct rule: finite_induct [OF True]) simp_all next case False then show ?thesis by simp qed lemma fps_deriv_maclauren_0: "(fps_nth_deriv k (f :: 'a::comm_semiring_1 fps)) $ 0 = of_nat (fact k) * f $ k" by (induct k arbitrary: f) (simp_all add: field_simps) lemma fps_deriv_lr_inverse: fixes x y :: "'a::ring_1" assumes "x * f$0 = 1" "f$0 * y = 1" \ \These assumptions imply x equals y, but no need to assume that.\ shows "fps_deriv (fps_left_inverse f x) = - fps_left_inverse f x * fps_deriv f * fps_left_inverse f x" and "fps_deriv (fps_right_inverse f y) = - fps_right_inverse f y * fps_deriv f * fps_right_inverse f y" proof- define L where "L \ fps_left_inverse f x" hence "fps_deriv (L * f) = 0" using fps_left_inverse[OF assms(1)] by simp with assms show "fps_deriv L = - L * fps_deriv f * L" using fps_right_inverse'[OF assms] by (simp add: minus_unique mult.assoc L_def) define R where "R \ fps_right_inverse f y" hence "fps_deriv (f * R) = 0" using fps_right_inverse[OF assms(2)] by simp hence 1: "f * fps_deriv R + fps_deriv f * R = 0" by simp have "R * f * fps_deriv R = - R * fps_deriv f * R" using iffD2[OF eq_neg_iff_add_eq_0, OF 1] by (simp add: mult.assoc) thus "fps_deriv R = - R * fps_deriv f * R" using fps_left_inverse'[OF assms] by (simp add: R_def) qed lemma fps_deriv_lr_inverse_comm: fixes x :: "'a::comm_ring_1" assumes "x * f$0 = 1" shows "fps_deriv (fps_left_inverse f x) = - fps_deriv f * (fps_left_inverse f x)\<^sup>2" and "fps_deriv (fps_right_inverse f x) = - fps_deriv f * (fps_right_inverse f x)\<^sup>2" using assms fps_deriv_lr_inverse[of x f x] by (simp_all add: mult.commute power2_eq_square) lemma fps_inverse_deriv_divring: fixes a :: "'a::division_ring fps" assumes "a$0 \ 0" shows "fps_deriv (inverse a) = - inverse a * fps_deriv a * inverse a" using assms fps_deriv_lr_inverse(2)[of "inverse (a$0)" a "inverse (a$0)"] by (simp add: fps_inverse_def) lemma fps_inverse_deriv: fixes a :: "'a::field fps" assumes "a$0 \ 0" shows "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2" using assms fps_deriv_lr_inverse_comm(2)[of "inverse (a$0)" a] by (simp add: fps_inverse_def) lemma fps_inverse_deriv': fixes a :: "'a::field fps" assumes a0: "a $ 0 \ 0" shows "fps_deriv (inverse a) = - fps_deriv a / a\<^sup>2" using fps_inverse_deriv[OF a0] a0 by (simp add: fps_divide_unit power2_eq_square fps_inverse_mult) (* FIXME: The last part of this proof should go through by simp once we have a proper theorem collection for simplifying division on rings *) lemma fps_divide_deriv: assumes "b dvd (a :: 'a :: field fps)" shows "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b^2" proof - have eq_divide_imp: "c \ 0 \ a * c = b \ a = b div c" for a b c :: "'a :: field fps" by (drule sym) (simp add: mult.assoc) from assms have "a = a / b * b" by simp also have "fps_deriv (a / b * b) = fps_deriv (a / b) * b + a / b * fps_deriv b" by simp finally have "fps_deriv (a / b) * b^2 = fps_deriv a * b - a * fps_deriv b" using assms by (simp add: power2_eq_square algebra_simps) thus ?thesis by (cases "b = 0") (simp_all add: eq_divide_imp) qed lemma fps_nth_deriv_fps_X[simp]: "fps_nth_deriv n fps_X = (if n = 0 then fps_X else if n=1 then 1 else 0)" by (cases n) simp_all subsection \Powers\ lemma fps_power_zeroth: "(a^n) $ 0 = (a$0)^n" by (induct n) auto lemma fps_power_zeroth_eq_one: "a$0 = 1 \ a^n $ 0 = 1" by (simp add: fps_power_zeroth) lemma fps_power_first: fixes a :: "'a::comm_semiring_1 fps" shows "(a^n) $ 1 = of_nat n * (a$0)^(n-1) * a$1" proof (cases n) case (Suc m) have "(a ^ Suc m) $ 1 = of_nat (Suc m) * (a$0)^(Suc m - 1) * a$1" proof (induct m) case (Suc k) hence "(a ^ Suc (Suc k)) $ 1 = a$0 * of_nat (Suc k) * (a $ 0)^k * a$1 + a$1 * ((a$0)^(Suc k))" using fps_mult_nth_1[of a] by (simp add: fps_power_zeroth[symmetric] mult.assoc) thus ?case by (simp add: algebra_simps) qed simp with Suc show ?thesis by simp qed simp lemma fps_power_first_eq: "a $ 0 = 1 \ a^n $ 1 = of_nat n * a$1" proof (induct n) case (Suc n) show ?case unfolding power_Suc fps_mult_nth using Suc.hyps[OF \a$0 = 1\] \a$0 = 1\ fps_power_zeroth_eq_one[OF \a$0=1\] by (simp add: algebra_simps) qed simp lemma fps_power_first_eq': assumes "a $ 1 = 1" shows "a^n $ 1 = of_nat n * (a$0)^(n-1)" proof (cases n) case (Suc m) from assms have "(a ^ Suc m) $ 1 = of_nat (Suc m) * (a$0)^(Suc m - 1)" using fps_mult_nth_1[of a] by (induct m) (simp_all add: algebra_simps mult_of_nat_commute fps_power_zeroth) with Suc show ?thesis by simp qed simp lemmas startsby_one_power = fps_power_zeroth_eq_one lemma startsby_zero_power: "a $ 0 = 0 \ n > 0 \ a^n $0 = 0" by (simp add: fps_power_zeroth zero_power) lemma startsby_power: "a $0 = v \ a^n $0 = v^n" by (simp add: fps_power_zeroth) lemma startsby_nonzero_power: fixes a :: "'a::semiring_1_no_zero_divisors fps" shows "a $ 0 \ 0 \ a^n $ 0 \ 0" by (simp add: startsby_power) lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::semiring_1_no_zero_divisors) \ n \ 0 \ a$0 = 0" proof show "a ^ n $ 0 = 0 \ n \ 0 \ a $ 0 = 0" proof assume a: "a^n $ 0 = 0" thus "a $ 0 = 0" using startsby_nonzero_power by auto have "n = 0 \ a^n $ 0 = 1" by simp with a show "n \ 0" by fastforce qed show "n \ 0 \ a $ 0 = 0 \ a ^ n $ 0 = 0" by (cases n) auto qed lemma startsby_zero_power_prefix: assumes a0: "a $ 0 = 0" shows "\n < k. a ^ k $ n = 0" proof (induct k rule: nat_less_induct, clarify) case (1 k) fix j :: nat assume j: "j < k" show "a ^ k $ j = 0" proof (cases k) case 0 with j show ?thesis by simp next case (Suc i) with 1 j have "\m\{0<..j}. a ^ i $ (j - m) = 0" by auto with Suc a0 show ?thesis by (simp add: fps_mult_nth sum.atLeast_Suc_atMost) qed qed lemma startsby_zero_sum_depends: assumes a0: "a $0 = 0" and kn: "n \ k" shows "sum (\i. (a ^ i)$k) {0 .. n} = sum (\i. (a ^ i)$k) {0 .. k}" proof (intro strip sum.mono_neutral_right) show "\i. i \ {0..n} - {0..k} \ a ^ i $ k = 0" by (simp add: a0 startsby_zero_power_prefix) qed (use kn in auto) lemma startsby_zero_power_nth_same: assumes a0: "a$0 = 0" shows "a^n $ n = (a$1) ^ n" proof (induct n) case (Suc n) have "\i\{Suc 1..Suc n}. a ^ n $ (Suc n - i) = 0" using a0 startsby_zero_power_prefix[of a n] by auto thus ?case using a0 Suc sum.atLeast_Suc_atMost[of 0 "Suc n" "\i. a $ i * a ^ n $ (Suc n - i)"] sum.atLeast_Suc_atMost[of 1 "Suc n" "\i. a $ i * a ^ n $ (Suc n - i)"] by (simp add: fps_mult_nth) qed simp lemma fps_lr_inverse_power: fixes a :: "'a::ring_1 fps" assumes "x * a$0 = 1" "a$0 * x = 1" shows "fps_left_inverse (a^n) (x^n) = fps_left_inverse a x ^ n" and "fps_right_inverse (a^n) (x^n) = fps_right_inverse a x ^ n" proof- from assms have xn: "\n. x^n * (a^n $ 0) = 1" "\n. (a^n $ 0) * x^n = 1" by (simp_all add: left_right_inverse_power fps_power_zeroth) show "fps_left_inverse (a^n) (x^n) = fps_left_inverse a x ^ n" proof (induct n) case 0 then show ?case by (simp add: fps_lr_inverse_one_one(1)) next case (Suc n) with assms show ?case using xn fps_lr_inverse_mult_ring1(1)[of x a "x^n" "a^n"] by (simp add: power_Suc2[symmetric]) qed moreover have "fps_right_inverse (a^n) (x^n) = fps_left_inverse (a^n) (x^n)" using xn by (intro fps_left_inverse_eq_fps_right_inverse[symmetric]) moreover have "fps_right_inverse a x = fps_left_inverse a x" using assms by (intro fps_left_inverse_eq_fps_right_inverse[symmetric]) ultimately show "fps_right_inverse (a^n) (x^n) = fps_right_inverse a x ^ n" by simp qed lemma fps_inverse_power: fixes a :: "'a::division_ring fps" shows "inverse (a^n) = inverse a ^ n" proof (cases "n=0" "a$0 = 0" rule: case_split[case_product case_split]) case False_True hence LHS: "inverse (a^n) = 0" and RHS: "inverse a ^ n = 0" by (simp_all add: startsby_zero_power) show ?thesis using trans_sym[OF LHS RHS] by fast next case False_False from False_False(2) show ?thesis by (simp add: fps_inverse_def fps_power_zeroth power_inverse fps_lr_inverse_power(2)[symmetric] ) qed auto lemma fps_deriv_power': fixes a :: "'a::comm_semiring_1 fps" shows "fps_deriv (a ^ n) = (of_nat n) * fps_deriv a * a ^ (n - 1)" proof (cases n) case (Suc m) moreover have "fps_deriv (a^Suc m) = of_nat (Suc m) * fps_deriv a * a^m" by (induct m) (simp_all add: algebra_simps) ultimately show ?thesis by simp qed simp lemma fps_deriv_power: fixes a :: "'a::comm_semiring_1 fps" shows "fps_deriv (a ^ n) = fps_const (of_nat n) * fps_deriv a * a ^ (n - 1)" by (simp add: fps_deriv_power' fps_of_nat) subsection \Integration\ definition fps_integral :: "'a::{semiring_1,inverse} fps \ 'a \ 'a fps" where "fps_integral a a0 = Abs_fps (\n. if n=0 then a0 else inverse (of_nat n) * a$(n - 1))" abbreviation "fps_integral0 a \ fps_integral a 0" lemma fps_integral_nth_0_Suc [simp]: fixes a :: "'a::{semiring_1,inverse} fps" shows "fps_integral a a0 $ 0 = a0" and "fps_integral a a0 $ Suc n = inverse (of_nat (Suc n)) * a $ n" by (auto simp: fps_integral_def) lemma fps_integral_conv_plus_const: "fps_integral a a0 = fps_integral a 0 + fps_const a0" unfolding fps_integral_def by (intro fps_ext) simp lemma fps_deriv_fps_integral: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_deriv (fps_integral a a0) = a" proof (intro fps_ext) fix n have "(of_nat (Suc n) :: 'a) \ 0" by (rule of_nat_neq_0) hence "of_nat (Suc n) * inverse (of_nat (Suc n) :: 'a) = 1" by simp moreover have "fps_deriv (fps_integral a a0) $ n = of_nat (Suc n) * inverse (of_nat (Suc n)) * a $ n" by (simp add: mult.assoc) ultimately show "fps_deriv (fps_integral a a0) $ n = a $ n" by simp qed lemma fps_integral0_deriv: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_integral0 (fps_deriv a) = a - fps_const (a$0)" proof (intro fps_ext) fix n show "fps_integral0 (fps_deriv a) $ n = (a - fps_const (a$0)) $ n" proof (cases n) case (Suc m) have "(of_nat (Suc m) :: 'a) \ 0" by (rule of_nat_neq_0) hence "inverse (of_nat (Suc m) :: 'a) * of_nat (Suc m) = 1" by simp moreover have "fps_integral0 (fps_deriv a) $ Suc m = inverse (of_nat (Suc m)) * of_nat (Suc m) * a $ (Suc m)" by (simp add: mult.assoc) ultimately show ?thesis using Suc by simp qed simp qed lemma fps_integral_deriv: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_integral (fps_deriv a) (a$0) = a" using fps_integral_conv_plus_const[of "fps_deriv a" "a$0"] by (simp add: fps_integral0_deriv) lemma fps_integral0_zero: "fps_integral0 (0::'a::{semiring_1,inverse} fps) = 0" by (intro fps_ext) (simp add: fps_integral_def) lemma fps_integral0_fps_const': fixes c :: "'a::{semiring_1,inverse}" assumes "inverse (1::'a) = 1" shows "fps_integral0 (fps_const c) = fps_const c * fps_X" proof (intro fps_ext) fix n show "fps_integral0 (fps_const c) $ n = (fps_const c * fps_X) $ n" by (cases n) (simp_all add: assms mult_delta_right) qed lemma fps_integral0_fps_const: fixes c :: "'a::division_ring" shows "fps_integral0 (fps_const c) = fps_const c * fps_X" by (rule fps_integral0_fps_const'[OF inverse_1]) lemma fps_integral0_one': assumes "inverse (1::'a::{semiring_1,inverse}) = 1" shows "fps_integral0 (1::'a fps) = fps_X" using assms fps_integral0_fps_const'[of "1::'a"] by simp lemma fps_integral0_one: "fps_integral0 (1::'a::division_ring fps) = fps_X" by (rule fps_integral0_one'[OF inverse_1]) lemma fps_integral0_fps_const_mult_left: fixes a :: "'a::division_ring fps" shows "fps_integral0 (fps_const c * a) = fps_const c * fps_integral0 a" proof (intro fps_ext) fix n show "fps_integral0 (fps_const c * a) $ n = (fps_const c * fps_integral0 a) $ n" using mult_inverse_of_nat_commute[of n c, symmetric] mult.assoc[of "inverse (of_nat n)" c "a$(n-1)"] mult.assoc[of c "inverse (of_nat n)" "a$(n-1)"] by (simp add: fps_integral_def) qed lemma fps_integral0_fps_const_mult_right: fixes a :: "'a::{semiring_1,inverse} fps" shows "fps_integral0 (a * fps_const c) = fps_integral0 a * fps_const c" by (intro fps_ext) (simp add: fps_integral_def algebra_simps) lemma fps_integral0_neg: fixes a :: "'a::{ring_1,inverse} fps" shows "fps_integral0 (-a) = - fps_integral0 a" using fps_integral0_fps_const_mult_right[of a "-1"] by (simp add: fps_const_neg[symmetric]) lemma fps_integral0_add: "fps_integral0 (a+b) = fps_integral0 a + fps_integral0 b" by (intro fps_ext) (simp add: fps_integral_def algebra_simps) lemma fps_integral0_linear: fixes a b :: "'a::division_ring" shows "fps_integral0 (fps_const a * f + fps_const b * g) = fps_const a * fps_integral0 f + fps_const b * fps_integral0 g" by (simp add: fps_integral0_add fps_integral0_fps_const_mult_left) lemma fps_integral0_linear2: "fps_integral0 (f * fps_const a + g * fps_const b) = fps_integral0 f * fps_const a + fps_integral0 g * fps_const b" by (simp add: fps_integral0_add fps_integral0_fps_const_mult_right) lemma fps_integral_linear: fixes a b a0 b0 :: "'a::division_ring" shows "fps_integral (fps_const a * f + fps_const b * g) (a*a0 + b*b0) = fps_const a * fps_integral f a0 + fps_const b * fps_integral g b0" using fps_integral_conv_plus_const[of "fps_const a * f + fps_const b * g" "a*a0 + b*b0" ] fps_integral_conv_plus_const[of f a0] fps_integral_conv_plus_const[of g b0] by (simp add: fps_integral0_linear algebra_simps) lemma fps_integral0_sub: fixes a b :: "'a::{ring_1,inverse} fps" shows "fps_integral0 (a-b) = fps_integral0 a - fps_integral0 b" using fps_integral0_linear2[of a 1 b "-1"] by (simp add: fps_const_neg[symmetric]) lemma fps_integral0_of_nat: "fps_integral0 (of_nat n :: 'a::division_ring fps) = of_nat n * fps_X" using fps_integral0_fps_const[of "of_nat n :: 'a"] by (simp add: fps_of_nat) lemma fps_integral0_sum: "fps_integral0 (sum f S) = sum (\i. fps_integral0 (f i)) S" proof (cases "finite S") case True show ?thesis by (induct rule: finite_induct [OF True]) (simp_all add: fps_integral0_zero fps_integral0_add) qed (simp add: fps_integral0_zero) lemma fps_integral0_by_parts: fixes a b :: "'a::{division_ring,ring_char_0} fps" shows "fps_integral0 (a * b) = a * fps_integral0 b - fps_integral0 (fps_deriv a * fps_integral0 b)" proof- have "fps_integral0 (fps_deriv (a * fps_integral0 b)) = a * fps_integral0 b" using fps_integral0_deriv[of "(a * fps_integral0 b)"] by simp moreover have "fps_integral0 (a * b) = fps_integral0 (fps_deriv (a * fps_integral0 b)) - fps_integral0 (fps_deriv a * fps_integral0 b)" by (auto simp: fps_deriv_fps_integral fps_integral0_sub[symmetric]) ultimately show ?thesis by simp qed lemma fps_integral0_fps_X: "fps_integral0 (fps_X::'a::{semiring_1,inverse} fps) = fps_const (inverse (of_nat 2)) * fps_X\<^sup>2" by (intro fps_ext) (auto simp: fps_integral_def) lemma fps_integral0_fps_X_power: "fps_integral0 ((fps_X::'a::{semiring_1,inverse} fps) ^ n) = fps_const (inverse (of_nat (Suc n))) * fps_X ^ Suc n" proof (intro fps_ext) fix k show "fps_integral0 ((fps_X::'a fps) ^ n) $ k = (fps_const (inverse (of_nat (Suc n))) * fps_X ^ Suc n) $ k" by (cases k) simp_all qed subsection \Composition of FPSs\ definition fps_compose :: "'a::semiring_1 fps \ 'a fps \ 'a fps" (infixl "oo" 55) where "a oo b = Abs_fps (\n. sum (\i. a$i * (b^i$n)) {0..n})" lemma fps_compose_nth: "(a oo b)$n = sum (\i. a$i * (b^i$n)) {0..n}" by (simp add: fps_compose_def) lemma fps_compose_nth_0 [simp]: "(f oo g) $ 0 = f $ 0" by (simp add: fps_compose_nth) lemma fps_compose_fps_X[simp]: "a oo fps_X = (a :: 'a::comm_ring_1 fps)" by (simp add: fps_ext fps_compose_def mult_delta_right) lemma fps_const_compose[simp]: "fps_const (a::'a::comm_ring_1) oo b = fps_const a" by (simp add: fps_eq_iff fps_compose_nth mult_delta_left) lemma numeral_compose[simp]: "(numeral k :: 'a::comm_ring_1 fps) oo b = numeral k" unfolding numeral_fps_const by simp lemma neg_numeral_compose[simp]: "(- numeral k :: 'a::comm_ring_1 fps) oo b = - numeral k" unfolding neg_numeral_fps_const by simp lemma fps_X_fps_compose_startby0[simp]: "a$0 = 0 \ fps_X oo a = (a :: 'a::comm_ring_1 fps)" by (simp add: fps_eq_iff fps_compose_def mult_delta_left not_le) subsection \Rules from Herbert Wilf's Generatingfunctionology\ subsubsection \Rule 1\ (* {a_{n+k}}_0^infty Corresponds to (f - sum (\i. a_i * x^i))/x^h, for h>0*) lemma fps_power_mult_eq_shift: "fps_X^Suc k * Abs_fps (\n. a (n + Suc k)) = Abs_fps a - sum (\i. fps_const (a i :: 'a::comm_ring_1) * fps_X^i) {0 .. k}" (is "?lhs = ?rhs") proof - have "?lhs $ n = ?rhs $ n" for n :: nat proof - have "?lhs $ n = (if n < Suc k then 0 else a n)" unfolding fps_X_power_mult_nth by auto also have "\ = ?rhs $ n" proof (induct k) case 0 then show ?case by (simp add: fps_sum_nth) next case (Suc k) have "(Abs_fps a - sum (\i. fps_const (a i :: 'a) * fps_X^i) {0 .. Suc k})$n = (Abs_fps a - sum (\i. fps_const (a i :: 'a) * fps_X^i) {0 .. k} - fps_const (a (Suc k)) * fps_X^ Suc k) $ n" by (simp add: field_simps) also have "\ = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * fps_X^ Suc k)$n" using Suc.hyps[symmetric] unfolding fps_sub_nth by simp also have "\ = (if n < Suc (Suc k) then 0 else a n)" unfolding fps_X_power_mult_right_nth by (simp add: not_less le_less_Suc_eq) finally show ?case by simp qed finally show ?thesis . qed then show ?thesis by (simp add: fps_eq_iff) qed subsubsection \Rule 2\ (* We can not reach the form of Wilf, but still near to it using rewrite rules*) (* If f reprents {a_n} and P is a polynomial, then P(xD) f represents {P(n) a_n}*) definition "fps_XD = (*) fps_X \ fps_deriv" lemma fps_XD_add[simp]:"fps_XD (a + b) = fps_XD a + fps_XD (b :: 'a::comm_ring_1 fps)" by (simp add: fps_XD_def field_simps) lemma fps_XD_mult_const[simp]:"fps_XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * fps_XD a" by (simp add: fps_XD_def field_simps) lemma fps_XD_linear[simp]: "fps_XD (fps_const c * a + fps_const d * b) = fps_const c * fps_XD a + fps_const d * fps_XD (b :: 'a::comm_ring_1 fps)" by simp lemma fps_XDN_linear: "(fps_XD ^^ n) (fps_const c * a + fps_const d * b) = fps_const c * (fps_XD ^^ n) a + fps_const d * (fps_XD ^^ n) (b :: 'a::comm_ring_1 fps)" by (induct n) simp_all lemma fps_mult_fps_X_deriv_shift: "fps_X* fps_deriv a = Abs_fps (\n. of_nat n* a$n)" by (simp add: fps_eq_iff) lemma fps_mult_fps_XD_shift: "(fps_XD ^^ k) (a :: 'a::comm_ring_1 fps) = Abs_fps (\n. (of_nat n ^ k) * a$n)" by (induct k arbitrary: a) (simp_all add: fps_XD_def fps_eq_iff field_simps del: One_nat_def) subsubsection \Rule 3\ text \Rule 3 is trivial and is given by \fps_times_def\.\ subsubsection \Rule 5 --- summation and "division" by (1 - fps_X)\ lemma fps_divide_fps_X_minus1_sum_lemma: "a = ((1::'a::ring_1 fps) - fps_X) * Abs_fps (\n. sum (\i. a $ i) {0..n})" proof (rule fps_ext) define f g :: "'a fps" where "f \ 1 - fps_X" and "g \ Abs_fps (\n. sum (\i. a $ i) {0..n})" fix n show "a $ n= (f * g) $ n" proof (cases n) case (Suc m) hence "(f * g) $ n = g $ Suc m - g $ m" using fps_mult_nth[of f g "Suc m"] sum.atLeast_Suc_atMost[of 0 "Suc m" "\i. f $ i * g $ (Suc m - i)"] sum.atLeast_Suc_atMost[of 1 "Suc m" "\i. f $ i * g $ (Suc m - i)"] by (simp add: f_def) with Suc show ?thesis by (simp add: g_def) qed (simp add: f_def g_def) qed lemma fps_divide_fps_X_minus1_sum_ring1: assumes "inverse 1 = (1::'a::{ring_1,inverse})" shows "a /((1::'a fps) - fps_X) = Abs_fps (\n. sum (\i. a $ i) {0..n})" proof- from assms have "a /((1::'a fps) - fps_X) = a * Abs_fps (\n. 1)" by (simp add: fps_divide_def fps_inverse_def fps_lr_inverse_one_minus_fps_X(2)) thus ?thesis by (auto intro: fps_ext simp: fps_mult_nth) qed lemma fps_divide_fps_X_minus1_sum: "a /((1::'a::division_ring fps) - fps_X) = Abs_fps (\n. sum (\i. a $ i) {0..n})" using fps_divide_fps_X_minus1_sum_ring1[of a] by simp subsubsection \Rule 4 in its more general form: generalizes Rule 3 for an arbitrary finite product of FPS, also the relvant instance of powers of a FPS\ definition "natpermute n k = {l :: nat list. length l = k \ sum_list l = n}" lemma natlist_trivial_1: "natpermute n 1 = {[n]}" proof - have "\length xs = 1; n = sum_list xs\ \ xs = [sum_list xs]" for xs by (cases xs) auto then show ?thesis by (auto simp add: natpermute_def) qed lemma natlist_trivial_Suc0 [simp]: "natpermute n (Suc 0) = {[n]}" using natlist_trivial_1 by force lemma append_natpermute_less_eq: assumes "xs @ ys \ natpermute n k" shows "sum_list xs \ n" and "sum_list ys \ n" proof - from assms have "sum_list (xs @ ys) = n" by (simp add: natpermute_def) then have "sum_list xs + sum_list ys = n" by simp then show "sum_list xs \ n" and "sum_list ys \ n" by simp_all qed lemma natpermute_split: assumes "h \ k" shows "natpermute n k = (\m \{0..n}. {l1 @ l2 |l1 l2. l1 \ natpermute m h \ l2 \ natpermute (n - m) (k - h)})" (is "?L = ?R" is "_ = (\m \{0..n}. ?S m)") proof show "?R \ ?L" proof fix l assume l: "l \ ?R" from l obtain m xs ys where h: "m \ {0..n}" and xs: "xs \ natpermute m h" and ys: "ys \ natpermute (n - m) (k - h)" and leq: "l = xs@ys" by blast from xs have xs': "sum_list xs = m" by (simp add: natpermute_def) from ys have ys': "sum_list ys = n - m" by (simp add: natpermute_def) show "l \ ?L" using leq xs ys h using assms by (force simp add: natpermute_def) qed show "?L \ ?R" proof fix l assume l: "l \ natpermute n k" let ?xs = "take h l" let ?ys = "drop h l" let ?m = "sum_list ?xs" from l have ls: "sum_list (?xs @ ?ys) = n" by (simp add: natpermute_def) have xs: "?xs \ natpermute ?m h" using l assms by (simp add: natpermute_def) have l_take_drop: "sum_list l = sum_list (take h l @ drop h l)" by simp then have ys: "?ys \ natpermute (n - ?m) (k - h)" using l assms ls by (auto simp add: natpermute_def simp del: append_take_drop_id) from ls have m: "?m \ {0..n}" by (simp add: l_take_drop del: append_take_drop_id) have "sum_list (take h l) \ sum_list l" using l_take_drop ls m by presburger with xs ys ls l show "l \ ?R" by simp (metis append_take_drop_id m) qed qed lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})" by (auto simp add: natpermute_def) lemma natpermute_0'[simp]: "natpermute 0 k = (if k = 0 then {[]} else {replicate k 0})" by (auto simp add: set_replicate_conv_if natpermute_def replicate_length_same) lemma natpermute_finite: "finite (natpermute n k)" proof (induct k arbitrary: n) case 0 then show ?case by (simp add: natpermute_0) next case (Suc k) then show ?case using natpermute_split [of k "Suc k"] finite_UN_I by simp qed lemma natpermute_contain_maximal: "{xs \ natpermute n (k + 1). n \ set xs} = (\i\{0 .. k}. {(replicate (k + 1) 0) [i:=n]})" (is "?A = ?B") proof show "?A \ ?B" proof fix xs assume "xs \ ?A" then have H: "xs \ natpermute n (k + 1)" and n: "n \ set xs" by blast+ then obtain i where i: "i \ {0.. k}" "xs!i = n" unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def) have eqs: "({0..k} - {i}) \ {i} = {0..k}" using i by auto have f: "finite({0..k} - {i})" "finite {i}" by auto have d: "({0..k} - {i}) \ {i} = {}" using i by auto from H have "n = sum (nth xs) {0..k}" by (auto simp add: natpermute_def atLeastLessThanSuc_atLeastAtMost sum_list_sum_nth) also have "\ = n + sum (nth xs) ({0..k} - {i})" unfolding sum.union_disjoint[OF f d, unfolded eqs] using i by simp finally have zxs: "\ j\ {0..k} - {i}. xs!j = 0" by auto from H have xsl: "length xs = k+1" by (simp add: natpermute_def) from i have i': "i < length (replicate (k+1) 0)" "i < k+1" unfolding length_replicate by presburger+ have "xs = (replicate (k+1) 0) [i := n]" proof (rule nth_equalityI) show "length xs = length ((replicate (k + 1) 0)[i := n])" by (metis length_list_update length_replicate xsl) show "xs ! j = (replicate (k + 1) 0)[i := n] ! j" if "j < length xs" for j proof (cases "j = i") case True then show ?thesis by (metis i'(1) i(2) nth_list_update) next case False with that show ?thesis by (simp add: xsl zxs del: replicate.simps split: nat.split) qed qed then show "xs \ ?B" using i by blast qed show "?B \ ?A" proof fix xs assume "xs \ ?B" then obtain i where i: "i \ {0..k}" and xs: "xs = (replicate (k + 1) 0) [i:=n]" by auto have nxs: "n \ set xs" unfolding xs using set_update_memI i by (metis Suc_eq_plus1 atLeast0AtMost atMost_iff le_simps(2) length_replicate) have xsl: "length xs = k + 1" by (simp only: xs length_replicate length_list_update) have "sum_list xs = sum (nth xs) {0.. = sum (\j. if j = i then n else 0) {0..< k+1}" by (rule sum.cong) (simp_all add: xs del: replicate.simps) also have "\ = n" using i by simp finally have "xs \ natpermute n (k + 1)" using xsl unfolding natpermute_def mem_Collect_eq by blast then show "xs \ ?A" using nxs by blast qed qed text \The general form.\ lemma fps_prod_nth: fixes m :: nat and a :: "nat \ 'a::comm_ring_1 fps" shows "(prod a {0 .. m}) $ n = sum (\v. prod (\j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))" (is "?P m n") proof (induct m arbitrary: n rule: nat_less_induct) fix m n assume H: "\m' < m. \n. ?P m' n" show "?P m n" proof (cases m) case 0 then show ?thesis by simp next case (Suc k) then have km: "k < m" by arith have u0: "{0 .. k} \ {m} = {0..m}" using Suc by (simp add: set_eq_iff) presburger have f0: "finite {0 .. k}" "finite {m}" by auto have d0: "{0 .. k} \ {m} = {}" using Suc by auto have "(prod a {0 .. m}) $ n = (prod a {0 .. k} * a m) $ n" unfolding prod.union_disjoint[OF f0 d0, unfolded u0] by simp also have "\ = (\i = 0..n. (\v\natpermute i (k + 1). (\j = 0..k. a j $ v ! j) * a m $ (n - i)))" unfolding fps_mult_nth H[rule_format, OF km] sum_distrib_right .. also have "... = (\i = 0..n. \v\(\l1. l1 @ [n - i]) ` natpermute i (Suc k). (\j = 0..k. a j $ v ! j) * a (Suc k) $ v ! Suc k)" by (intro sum.cong [OF refl sym] sum.reindex_cong) (auto simp: inj_on_def natpermute_def nth_append Suc) also have "... = (\v\(\x\{0..n}. {l1 @ [n - x] |l1. l1 \ natpermute x (Suc k)}). (\j = 0..k. a j $ v ! j) * a (Suc k) $ v ! Suc k)" by (subst sum.UNION_disjoint) (auto simp add: natpermute_finite setcompr_eq_image) also have "\ = (\v\natpermute n (m + 1). \j\{0..m}. a j $ v ! j)" using natpermute_split[of m "m + 1"] by (simp add: Suc) finally show ?thesis . qed qed text \The special form for powers.\ lemma fps_power_nth_Suc: fixes m :: nat and a :: "'a::comm_ring_1 fps" shows "(a ^ Suc m)$n = sum (\v. prod (\j. a $ (v!j)) {0..m}) (natpermute n (m+1))" proof - have th0: "a^Suc m = prod (\i. a) {0..m}" by (simp add: prod_constant) show ?thesis unfolding th0 fps_prod_nth .. qed lemma fps_power_nth: fixes m :: nat and a :: "'a::comm_ring_1 fps" shows "(a ^m)$n = (if m=0 then 1$n else sum (\v. prod (\j. a $ (v!j)) {0..m - 1}) (natpermute n m))" by (cases m) (simp_all add: fps_power_nth_Suc del: power_Suc) lemmas fps_nth_power_0 = fps_power_zeroth lemma natpermute_max_card: assumes n0: "n \ 0" shows "card {xs \ natpermute n (k + 1). n \ set xs} = k + 1" unfolding natpermute_contain_maximal proof - let ?A = "\i. {(replicate (k + 1) 0)[i := n]}" let ?K = "{0 ..k}" have fK: "finite ?K" by simp have fAK: "\i\?K. finite (?A i)" by auto have d: "\i\ ?K. \j\ ?K. i \ j \ {(replicate (k + 1) 0)[i := n]} \ {(replicate (k + 1) 0)[j := n]} = {}" proof clarify fix i j assume i: "i \ ?K" and j: "j \ ?K" and ij: "i \ j" have False if eq: "(replicate (k+1) 0)[i:=n] = (replicate (k+1) 0)[j:= n]" proof - have "(replicate (k+1) 0) [i:=n] ! i = n" using i by (simp del: replicate.simps) moreover have "(replicate (k+1) 0) [j:=n] ! i = 0" using i ij by (simp del: replicate.simps) ultimately show ?thesis using eq n0 by (simp del: replicate.simps) qed then show "{(replicate (k + 1) 0)[i := n]} \ {(replicate (k + 1) 0)[j := n]} = {}" by auto qed from card_UN_disjoint[OF fK fAK d] show "card (\i\{0..k}. {(replicate (k + 1) 0)[i := n]}) = k + 1" by simp qed lemma fps_power_Suc_nth: fixes f :: "'a :: comm_ring_1 fps" assumes k: "k > 0" shows "(f ^ Suc m) $ k = of_nat (Suc m) * (f $ k * (f $ 0) ^ m) + (\v\{v\natpermute k (m+1). k \ set v}. \j = 0..m. f $ v ! j)" proof - define A B where "A = {v\natpermute k (m+1). k \ set v}" and "B = {v\natpermute k (m+1). k \ set v}" have [simp]: "finite A" "finite B" "A \ B = {}" by (auto simp: A_def B_def natpermute_finite) from natpermute_max_card[of k m] k have card_A: "card A = m + 1" by (simp add: A_def) { fix v assume v: "v \ A" from v have [simp]: "length v = Suc m" by (simp add: A_def natpermute_def) from v have "\j. j \ m \ v ! j = k" by (auto simp: set_conv_nth A_def natpermute_def less_Suc_eq_le) - then guess j by (elim exE conjE) note j = this + then obtain j where j: "j \ m" "v ! j = k" by auto from v have "k = sum_list v" by (simp add: A_def natpermute_def) also have "\ = (\i=0..m. v ! i)" by (simp add: sum_list_sum_nth atLeastLessThanSuc_atLeastAtMost del: sum.op_ivl_Suc) also from j have "{0..m} = insert j ({0..m}-{j})" by auto also from j have "(\i\\. v ! i) = k + (\i\{0..m}-{j}. v ! i)" by (subst sum.insert) simp_all finally have "(\i\{0..m}-{j}. v ! i) = 0" by simp hence zero: "v ! i = 0" if "i \ {0..m}-{j}" for i using that by (subst (asm) sum_eq_0_iff) auto from j have "{0..m} = insert j ({0..m} - {j})" by auto also from j have "(\i\\. f $ (v ! i)) = f $ k * (\i\{0..m} - {j}. f $ (v ! i))" by (subst prod.insert) auto also have "(\i\{0..m} - {j}. f $ (v ! i)) = (\i\{0..m} - {j}. f $ 0)" by (intro prod.cong) (simp_all add: zero) also from j have "\ = (f $ 0) ^ m" by (subst prod_constant) simp_all finally have "(\j = 0..m. f $ (v ! j)) = f $ k * (f $ 0) ^ m" . } note A = this have "(f ^ Suc m) $ k = (\v\natpermute k (m + 1). \j = 0..m. f $ v ! j)" by (rule fps_power_nth_Suc) also have "natpermute k (m+1) = A \ B" unfolding A_def B_def by blast also have "(\v\\. \j = 0..m. f $ (v ! j)) = (\v\A. \j = 0..m. f $ (v ! j)) + (\v\B. \j = 0..m. f $ (v ! j))" by (intro sum.union_disjoint) simp_all also have "(\v\A. \j = 0..m. f $ (v ! j)) = of_nat (Suc m) * (f $ k * (f $ 0) ^ m)" by (simp add: A card_A) finally show ?thesis by (simp add: B_def) qed lemma fps_power_Suc_eqD: fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ Suc m = g ^ Suc m" "f $ 0 = g $ 0" "f $ 0 \ 0" shows "f = g" proof (rule fps_ext) fix k :: nat show "f $ k = g $ k" proof (induction k rule: less_induct) case (less k) show ?case proof (cases "k = 0") case False let ?h = "\f. (\v | v \ natpermute k (m + 1) \ k \ set v. \j = 0..m. f $ v ! j)" from False fps_power_Suc_nth[of k f m] fps_power_Suc_nth[of k g m] have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h f = g $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h g" using assms by (simp add: mult_ac del: power_Suc of_nat_Suc) also have "v ! i < k" if "v \ {v\natpermute k (m+1). k \ set v}" "i \ m" for v i using that elem_le_sum_list[of i v] unfolding natpermute_def by (auto simp: set_conv_nth dest!: spec[of _ i]) hence "?h f = ?h g" by (intro sum.cong refl prod.cong less lessI) (simp add: natpermute_def) finally have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) = g $ k * (of_nat (Suc m) * (f $ 0) ^ m)" by simp with assms show "f $ k = g $ k" by (subst (asm) mult_right_cancel) (auto simp del: of_nat_Suc) qed (simp_all add: assms) qed qed lemma fps_power_Suc_eqD': fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ Suc m = g ^ Suc m" "f $ subdegree f = g $ subdegree g" shows "f = g" proof (cases "f = 0") case False have "Suc m * subdegree f = subdegree (f ^ Suc m)" by (rule subdegree_power [symmetric]) also have "f ^ Suc m = g ^ Suc m" by fact also have "subdegree \ = Suc m * subdegree g" by (rule subdegree_power) finally have [simp]: "subdegree f = subdegree g" by (subst (asm) Suc_mult_cancel1) have "fps_shift (subdegree f) f * fps_X ^ subdegree f = f" by (rule subdegree_decompose [symmetric]) also have "\ ^ Suc m = g ^ Suc m" by fact also have "g = fps_shift (subdegree g) g * fps_X ^ subdegree g" by (rule subdegree_decompose) also have "subdegree f = subdegree g" by fact finally have "fps_shift (subdegree g) f ^ Suc m = fps_shift (subdegree g) g ^ Suc m" by (simp add: algebra_simps power_mult_distrib del: power_Suc) hence "fps_shift (subdegree g) f = fps_shift (subdegree g) g" by (rule fps_power_Suc_eqD) (insert assms False, auto) with subdegree_decompose[of f] subdegree_decompose[of g] show ?thesis by simp qed (insert assms, simp_all) lemma fps_power_eqD': fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ m = g ^ m" "f $ subdegree f = g $ subdegree g" "m > 0" shows "f = g" using fps_power_Suc_eqD'[of f "m-1" g] assms by simp lemma fps_power_eqD: fixes f g :: "'a :: {idom,semiring_char_0} fps" assumes "f ^ m = g ^ m" "f $ 0 = g $ 0" "f $ 0 \ 0" "m > 0" shows "f = g" by (rule fps_power_eqD'[of f m g]) (insert assms, simp_all) lemma fps_compose_inj_right: assumes a0: "a$0 = (0::'a::idom)" and a1: "a$1 \ 0" shows "(b oo a = c oo a) \ b = c" (is "?lhs \?rhs") proof show ?lhs if ?rhs using that by simp show ?rhs if ?lhs proof - have "b$n = c$n" for n proof (induct n rule: nat_less_induct) fix n assume H: "\m?lhs\ have "(b oo a)$n = (c oo a)$n" by simp then show ?thesis using 0 by (simp add: fps_compose_nth) next case (Suc n1) have f: "finite {0 .. n1}" "finite {n}" by simp_all have eq: "{0 .. n1} \ {n} = {0 .. n}" using Suc by auto have d: "{0 .. n1} \ {n} = {}" using Suc by auto have seq: "(\i = 0..n1. b $ i * a ^ i $ n) = (\i = 0..n1. c $ i * a ^ i $ n)" using H Suc by auto have th0: "(b oo a) $n = (\i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n" unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] seq using startsby_zero_power_nth_same[OF a0] by simp have th1: "(c oo a) $n = (\i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n" unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] using startsby_zero_power_nth_same[OF a0] by simp from \?lhs\[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1 show ?thesis by auto qed qed then show ?rhs by (simp add: fps_eq_iff) qed qed subsection \Radicals\ declare prod.cong [fundef_cong] function radical :: "(nat \ 'a \ 'a) \ nat \ 'a::field fps \ nat \ 'a" where "radical r 0 a 0 = 1" | "radical r 0 a (Suc n) = 0" | "radical r (Suc k) a 0 = r (Suc k) (a$0)" | "radical r (Suc k) a (Suc n) = (a$ Suc n - sum (\xs. prod (\j. radical r (Suc k) a (xs ! j)) {0..k}) {xs. xs \ natpermute (Suc n) (Suc k) \ Suc n \ set xs}) / (of_nat (Suc k) * (radical r (Suc k) a 0)^k)" by pat_completeness auto termination radical proof let ?R = "measure (\(r, k, a, n). n)" { show "wf ?R" by auto next fix r :: "nat \ 'a \ 'a" and a :: "'a fps" and k n xs i assume xs: "xs \ {xs \ natpermute (Suc n) (Suc k). Suc n \ set xs}" and i: "i \ {0..k}" have False if c: "Suc n \ xs ! i" proof - from xs i have "xs !i \ Suc n" by (simp add: in_set_conv_nth natpermute_def) with c have c': "Suc n < xs!i" by arith have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1.. ({i} \ {i+1 ..< Suc k}) = {}" "{i} \ {i+1..< Suc k} = {}" by auto have eqs: "{0.. ({i} \ {i+1 ..< Suc k})" using i by auto from xs have "Suc n = sum_list xs" by (simp add: natpermute_def) also have "\ = sum (nth xs) {0.. = xs!i + sum (nth xs) {0.. ?R" using not_less by auto next fix r :: "nat \ 'a \ 'a" and a :: "'a fps" and k n show "((r, Suc k, a, 0), r, Suc k, a, Suc n) \ ?R" by simp } qed definition "fps_radical r n a = Abs_fps (radical r n a)" lemma radical_0 [simp]: "\n. 0 < n \ radical r 0 a n = 0" using radical.elims by blast lemma fps_radical0[simp]: "fps_radical r 0 a = 1" by (auto simp add: fps_eq_iff fps_radical_def) lemma fps_radical_nth_0[simp]: "fps_radical r n a $ 0 = (if n = 0 then 1 else r n (a$0))" by (cases n) (simp_all add: fps_radical_def) lemma fps_radical_power_nth[simp]: assumes r: "(r k (a$0)) ^ k = a$0" shows "fps_radical r k a ^ k $ 0 = (if k = 0 then 1 else a$0)" proof (cases k) case 0 then show ?thesis by simp next case (Suc h) have eq1: "fps_radical r k a ^ k $ 0 = (\j\{0..h}. fps_radical r k a $ (replicate k 0) ! j)" unfolding fps_power_nth Suc by simp also have "\ = (\j\{0..h}. r k (a$0))" proof (rule prod.cong [OF refl]) show "fps_radical r k a $ replicate k 0 ! j = r k (a $ 0)" if "j \ {0..h}" for j proof - have "j < Suc h" using that by presburger then show ?thesis by (metis Suc fps_radical_nth_0 nth_replicate old.nat.distinct(2)) qed qed also have "\ = a$0" using r Suc by simp finally show ?thesis using Suc by simp qed lemma power_radical: fixes a:: "'a::field_char_0 fps" assumes a0: "a$0 \ 0" shows "(r (Suc k) (a$0)) ^ Suc k = a$0 \ (fps_radical r (Suc k) a) ^ (Suc k) = a" (is "?lhs \ ?rhs") proof let ?r = "fps_radical r (Suc k) a" show ?rhs if r0: ?lhs proof - from a0 r0 have r00: "r (Suc k) (a$0) \ 0" by auto have "?r ^ Suc k $ z = a$z" for z proof (induct z rule: nat_less_induct) fix n assume H: "\m 0" by simp let ?Pnk = "natpermute n (k + 1)" let ?Pnkn = "{xs \ ?Pnk. n \ set xs}" let ?Pnknn = "{xs \ ?Pnk. n \ set xs}" have eq: "?Pnkn \ ?Pnknn = ?Pnk" by blast have d: "?Pnkn \ ?Pnknn = {}" by blast have f: "finite ?Pnkn" "finite ?Pnknn" using finite_Un[of ?Pnkn ?Pnknn, unfolded eq] by (metis natpermute_finite)+ let ?f = "\v. \j\{0..k}. ?r $ v ! j" have "sum ?f ?Pnkn = sum (\v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn" proof (rule sum.cong) fix v assume v: "v \ {xs \ natpermute n (k + 1). n \ set xs}" let ?ths = "(\j\{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k" from v obtain i where i: "i \ {0..k}" "v = (replicate (k+1) 0) [i:= n]" unfolding natpermute_contain_maximal by auto have "(\j\{0..k}. fps_radical r (Suc k) a $ v ! j) = (\j\{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))" using i r0 by (auto simp del: replicate.simps intro: prod.cong) also have "\ = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k" using i r0 by (simp add: prod_gen_delta) finally show ?ths . qed rule then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k" by (simp add: natpermute_max_card[OF \n \ 0\, simplified]) also have "\ = a$n - sum ?f ?Pnknn" unfolding Suc using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc) finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" . have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn" unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] .. also have "\ = a$n" unfolding fn by simp finally show ?thesis . qed qed then show ?thesis using r0 by (simp add: fps_eq_iff) qed show ?lhs if ?rhs proof - from that have "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0" by simp then show ?thesis unfolding fps_power_nth_Suc by (simp add: prod_constant del: replicate.simps) qed qed lemma radical_unique: assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0" and a0: "r (Suc k) (b$0 ::'a::field_char_0) = a$0" and b0: "b$0 \ 0" shows "a^(Suc k) = b \ a = fps_radical r (Suc k) b" (is "?lhs \ ?rhs" is "_ \ a = ?r") proof show ?lhs if ?rhs using that using power_radical[OF b0, of r k, unfolded r0] by simp show ?rhs if ?lhs proof - have r00: "r (Suc k) (b$0) \ 0" using b0 r0 by auto have ceq: "card {0..k} = Suc k" by simp from a0 have a0r0: "a$0 = ?r$0" by simp have "a $ n = ?r $ n" for n proof (induct n rule: nat_less_induct) fix n assume h: "\m 0" using Suc by simp let ?Pnk = "natpermute n (Suc k)" let ?Pnkn = "{xs \ ?Pnk. n \ set xs}" let ?Pnknn = "{xs \ ?Pnk. n \ set xs}" have eq: "?Pnkn \ ?Pnknn = ?Pnk" by blast have d: "?Pnkn \ ?Pnknn = {}" by blast have f: "finite ?Pnkn" "finite ?Pnknn" using finite_Un[of ?Pnkn ?Pnknn, unfolded eq] by (metis natpermute_finite)+ let ?f = "\v. \j\{0..k}. ?r $ v ! j" let ?g = "\v. \j\{0..k}. a $ v ! j" have "sum ?g ?Pnkn = sum (\v. a $ n * (?r$0)^k) ?Pnkn" proof (rule sum.cong) fix v assume v: "v \ {xs \ natpermute n (Suc k). n \ set xs}" let ?ths = "(\j\{0..k}. a $ v ! j) = a $ n * (?r$0)^k" from v obtain i where i: "i \ {0..k}" "v = (replicate (k+1) 0) [i:= n]" unfolding Suc_eq_plus1 natpermute_contain_maximal by (auto simp del: replicate.simps) have "(\j\{0..k}. a $ v ! j) = (\j\{0..k}. if j = i then a $ n else r (Suc k) (b$0))" using i a0 by (auto simp del: replicate.simps intro: prod.cong) also have "\ = a $ n * (?r $ 0)^k" using i by (simp add: prod_gen_delta) finally show ?ths . qed rule then have th0: "sum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k" by (simp add: natpermute_max_card[OF nz, simplified]) have th1: "sum ?g ?Pnknn = sum ?f ?Pnknn" proof (rule sum.cong, rule refl, rule prod.cong, simp) fix xs i assume xs: "xs \ ?Pnknn" and i: "i \ {0..k}" have False if c: "n \ xs ! i" proof - from xs i have "xs ! i \ n" by (simp add: in_set_conv_nth natpermute_def) with c have c': "n < xs!i" by arith have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1.. ({i} \ {i+1 ..< Suc k}) = {}" "{i} \ {i+1..< Suc k} = {}" by auto have eqs: "{0.. ({i} \ {i+1 ..< Suc k})" using i by auto from xs have "n = sum_list xs" by (simp add: natpermute_def) also have "\ = sum (nth xs) {0.. = xs!i + sum (nth xs) {0..x::'a. of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x" by (simp add: field_simps del: of_nat_Suc) from \?lhs\ have "b$n = a^Suc k $ n" by (simp add: fps_eq_iff) also have "a ^ Suc k$n = sum ?g ?Pnkn + sum ?g ?Pnknn" unfolding fps_power_nth_Suc using sum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric], unfolded eq, of ?g] by simp also have "\ = of_nat (k+1) * a $ n * (?r $ 0)^k + sum ?f ?Pnknn" unfolding th0 th1 .. finally have \
: "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - sum ?f ?Pnknn" by simp have "a$n = (b$n - sum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)" apply (rule eq_divide_imp) using r00 \
by (simp_all add: ac_simps del: of_nat_Suc) then show ?thesis unfolding fps_radical_def Suc by (simp del: of_nat_Suc) qed qed then show ?rhs by (simp add: fps_eq_iff) qed qed lemma radical_power: assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0" and a0: "(a$0 :: 'a::field_char_0) \ 0" shows "(fps_radical r (Suc k) (a ^ Suc k)) = a" proof - let ?ak = "a^ Suc k" have ak0: "?ak $ 0 = (a$0) ^ Suc k" by (simp add: fps_nth_power_0 del: power_Suc) from r0 have th0: "r (Suc k) (a ^ Suc k $ 0) ^ Suc k = a ^ Suc k $ 0" using ak0 by auto from r0 ak0 have th1: "r (Suc k) (a ^ Suc k $ 0) = a $ 0" by auto from ak0 a0 have ak00: "?ak $ 0 \0 " by auto from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis by metis qed lemma fps_deriv_radical': fixes a :: "'a::field_char_0 fps" assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \ 0" shows "fps_deriv (fps_radical r (Suc k) a) = fps_deriv a / ((of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)" proof - let ?r = "fps_radical r (Suc k) a" let ?w = "(of_nat (Suc k)) * ?r ^ k" from a0 r0 have r0': "r (Suc k) (a$0) \ 0" by auto from r0' have w0: "?w $ 0 \ 0" by (simp del: of_nat_Suc) note th0 = inverse_mult_eq_1[OF w0] let ?iw = "inverse ?w" from iffD1[OF power_radical[of a r], OF a0 r0] have "fps_deriv (?r ^ Suc k) = fps_deriv a" by simp then have "fps_deriv ?r * ?w = fps_deriv a" by (simp add: fps_deriv_power' ac_simps del: power_Suc) then have "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a" by simp with a0 r0 have "fps_deriv ?r * (?iw * ?w) = fps_deriv a / ?w" by (subst fps_divide_unit) (auto simp del: of_nat_Suc) then show ?thesis unfolding th0 by simp qed lemma fps_deriv_radical: fixes a :: "'a::field_char_0 fps" assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \ 0" shows "fps_deriv (fps_radical r (Suc k) a) = fps_deriv a / (fps_const (of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)" using fps_deriv_radical'[of r k a, OF r0 a0] by (simp add: fps_of_nat[symmetric]) lemma radical_mult_distrib: fixes a :: "'a::field_char_0 fps" assumes k: "k > 0" and ra0: "r k (a $ 0) ^ k = a $ 0" and rb0: "r k (b $ 0) ^ k = b $ 0" and a0: "a $ 0 \ 0" and b0: "b $ 0 \ 0" shows "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0) \ fps_radical r k (a * b) = fps_radical r k a * fps_radical r k b" (is "?lhs \ ?rhs") proof show ?rhs if r0': ?lhs proof - from r0' have r0: "(r k ((a * b) $ 0)) ^ k = (a * b) $ 0" by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib) show ?thesis proof (cases k) case 0 then show ?thesis using r0' by simp next case (Suc h) let ?ra = "fps_radical r (Suc h) a" let ?rb = "fps_radical r (Suc h) b" have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0" using r0' Suc by (simp add: fps_mult_nth) have ab0: "(a*b) $ 0 \ 0" using a0 b0 by (simp add: fps_mult_nth) from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded Suc] th0 ab0, symmetric] iffD1[OF power_radical[of _ r], OF a0 ra0[unfolded Suc]] iffD1[OF power_radical[of _ r], OF b0 rb0[unfolded Suc]] Suc r0' show ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc) qed qed show ?lhs if ?rhs proof - from that have "(fps_radical r k (a * b)) $ 0 = (fps_radical r k a * fps_radical r k b) $ 0" by simp then show ?thesis using k by (simp add: fps_mult_nth) qed qed (* lemma radical_mult_distrib: fixes a:: "'a::field_char_0 fps" assumes ra0: "r k (a $ 0) ^ k = a $ 0" and rb0: "r k (b $ 0) ^ k = b $ 0" and r0': "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0)" and a0: "a$0 \ 0" and b0: "b$0 \ 0" shows "fps_radical r (k) (a*b) = fps_radical r (k) a * fps_radical r (k) (b)" proof- from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0" by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib) {assume "k=0" then have ?thesis by simp} moreover {fix h assume k: "k = Suc h" let ?ra = "fps_radical r (Suc h) a" let ?rb = "fps_radical r (Suc h) b" have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0" using r0' k by (simp add: fps_mult_nth) have ab0: "(a*b) $ 0 \ 0" using a0 b0 by (simp add: fps_mult_nth) from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded k] th0 ab0, symmetric] power_radical[of r, OF ra0[unfolded k] a0] power_radical[of r, OF rb0[unfolded k] b0] k have ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc)} ultimately show ?thesis by (cases k, auto) qed *) lemma radical_divide: fixes a :: "'a::field_char_0 fps" assumes kp: "k > 0" and ra0: "(r k (a $ 0)) ^ k = a $ 0" and rb0: "(r k (b $ 0)) ^ k = b $ 0" and a0: "a$0 \ 0" and b0: "b$0 \ 0" shows "r k ((a $ 0) / (b$0)) = r k (a$0) / r k (b $ 0) \ fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b" (is "?lhs = ?rhs") proof let ?r = "fps_radical r k" from kp obtain h where k: "k = Suc h" by (cases k) auto have ra0': "r k (a$0) \ 0" using a0 ra0 k by auto have rb0': "r k (b$0) \ 0" using b0 rb0 k by auto show ?lhs if ?rhs proof - from that have "?r (a/b) $ 0 = (?r a / ?r b)$0" by simp then show ?thesis using k a0 b0 rb0' by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse) qed show ?rhs if ?lhs proof - from a0 b0 have ab0[simp]: "(a/b)$0 = a$0 / b$0" by (simp add: fps_divide_def fps_mult_nth divide_inverse fps_inverse_def) have th0: "r k ((a/b)$0) ^ k = (a/b)$0" by (simp add: \?lhs\ power_divide ra0 rb0) from a0 b0 ra0' rb0' kp \?lhs\ have th1: "r k ((a / b) $ 0) = (fps_radical r k a / fps_radical r k b) $ 0" by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse) from a0 b0 ra0' rb0' kp have ab0': "(a / b) $ 0 \ 0" by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def nonzero_imp_inverse_nonzero) note tha[simp] = iffD1[OF power_radical[where r=r and k=h], OF a0 ra0[unfolded k], unfolded k[symmetric]] note thb[simp] = iffD1[OF power_radical[where r=r and k=h], OF b0 rb0[unfolded k], unfolded k[symmetric]] from b0 rb0' have th2: "(?r a / ?r b)^k = a/b" by (simp add: fps_divide_unit power_mult_distrib fps_inverse_power[symmetric]) from iffD1[OF radical_unique[where r=r and a="?r a / ?r b" and b="a/b" and k=h], symmetric, unfolded k[symmetric], OF th0 th1 ab0' th2] show ?thesis . qed qed lemma radical_inverse: fixes a :: "'a::field_char_0 fps" assumes k: "k > 0" and ra0: "r k (a $ 0) ^ k = a $ 0" and r1: "(r k 1)^k = 1" and a0: "a$0 \ 0" shows "r k (inverse (a $ 0)) = r k 1 / (r k (a $ 0)) \ fps_radical r k (inverse a) = fps_radical r k 1 / fps_radical r k a" using radical_divide[where k=k and r=r and a=1 and b=a, OF k ] ra0 r1 a0 by (simp add: divide_inverse fps_divide_def) subsection \Derivative of composition\ lemma fps_compose_deriv: fixes a :: "'a::idom fps" assumes b0: "b$0 = 0" shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * fps_deriv b" proof - have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n" for n proof - have "(fps_deriv (a oo b))$n = sum (\i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}" by (simp add: fps_compose_def field_simps sum_distrib_left del: of_nat_Suc) also have "\ = sum (\i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}" by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc) also have "\ = sum (\i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}" unfolding fps_mult_left_const_nth by (simp add: field_simps) also have "\ = sum (\i. of_nat i * a$i * (sum (\j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}" unfolding fps_mult_nth .. also have "\ = sum (\i. of_nat i * a$i * (sum (\j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}" by (intro sum.mono_neutral_right) (auto simp add: mult_delta_left not_le) also have "\ = sum (\i. of_nat (i + 1) * a$(i+1) * (sum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" unfolding fps_deriv_nth by (rule sum.reindex_cong [of Suc]) (simp_all add: mult.assoc) finally have th0: "(fps_deriv (a oo b))$n = sum (\i. of_nat (i + 1) * a$(i+1) * (sum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" . have "(((fps_deriv a) oo b) * (fps_deriv b))$n = sum (\i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}" unfolding fps_mult_nth by (simp add: ac_simps) also have "\ = sum (\i. sum (\j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}" unfolding fps_deriv_nth fps_compose_nth sum_distrib_left mult.assoc by (auto simp: subset_eq b0 startsby_zero_power_prefix sum.mono_neutral_left intro: sum.cong) also have "\ = sum (\i. of_nat (i + 1) * a$(i+1) * (sum (\j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" unfolding sum_distrib_left by (subst sum.swap) (force intro: sum.cong) finally show ?thesis unfolding th0 by simp qed then show ?thesis by (simp add: fps_eq_iff) qed subsection \Finite FPS (i.e. polynomials) and fps_X\ lemma fps_poly_sum_fps_X: assumes "\i > n. a$i = 0" shows "a = sum (\i. fps_const (a$i) * fps_X^i) {0..n}" (is "a = ?r") proof - have "a$i = ?r$i" for i unfolding fps_sum_nth fps_mult_left_const_nth fps_X_power_nth by (simp add: mult_delta_right assms) then show ?thesis unfolding fps_eq_iff by blast qed subsection \Compositional inverses\ fun compinv :: "'a fps \ nat \ 'a::field" where "compinv a 0 = fps_X$0" | "compinv a (Suc n) = (fps_X$ Suc n - sum (\i. (compinv a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n" definition "fps_inv a = Abs_fps (compinv a)" lemma fps_inv: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "fps_inv a oo a = fps_X" proof - let ?i = "fps_inv a oo a" have "?i $n = fps_X$n" for n proof (induct n rule: nat_less_induct) fix n assume h: "\mi. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1" by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc) also have "\ = sum (\i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + (fps_X$ Suc n1 - sum (\i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})" using a0 a1 Suc by (simp add: fps_inv_def) also have "\ = fps_X$n" using Suc by simp finally show ?thesis . qed qed then show ?thesis by (simp add: fps_eq_iff) qed fun gcompinv :: "'a fps \ 'a fps \ nat \ 'a::field" where "gcompinv b a 0 = b$0" | "gcompinv b a (Suc n) = (b$ Suc n - sum (\i. (gcompinv b a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n" definition "fps_ginv b a = Abs_fps (gcompinv b a)" lemma fps_ginv: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "fps_ginv b a oo a = b" proof - let ?i = "fps_ginv b a oo a" have "?i $n = b$n" for n proof (induct n rule: nat_less_induct) fix n assume h: "\mi. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1" by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc) also have "\ = sum (\i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + (b$ Suc n1 - sum (\i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})" using a0 a1 Suc by (simp add: fps_ginv_def) also have "\ = b$n" using Suc by simp finally show ?thesis . qed qed then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_inv_ginv: "fps_inv = fps_ginv fps_X" proof - have "compinv x n = gcompinv fps_X x n" for n and x :: "'a fps" proof (induction n rule: nat_less_induct) case (1 n) then show ?case by (cases n) auto qed then show ?thesis by (auto simp add: fun_eq_iff fps_eq_iff fps_inv_def fps_ginv_def) qed lemma fps_compose_1[simp]: "1 oo a = 1" by (simp add: fps_eq_iff fps_compose_nth mult_delta_left) lemma fps_compose_0[simp]: "0 oo a = 0" by (simp add: fps_eq_iff fps_compose_nth) lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a $ 0)" by (simp add: fps_eq_iff fps_compose_nth power_0_left sum.neutral) lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)" by (simp add: fps_eq_iff fps_compose_nth field_simps sum.distrib) lemma fps_compose_sum_distrib: "(sum f S) oo a = sum (\i. f i oo a) S" proof (cases "finite S") case True show ?thesis proof (rule finite_induct[OF True]) show "sum f {} oo a = (\i\{}. f i oo a)" by simp next fix x F assume fF: "finite F" and xF: "x \ F" and h: "sum f F oo a = sum (\i. f i oo a) F" show "sum f (insert x F) oo a = sum (\i. f i oo a) (insert x F)" using fF xF h by (simp add: fps_compose_add_distrib) qed next case False then show ?thesis by simp qed lemma convolution_eq: "sum (\i. a (i :: nat) * b (n - i)) {0 .. n} = sum (\(i,j). a i * b j) {(i,j). i \ n \ j \ n \ i + j = n}" by (rule sum.reindex_bij_witness[where i=fst and j="\i. (i, n - i)"]) auto lemma product_composition_lemma: assumes c0: "c$0 = (0::'a::idom)" and d0: "d$0 = 0" shows "((a oo c) * (b oo d))$n = sum (\(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \ n}" (is "?l = ?r") proof - let ?S = "{(k::nat, m::nat). k + m \ n}" have s: "?S \ {0..n} \ {0..n}" by (simp add: subset_eq) have f: "finite {(k::nat, m::nat). k + m \ n}" by (auto intro: finite_subset[OF s]) have "?r = (\(k, m) \ {(k, m). k + m \ n}. \j = 0..n. a $ k * b $ m * (c ^ k $ j * d ^ m $ (n - j)))" by (simp add: fps_mult_nth sum_distrib_left) also have "\ = (\i = 0..n. \(k,m)\{(k,m). k+m \ n}. a $ k * c ^ k $ i * b $ m * d ^ m $ (n-i))" unfolding sum.swap [where A = "{0..n}"] by (auto simp add: field_simps intro: sum.cong) also have "... = (\i = 0..n. \q = 0..i. \j = 0..n - i. a $ q * c ^ q $ i * (b $ j * d ^ j $ (n - i)))" apply (rule sum.cong [OF refl]) apply (simp add: sum.cartesian_product mult.assoc) apply (rule sum.mono_neutral_right[OF f], force) by clarsimp (meson c0 d0 leI startsby_zero_power_prefix) also have "\ = ?l" by (simp add: fps_mult_nth fps_compose_nth sum_product) finally show ?thesis by simp qed lemma sum_pair_less_iff: "sum (\((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \ n} = sum (\s. sum (\i. a i * b (s - i) * c s) {0..s}) {0..n}" (is "?l = ?r") proof - have th0: "{(k, m). k + m \ n} = (\s\{0..n}. \i\{0..s}. {(i, s - i)})" by auto show "?l = ?r" unfolding th0 by (simp add: sum.UNION_disjoint eq_diff_iff disjoint_iff) qed lemma fps_compose_mult_distrib_lemma: assumes c0: "c$0 = (0::'a::idom)" shows "((a oo c) * (b oo c))$n = sum (\s. sum (\i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}" unfolding product_composition_lemma[OF c0 c0] power_add[symmetric] unfolding sum_pair_less_iff[where a = "\k. a$k" and b="\m. b$m" and c="\s. (c ^ s)$n" and n = n] .. lemma fps_compose_mult_distrib: assumes c0: "c $ 0 = (0::'a::idom)" shows "(a * b) oo c = (a oo c) * (b oo c)" proof (clarsimp simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0]) show "(a * b oo c) $ n = (\s = 0..n. \i = 0..s. a $ i * b $ (s - i) * c ^ s $ n)" for n by (simp add: fps_compose_nth fps_mult_nth sum_distrib_right) qed lemma fps_compose_prod_distrib: assumes c0: "c$0 = (0::'a::idom)" shows "prod a S oo c = prod (\k. a k oo c) S" proof (induct S rule: infinite_finite_induct) next case (insert) then show ?case by (simp add: fps_compose_mult_distrib[OF c0]) qed auto lemma fps_compose_divide: assumes [simp]: "g dvd f" "h $ 0 = 0" shows "fps_compose f h = fps_compose (f / g :: 'a :: field fps) h * fps_compose g h" proof - have "f = (f / g) * g" by simp also have "fps_compose \ h = fps_compose (f / g) h * fps_compose g h" by (subst fps_compose_mult_distrib) simp_all finally show ?thesis . qed lemma fps_compose_divide_distrib: assumes "g dvd f" "h $ 0 = 0" "fps_compose g h \ 0" shows "fps_compose (f / g :: 'a :: field fps) h = fps_compose f h / fps_compose g h" using fps_compose_divide[OF assms(1,2)] assms(3) by simp lemma fps_compose_power: assumes c0: "c$0 = (0::'a::idom)" shows "(a oo c)^n = a^n oo c" proof (cases n) case 0 then show ?thesis by simp next case (Suc m) have "(\n = 0..m. a) oo c = (\n = 0..m. a oo c)" using c0 fps_compose_prod_distrib by blast moreover have th0: "a^n = prod (\k. a) {0..m}" "(a oo c) ^ n = prod (\k. a oo c) {0..m}" by (simp_all add: prod_constant Suc) ultimately show ?thesis by presburger qed lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)" by (simp add: fps_eq_iff fps_compose_nth field_simps sum_negf[symmetric]) lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)" using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus) lemma fps_X_fps_compose: "fps_X oo a = Abs_fps (\n. if n = 0 then (0::'a::comm_ring_1) else a$n)" by (simp add: fps_eq_iff fps_compose_nth mult_delta_left) lemma fps_inverse_compose: assumes b0: "(b$0 :: 'a::field) = 0" and a0: "a$0 \ 0" shows "inverse a oo b = inverse (a oo b)" proof - let ?ia = "inverse a" let ?ab = "a oo b" let ?iab = "inverse ?ab" from a0 have ia0: "?ia $ 0 \ 0" by simp from a0 have ab0: "?ab $ 0 \ 0" by (simp add: fps_compose_def) have "(?ia oo b) * (a oo b) = 1" unfolding fps_compose_mult_distrib[OF b0, symmetric] unfolding inverse_mult_eq_1[OF a0] fps_compose_1 .. then have "(?ia oo b) * (a oo b) * ?iab = 1 * ?iab" by simp then have "(?ia oo b) * (?iab * (a oo b)) = ?iab" by simp then show ?thesis unfolding inverse_mult_eq_1[OF ab0] by simp qed lemma fps_divide_compose: assumes c0: "(c$0 :: 'a::field) = 0" and b0: "b$0 \ 0" shows "(a/b) oo c = (a oo c) / (b oo c)" using b0 c0 by (simp add: fps_divide_unit fps_inverse_compose fps_compose_mult_distrib) lemma gp: assumes a0: "a$0 = (0::'a::field)" shows "(Abs_fps (\n. 1)) oo a = 1/(1 - a)" (is "?one oo a = _") proof - have o0: "?one $ 0 \ 0" by simp have th0: "(1 - fps_X) $ 0 \ (0::'a)" by simp from fps_inverse_gp[where ?'a = 'a] have "inverse ?one = 1 - fps_X" by (simp add: fps_eq_iff) then have "inverse (inverse ?one) = inverse (1 - fps_X)" by simp then have th: "?one = 1/(1 - fps_X)" unfolding fps_inverse_idempotent[OF o0] by (simp add: fps_divide_def) show ?thesis unfolding th unfolding fps_divide_compose[OF a0 th0] fps_compose_1 fps_compose_sub_distrib fps_X_fps_compose_startby0[OF a0] .. qed lemma fps_compose_radical: assumes b0: "b$0 = (0::'a::field_char_0)" and ra0: "r (Suc k) (a$0) ^ Suc k = a$0" and a0: "a$0 \ 0" shows "fps_radical r (Suc k) a oo b = fps_radical r (Suc k) (a oo b)" proof - let ?r = "fps_radical r (Suc k)" let ?ab = "a oo b" have ab0: "?ab $ 0 = a$0" by (simp add: fps_compose_def) from ab0 a0 ra0 have rab0: "?ab $ 0 \ 0" "r (Suc k) (?ab $ 0) ^ Suc k = ?ab $ 0" by simp_all have th00: "r (Suc k) ((a oo b) $ 0) = (fps_radical r (Suc k) a oo b) $ 0" by (simp add: ab0 fps_compose_def) have th0: "(?r a oo b) ^ (Suc k) = a oo b" unfolding fps_compose_power[OF b0] unfolding iffD1[OF power_radical[of a r k], OF a0 ra0] .. from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0] show ?thesis . qed lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b" by (simp add: fps_eq_iff fps_compose_nth sum_distrib_left mult.assoc) lemma fps_const_mult_apply_right: "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b" by (simp add: fps_const_mult_apply_left mult.commute) lemma fps_compose_assoc: assumes c0: "c$0 = (0::'a::idom)" and b0: "b$0 = 0" shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r") proof - have "?l$n = ?r$n" for n proof - have "?l$n = (sum (\i. (fps_const (a$i) * b^i) oo c) {0..n})$n" by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left sum_distrib_left mult.assoc fps_sum_nth) also have "\ = ((sum (\i. fps_const (a$i) * b^i) {0..n}) oo c)$n" by (simp add: fps_compose_sum_distrib) also have "... = (\i = 0..n. \j = 0..n. a $ j * (b ^ j $ i * c ^ i $ n))" by (simp add: fps_compose_nth fps_sum_nth sum_distrib_right mult.assoc) also have "... = (\i = 0..n. \j = 0..i. a $ j * (b ^ j $ i * c ^ i $ n))" by (intro sum.cong [OF refl] sum.mono_neutral_right; simp add: b0 startsby_zero_power_prefix) also have "\ = ?r$n" by (simp add: fps_compose_nth sum_distrib_right mult.assoc) finally show ?thesis . qed then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_X_power_compose: assumes a0: "a$0=0" shows "fps_X^k oo a = (a::'a::idom fps)^k" (is "?l = ?r") proof (cases k) case 0 then show ?thesis by simp next case (Suc h) have "?l $ n = ?r $n" for n proof - consider "k > n" | "k \ n" by arith then show ?thesis proof cases case 1 then show ?thesis using a0 startsby_zero_power_prefix[OF a0] Suc by (simp add: fps_compose_nth del: power_Suc) next case 2 then show ?thesis by (simp add: fps_compose_nth mult_delta_left) qed qed then show ?thesis unfolding fps_eq_iff by blast qed lemma fps_inv_right: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "a oo fps_inv a = fps_X" proof - let ?ia = "fps_inv a" let ?iaa = "a oo fps_inv a" have th0: "?ia $ 0 = 0" by (simp add: fps_inv_def) have th1: "?iaa $ 0 = 0" using a0 a1 by (simp add: fps_inv_def fps_compose_nth) have th2: "fps_X$0 = 0" by simp from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo fps_X" by simp then have "(a oo fps_inv a) oo a = fps_X oo a" by (simp add: fps_compose_assoc[OF a0 th0] fps_X_fps_compose_startby0[OF a0]) with fps_compose_inj_right[OF a0 a1] show ?thesis by simp qed lemma fps_inv_deriv: assumes a0: "a$0 = (0::'a::field)" and a1: "a$1 \ 0" shows "fps_deriv (fps_inv a) = inverse (fps_deriv a oo fps_inv a)" proof - let ?ia = "fps_inv a" let ?d = "fps_deriv a oo ?ia" let ?dia = "fps_deriv ?ia" have ia0: "?ia$0 = 0" by (simp add: fps_inv_def) have th0: "?d$0 \ 0" using a1 by (simp add: fps_compose_nth) from fps_inv_right[OF a0 a1] have "?d * ?dia = 1" by (simp add: fps_compose_deriv[OF ia0, of a, symmetric] ) then have "inverse ?d * ?d * ?dia = inverse ?d * 1" by simp with inverse_mult_eq_1 [OF th0] show "?dia = inverse ?d" by simp qed lemma fps_inv_idempotent: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" shows "fps_inv (fps_inv a) = a" proof - let ?r = "fps_inv" have ra0: "?r a $ 0 = 0" by (simp add: fps_inv_def) from a1 have ra1: "?r a $ 1 \ 0" by (simp add: fps_inv_def field_simps) have fps_X0: "fps_X$0 = 0" by simp from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = fps_X" . then have "?r (?r a) oo ?r a oo a = fps_X oo a" by simp then have "?r (?r a) oo (?r a oo a) = a" unfolding fps_X_fps_compose_startby0[OF a0] unfolding fps_compose_assoc[OF a0 ra0, symmetric] . then show ?thesis unfolding fps_inv[OF a0 a1] by simp qed lemma fps_ginv_ginv: assumes a0: "a$0 = 0" and a1: "a$1 \ 0" and c0: "c$0 = 0" and c1: "c$1 \ 0" shows "fps_ginv b (fps_ginv c a) = b oo a oo fps_inv c" proof - let ?r = "fps_ginv" from c0 have rca0: "?r c a $0 = 0" by (simp add: fps_ginv_def) from a1 c1 have rca1: "?r c a $ 1 \ 0" by (simp add: fps_ginv_def field_simps) from fps_ginv[OF rca0 rca1] have "?r b (?r c a) oo ?r c a = b" . then have "?r b (?r c a) oo ?r c a oo a = b oo a" by simp then have "?r b (?r c a) oo (?r c a oo a) = b oo a" by (simp add: a0 fps_compose_assoc rca0) then have "?r b (?r c a) oo c = b oo a" unfolding fps_ginv[OF a0 a1] . then have "?r b (?r c a) oo c oo fps_inv c= b oo a oo fps_inv c" by simp then have "?r b (?r c a) oo (c oo fps_inv c) = b oo a oo fps_inv c" by (metis c0 c1 fps_compose_assoc fps_compose_nth_0 fps_inv fps_inv_right) then show ?thesis unfolding fps_inv_right[OF c0 c1] by simp qed lemma fps_ginv_deriv: assumes a0:"a$0 = (0::'a::field)" and a1: "a$1 \ 0" shows "fps_deriv (fps_ginv b a) = (fps_deriv b / fps_deriv a) oo fps_ginv fps_X a" proof - let ?ia = "fps_ginv b a" let ?ifps_Xa = "fps_ginv fps_X a" let ?d = "fps_deriv" let ?dia = "?d ?ia" have ifps_Xa0: "?ifps_Xa $ 0 = 0" by (simp add: fps_ginv_def) have da0: "?d a $ 0 \ 0" using a1 by simp from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b" by simp then have "(?d ?ia oo a) * ?d a = ?d b" unfolding fps_compose_deriv[OF a0] . then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)" by simp with a1 have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a" by (simp add: fps_divide_unit) then have "(?d ?ia oo a) oo ?ifps_Xa = (?d b / ?d a) oo ?ifps_Xa" unfolding inverse_mult_eq_1[OF da0] by simp then have "?d ?ia oo (a oo ?ifps_Xa) = (?d b / ?d a) oo ?ifps_Xa" unfolding fps_compose_assoc[OF ifps_Xa0 a0] . then show ?thesis unfolding fps_inv_ginv[symmetric] unfolding fps_inv_right[OF a0 a1] by simp qed lemma fps_compose_linear: "fps_compose (f :: 'a :: comm_ring_1 fps) (fps_const c * fps_X) = Abs_fps (\n. c^n * f $ n)" by (simp add: fps_eq_iff fps_compose_def power_mult_distrib if_distrib cong: if_cong) lemma fps_compose_uminus': "fps_compose f (-fps_X :: 'a :: comm_ring_1 fps) = Abs_fps (\n. (-1)^n * f $ n)" using fps_compose_linear[of f "-1"] by (simp only: fps_const_neg [symmetric] fps_const_1_eq_1) simp subsection \Elementary series\ subsubsection \Exponential series\ definition "fps_exp x = Abs_fps (\n. x^n / of_nat (fact n))" lemma fps_exp_deriv[simp]: "fps_deriv (fps_exp a) = fps_const (a::'a::field_char_0) * fps_exp a" (is "?l = ?r") proof - have "?l$n = ?r $ n" for n using of_nat_neq_0 by (auto simp add: fps_exp_def divide_simps) then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_exp_unique_ODE: "fps_deriv a = fps_const c * a \ a = fps_const (a$0) * fps_exp (c::'a::field_char_0)" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs proof - from that have th: "\n. a $ Suc n = c * a$n / of_nat (Suc n)" by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc) have th': "a$n = a$0 * c ^ n/ (fact n)" for n proof (induct n) case 0 then show ?case by simp next case Suc then show ?case by (simp add: th divide_simps) qed show ?thesis by (auto simp add: fps_eq_iff fps_const_mult_left fps_exp_def intro: th') qed show ?lhs if ?rhs using that by (metis fps_exp_deriv fps_deriv_mult_const_left mult.left_commute) qed lemma fps_exp_add_mult: "fps_exp (a + b) = fps_exp (a::'a::field_char_0) * fps_exp b" (is "?l = ?r") proof - have "fps_deriv ?r = fps_const (a + b) * ?r" by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add) then have "?r = ?l" by (simp only: fps_exp_unique_ODE) (simp add: fps_mult_nth fps_exp_def) then show ?thesis .. qed lemma fps_exp_nth[simp]: "fps_exp a $ n = a^n / of_nat (fact n)" by (simp add: fps_exp_def) lemma fps_exp_0[simp]: "fps_exp (0::'a::field) = 1" by (simp add: fps_eq_iff power_0_left) lemma fps_exp_neg: "fps_exp (- a) = inverse (fps_exp (a::'a::field_char_0))" proof - from fps_exp_add_mult[of a "- a"] have th0: "fps_exp a * fps_exp (- a) = 1" by simp from fps_inverse_unique[OF th0] show ?thesis by simp qed lemma fps_exp_nth_deriv[simp]: "fps_nth_deriv n (fps_exp (a::'a::field_char_0)) = (fps_const a)^n * (fps_exp a)" by (induct n) auto lemma fps_X_compose_fps_exp[simp]: "fps_X oo fps_exp (a::'a::field) = fps_exp a - 1" by (simp add: fps_eq_iff fps_X_fps_compose) lemma fps_inv_fps_exp_compose: assumes a: "a \ 0" shows "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = fps_X" and "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_X" proof - let ?b = "fps_exp a - 1" have b0: "?b $ 0 = 0" by simp have b1: "?b $ 1 \ 0" by (simp add: a) from fps_inv[OF b0 b1] show "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = fps_X" . from fps_inv_right[OF b0 b1] show "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_X" . qed lemma fps_exp_power_mult: "(fps_exp (c::'a::field_char_0))^n = fps_exp (of_nat n * c)" by (induct n) (simp_all add: field_simps fps_exp_add_mult) lemma radical_fps_exp: assumes r: "r (Suc k) 1 = 1" shows "fps_radical r (Suc k) (fps_exp (c::'a::field_char_0)) = fps_exp (c / of_nat (Suc k))" proof - let ?ck = "(c / of_nat (Suc k))" let ?r = "fps_radical r (Suc k)" have eq0[simp]: "?ck * of_nat (Suc k) = c" "of_nat (Suc k) * ?ck = c" by (simp_all del: of_nat_Suc) have th0: "fps_exp ?ck ^ (Suc k) = fps_exp c" unfolding fps_exp_power_mult eq0 .. have th: "r (Suc k) (fps_exp c $0) ^ Suc k = fps_exp c $ 0" "r (Suc k) (fps_exp c $ 0) = fps_exp ?ck $ 0" "fps_exp c $ 0 \ 0" using r by simp_all from th0 radical_unique[where r=r and k=k, OF th] show ?thesis by auto qed lemma fps_exp_compose_linear [simp]: "fps_exp (d::'a::field_char_0) oo (fps_const c * fps_X) = fps_exp (c * d)" by (simp add: fps_compose_linear fps_exp_def fps_eq_iff power_mult_distrib) lemma fps_fps_exp_compose_minus [simp]: "fps_compose (fps_exp c) (-fps_X) = fps_exp (-c :: 'a :: field_char_0)" using fps_exp_compose_linear[of c "-1 :: 'a"] unfolding fps_const_neg [symmetric] fps_const_1_eq_1 by simp lemma fps_exp_eq_iff [simp]: "fps_exp c = fps_exp d \ c = (d :: 'a :: field_char_0)" proof assume "fps_exp c = fps_exp d" from arg_cong[of _ _ "\F. F $ 1", OF this] show "c = d" by simp qed simp_all lemma fps_exp_eq_fps_const_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = fps_const c' \ c = 0 \ c' = 1" proof assume "c = 0 \ c' = 1" thus "fps_exp c = fps_const c'" by (simp add: fps_eq_iff) next assume "fps_exp c = fps_const c'" from arg_cong[of _ _ "\F. F $ 1", OF this] arg_cong[of _ _ "\F. F $ 0", OF this] show "c = 0 \ c' = 1" by simp_all qed lemma fps_exp_neq_0 [simp]: "\fps_exp (c :: 'a :: field_char_0) = 0" unfolding fps_const_0_eq_0 [symmetric] fps_exp_eq_fps_const_iff by simp lemma fps_exp_eq_1_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = 1 \ c = 0" unfolding fps_const_1_eq_1 [symmetric] fps_exp_eq_fps_const_iff by simp lemma fps_exp_neq_numeral_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = numeral n \ c = 0 \ n = Num.One" unfolding numeral_fps_const fps_exp_eq_fps_const_iff by simp subsubsection \Logarithmic series\ lemma Abs_fps_if_0: "Abs_fps (\n. if n = 0 then (v::'a::ring_1) else f n) = fps_const v + fps_X * Abs_fps (\n. f (Suc n))" by (simp add: fps_eq_iff) definition fps_ln :: "'a::field_char_0 \ 'a fps" where "fps_ln c = fps_const (1/c) * Abs_fps (\n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n)" lemma fps_ln_deriv: "fps_deriv (fps_ln c) = fps_const (1/c) * inverse (1 + fps_X)" unfolding fps_inverse_fps_X_plus1 by (simp add: fps_ln_def fps_eq_iff del: of_nat_Suc) lemma fps_ln_nth: "fps_ln c $ n = (if n = 0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))" by (simp add: fps_ln_def field_simps) lemma fps_ln_0 [simp]: "fps_ln c $ 0 = 0" by (simp add: fps_ln_def) lemma fps_ln_fps_exp_inv: fixes a :: "'a::field_char_0" assumes a: "a \ 0" shows "fps_ln a = fps_inv (fps_exp a - 1)" (is "?l = ?r") proof - let ?b = "fps_exp a - 1" have b0: "?b $ 0 = 0" by simp have b1: "?b $ 1 \ 0" by (simp add: a) have "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) = (fps_const a * (fps_exp a - 1) + fps_const a) oo fps_inv (fps_exp a - 1)" by (simp add: field_simps) also have "\ = fps_const a * (fps_X + 1)" by (simp add: fps_compose_add_distrib fps_inv_right[OF b0 b1] distrib_left flip: fps_const_mult_apply_left) finally have eq: "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_const a * (fps_X + 1)" . from fps_inv_deriv[OF b0 b1, unfolded eq] have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (fps_X + 1)" using a by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult) then have "fps_deriv ?l = fps_deriv ?r" by (simp add: fps_ln_deriv add.commute fps_divide_def divide_inverse) then show ?thesis unfolding fps_deriv_eq_iff by (simp add: fps_ln_nth fps_inv_def) qed lemma fps_ln_mult_add: assumes c0: "c\0" and d0: "d\0" shows "fps_ln c + fps_ln d = fps_const (c+d) * fps_ln (c*d)" (is "?r = ?l") proof- from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps) have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + fps_X)" by (simp add: fps_ln_deriv fps_const_add[symmetric] algebra_simps del: fps_const_add) also have "\ = fps_deriv ?l" by (simp add: eq fps_ln_deriv) finally show ?thesis unfolding fps_deriv_eq_iff by simp qed lemma fps_X_dvd_fps_ln [simp]: "fps_X dvd fps_ln c" proof - have "fps_ln c = fps_X * Abs_fps (\n. (-1) ^ n / (of_nat (Suc n) * c))" by (intro fps_ext) (simp add: fps_ln_def of_nat_diff) thus ?thesis by simp qed subsubsection \Binomial series\ definition "fps_binomial a = Abs_fps (\n. a gchoose n)" lemma fps_binomial_nth[simp]: "fps_binomial a $ n = a gchoose n" by (simp add: fps_binomial_def) lemma fps_binomial_ODE_unique: fixes c :: "'a::field_char_0" shows "fps_deriv a = (fps_const c * a) / (1 + fps_X) \ a = fps_const (a$0) * fps_binomial c" (is "?lhs \ ?rhs") proof let ?da = "fps_deriv a" let ?x1 = "(1 + fps_X):: 'a fps" let ?l = "?x1 * ?da" let ?r = "fps_const c * a" have eq: "?l = ?r \ ?lhs" proof - have x10: "?x1 $ 0 \ 0" by simp have "?l = ?r \ inverse ?x1 * ?l = inverse ?x1 * ?r" by simp also have "\ \ ?da = (fps_const c * a) / ?x1" unfolding fps_divide_def mult.assoc[symmetric] inverse_mult_eq_1[OF x10] by (simp add: field_simps) finally show ?thesis . qed show ?rhs if ?lhs proof - from eq that have h: "?l = ?r" .. have th0: "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" for n proof - from h have "?l $ n = ?r $ n" by simp then show ?thesis by (simp add: field_simps del: of_nat_Suc split: if_split_asm) qed have th1: "a $ n = (c gchoose n) * a $ 0" for n proof (induct n) case 0 then show ?case by simp next case (Suc m) have "(c - of_nat m) * (c gchoose m) = (c gchoose Suc m) * of_nat (Suc m)" by (metis gbinomial_absorb_comp gbinomial_absorption mult.commute) with Suc show ?case unfolding th0 by (simp add: divide_simps del: of_nat_Suc) qed show ?thesis by (metis expand_fps_eq fps_binomial_nth fps_mult_right_const_nth mult.commute th1) qed show ?lhs if ?rhs proof - have th00: "x * (a $ 0 * y) = a $ 0 * (x * y)" for x y by (simp add: mult.commute) have "?l = (1 + fps_X) * fps_deriv (fps_const (a $ 0) * fps_binomial c)" using that by auto also have "... = fps_const c * (fps_const (a $ 0) * fps_binomial c)" proof (clarsimp simp add: fps_eq_iff algebra_simps) show "a $ 0 * (c gchoose Suc n) + (of_nat n * ((c gchoose n) * a $ 0) + of_nat n * (a $ 0 * (c gchoose Suc n))) = c * ((c gchoose n) * a $ 0)" for n unfolding mult.assoc[symmetric] by (simp add: field_simps gbinomial_mult_1) qed also have "... = ?r" using that by auto finally have "?l = ?r" . with eq show ?thesis .. qed qed lemma fps_binomial_ODE_unique': "(fps_deriv a = fps_const c * a / (1 + fps_X) \ a $ 0 = 1) \ (a = fps_binomial c)" by (subst fps_binomial_ODE_unique) auto lemma fps_binomial_deriv: "fps_deriv (fps_binomial c) = fps_const c * fps_binomial c / (1 + fps_X)" proof - let ?a = "fps_binomial c" have th0: "?a = fps_const (?a$0) * ?a" by (simp) from iffD2[OF fps_binomial_ODE_unique, OF th0] show ?thesis . qed lemma fps_binomial_add_mult: "fps_binomial (c+d) = fps_binomial c * fps_binomial d" (is "?l = ?r") proof - let ?P = "?r - ?l" let ?b = "fps_binomial" let ?db = "\x. fps_deriv (?b x)" have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)" by simp also have "\ = inverse (1 + fps_X) * (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))" unfolding fps_binomial_deriv by (simp add: fps_divide_def field_simps) also have "\ = (fps_const (c + d)/ (1 + fps_X)) * ?P" by (simp add: field_simps fps_divide_unit fps_const_add[symmetric] del: fps_const_add) finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + fps_X)" by (simp add: fps_divide_def) have "?P = fps_const (?P$0) * ?b (c + d)" unfolding fps_binomial_ODE_unique[symmetric] using th0 by simp then have "?P = 0" by (simp add: fps_mult_nth) then show ?thesis by simp qed lemma fps_binomial_minus_one: "fps_binomial (- 1) = inverse (1 + fps_X)" (is "?l = inverse ?r") proof- have th: "?r$0 \ 0" by simp have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + fps_X)" by (simp add: fps_inverse_deriv[OF th] fps_divide_def power2_eq_square mult.commute fps_const_neg[symmetric] del: fps_const_neg) have eq: "inverse ?r $ 0 = 1" by (simp add: fps_inverse_def) from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + fps_X)" "- 1"] th'] eq show ?thesis by (simp add: fps_inverse_def) qed lemma fps_binomial_of_nat: "fps_binomial (of_nat n) = (1 + fps_X :: 'a :: field_char_0 fps) ^ n" proof (cases "n = 0") case [simp]: True have "fps_deriv ((1 + fps_X) ^ n :: 'a fps) = 0" by simp also have "\ = fps_const (of_nat n) * (1 + fps_X) ^ n / (1 + fps_X)" by (simp add: fps_binomial_def) finally show ?thesis by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) simp_all next case False have "fps_deriv ((1 + fps_X) ^ n :: 'a fps) = fps_const (of_nat n) * (1 + fps_X) ^ (n - 1)" by (simp add: fps_deriv_power) also have "(1 + fps_X :: 'a fps) $ 0 \ 0" by simp hence "(1 + fps_X :: 'a fps) \ 0" by (intro notI) (simp only: , simp) with False have "(1 + fps_X :: 'a fps) ^ (n - 1) = (1 + fps_X) ^ n / (1 + fps_X)" by (cases n) (simp_all ) also have "fps_const (of_nat n :: 'a) * ((1 + fps_X) ^ n / (1 + fps_X)) = fps_const (of_nat n) * (1 + fps_X) ^ n / (1 + fps_X)" by (simp add: unit_div_mult_swap) finally show ?thesis by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) (simp_all add: fps_power_nth) qed lemma fps_binomial_0 [simp]: "fps_binomial 0 = 1" using fps_binomial_of_nat[of 0] by simp lemma fps_binomial_power: "fps_binomial a ^ n = fps_binomial (of_nat n * a)" by (induction n) (simp_all add: fps_binomial_add_mult ring_distribs) lemma fps_binomial_1: "fps_binomial 1 = 1 + fps_X" using fps_binomial_of_nat[of 1] by simp lemma fps_binomial_minus_of_nat: "fps_binomial (- of_nat n) = inverse ((1 + fps_X :: 'a :: field_char_0 fps) ^ n)" by (rule sym, rule fps_inverse_unique) (simp add: fps_binomial_of_nat [symmetric] fps_binomial_add_mult [symmetric]) lemma one_minus_const_fps_X_power: "c \ 0 \ (1 - fps_const c * fps_X) ^ n = fps_compose (fps_binomial (of_nat n)) (-fps_const c * fps_X)" by (subst fps_binomial_of_nat) (simp add: fps_compose_power [symmetric] fps_compose_add_distrib fps_const_neg [symmetric] del: fps_const_neg) lemma one_minus_fps_X_const_neg_power: "inverse ((1 - fps_const c * fps_X) ^ n) = fps_compose (fps_binomial (-of_nat n)) (-fps_const c * fps_X)" proof (cases "c = 0") case False thus ?thesis by (subst fps_binomial_minus_of_nat) (simp add: fps_compose_power [symmetric] fps_inverse_compose fps_compose_add_distrib fps_const_neg [symmetric] del: fps_const_neg) qed simp lemma fps_X_plus_const_power: "c \ 0 \ (fps_X + fps_const c) ^ n = fps_const (c^n) * fps_compose (fps_binomial (of_nat n)) (fps_const (inverse c) * fps_X)" by (subst fps_binomial_of_nat) (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib fps_const_power [symmetric] power_mult_distrib [symmetric] algebra_simps inverse_mult_eq_1' del: fps_const_power) lemma fps_X_plus_const_neg_power: "c \ 0 \ inverse ((fps_X + fps_const c) ^ n) = fps_const (inverse c^n) * fps_compose (fps_binomial (-of_nat n)) (fps_const (inverse c) * fps_X)" by (subst fps_binomial_minus_of_nat) (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib fps_const_power [symmetric] power_mult_distrib [symmetric] fps_inverse_compose algebra_simps fps_const_inverse [symmetric] fps_inverse_mult [symmetric] fps_inverse_power [symmetric] inverse_mult_eq_1' del: fps_const_power) lemma one_minus_const_fps_X_neg_power': fixes c :: "'a :: field_char_0" assumes "n > 0" shows "inverse ((1 - fps_const c * fps_X) ^ n) = Abs_fps (\k. of_nat ((n + k - 1) choose k) * c^k)" proof - have \
: "\j. Abs_fps (\na. (- c) ^ na * fps_binomial (- of_nat n) $ na) $ j = Abs_fps (\k. of_nat (n + k - 1 choose k) * c ^ k) $ j" using assms by (simp add: gbinomial_minus binomial_gbinomial of_nat_diff flip: power_mult_distrib mult.assoc) show ?thesis apply (rule fps_ext) using \
by (metis (no_types, lifting) one_minus_fps_X_const_neg_power fps_const_neg fps_compose_linear fps_nth_Abs_fps) qed text \Vandermonde's Identity as a consequence.\ lemma gbinomial_Vandermonde: "sum (\k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n" proof - let ?ba = "fps_binomial a" let ?bb = "fps_binomial b" let ?bab = "fps_binomial (a + b)" from fps_binomial_add_mult[of a b] have "?bab $ n = (?ba * ?bb)$n" by simp then show ?thesis by (simp add: fps_mult_nth) qed lemma binomial_Vandermonde: "sum (\k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n" using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n] by (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric] of_nat_sum[symmetric] of_nat_add[symmetric] of_nat_eq_iff) lemma binomial_Vandermonde_same: "sum (\k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n" using binomial_Vandermonde[of n n n, symmetric] unfolding mult_2 by (metis atMost_atLeast0 choose_square_sum mult_2) lemma Vandermonde_pochhammer_lemma: fixes a :: "'a::field_char_0" assumes b: "\j. j b \ of_nat j" shows "sum (\k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) / (of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} = pochhammer (- (a + b)) n / pochhammer (- b) n" (is "?l = ?r") proof - let ?m1 = "\m. (- 1 :: 'a) ^ m" let ?f = "\m. of_nat (fact m)" let ?p = "\(x::'a). pochhammer (- x)" from b have bn0: "?p b n \ 0" unfolding pochhammer_eq_0_iff by simp have th00: "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" (is ?gchoose) "pochhammer (1 + b - of_nat n) k \ 0" (is ?pochhammer) if kn: "k \ {0..n}" for k proof - from kn have "k \ n" by simp have nz: "pochhammer (1 + b - of_nat n) n \ 0" proof assume "pochhammer (1 + b - of_nat n) n = 0" then have c: "pochhammer (b - of_nat n + 1) n = 0" by (simp add: algebra_simps) then obtain j where j: "j < n" "b - of_nat n + 1 = - of_nat j" unfolding pochhammer_eq_0_iff by blast from j have "b = of_nat n - of_nat j - of_nat 1" by (simp add: algebra_simps) then have "b = of_nat (n - j - 1)" using j kn by (simp add: of_nat_diff) then show False using \j < n\ j b by auto qed from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \ 0" by (rule pochhammer_neq_0_mono) consider "k = 0 \ n = 0" | "k \ 0" "n \ 0" by blast then have "b gchoose (n - k) = (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)" proof cases case 1 then show ?thesis using kn by (cases "k = 0") (simp_all add: gbinomial_pochhammer) next case neq: 2 then obtain m where m: "n = Suc m" by (cases n) auto from neq(1) obtain h where h: "k = Suc h" by (cases k) auto show ?thesis proof (cases "k = n") case True with pochhammer_minus'[where k=k and b=b] bn0 show ?thesis by (simp add: pochhammer_same) next case False with kn have kn': "k < n" by simp have "h \ m" using \k \ n\ h m by blast have m1nk: "?m1 n = prod (\i. - 1) {..m}" "?m1 k = prod (\i. - 1) {0..h}" by (simp_all add: m h) have bnz0: "pochhammer (b - of_nat n + 1) k \ 0" using bn0 kn unfolding pochhammer_eq_0_iff by (metis add.commute add_diff_eq nz' pochhammer_eq_0_iff) have eq1: "prod (\k. (1::'a) + of_nat m - of_nat k) {..h} = prod of_nat {Suc (m - h) .. Suc m}" using kn' h m by (intro prod.reindex_bij_witness[where i="\k. Suc m - k" and j="\k. Suc m - k"]) (auto simp: of_nat_diff) have "(\i = 0..x = n - k..k \ n\ using prod.atLeastLessThan_shift_bounds [where ?'a = 'a, of "\i. 1 + of_nat i" 0 "n - k" k] by (auto simp add: of_nat_diff field_simps) then have "fact (n - k) * pochhammer ((1::'a) + of_nat n - of_nat k) k = fact n" using \k \ n\ by (auto simp add: fact_split [of k n] pochhammer_prod field_simps) then have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))" by (simp add: pochhammer_minus field_simps) have "?m1 n * ?p b n = pochhammer (b - of_nat m) (Suc m)" by (simp add: pochhammer_minus field_simps m) also have "... = (\i = 0..m. b - of_nat i)" by (auto simp add: pochhammer_prod_rev of_nat_diff prod.atLeast_Suc_atMost_Suc_shift simp del: prod.cl_ivl_Suc) finally have th20: "?m1 n * ?p b n = prod (\i. b - of_nat i) {0..m}" . have "(\x = 0..h. b - of_nat m + of_nat (h - x)) = (\i = m - h..m. b - of_nat i)" using \h \ m\ prod.atLeastAtMost_shift_0 [of "m - h" m, where ?'a = 'a] by (auto simp add: of_nat_diff field_simps) then have th21:"pochhammer (b - of_nat n + 1) k = prod (\i. b - of_nat i) {n - k .. n - 1}" using kn by (simp add: pochhammer_prod_rev m h prod.atLeast_Suc_atMost_Suc_shift del: prod.op_ivl_Suc del: prod.cl_ivl_Suc) have "?m1 n * ?p b n = prod (\i. b - of_nat i) {0.. n - k - 1} * pochhammer (b - of_nat n + 1) k" using kn' m h unfolding th20 th21 by (auto simp flip: prod.union_disjoint intro: prod.cong) then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k = prod (\i. b - of_nat i) {0.. n - k - 1}" using nz' by (simp add: field_simps) have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) = ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)" using bnz0 by (simp add: field_simps) also have "\ = b gchoose (n - k)" unfolding th1 th2 using kn' m h by (auto simp: field_simps gbinomial_mult_fact intro: prod.cong) finally show ?thesis by simp qed qed then show ?gchoose and ?pochhammer using nz' by force+ qed have "?r = ((a + b) gchoose n) * (of_nat (fact n) / (?m1 n * pochhammer (- b) n))" unfolding gbinomial_pochhammer using bn0 by (auto simp add: field_simps) also have "\ = ?l" using bn0 unfolding gbinomial_Vandermonde[symmetric] apply (simp add: th00) by (simp add: gbinomial_pochhammer sum_distrib_right sum_distrib_left field_simps) finally show ?thesis by simp qed lemma Vandermonde_pochhammer: fixes a :: "'a::field_char_0" assumes c: "\i \ {0..< n}. c \ - of_nat i" shows "sum (\k. (pochhammer a k * pochhammer (- (of_nat n)) k) / (of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n" proof - let ?a = "- a" let ?b = "c + of_nat n - 1" have h: "?b \ of_nat j" if "j < n" for j proof - have "c \ - of_nat (n - j - 1)" using c that by auto with that show ?thesis by (auto simp add: algebra_simps of_nat_diff) qed have th0: "pochhammer (- (?a + ?b)) n = (- 1)^n * pochhammer (c - a) n" unfolding pochhammer_minus by (simp add: algebra_simps) have th1: "pochhammer (- ?b) n = (- 1)^n * pochhammer c n" unfolding pochhammer_minus by simp have nz: "pochhammer c n \ 0" using c by (simp add: pochhammer_eq_0_iff) from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1] show ?thesis using nz by (simp add: field_simps sum_distrib_left) qed subsubsection \Formal trigonometric functions\ definition "fps_sin (c::'a::field_char_0) = Abs_fps (\n. if even n then 0 else (- 1) ^((n - 1) div 2) * c^n /(of_nat (fact n)))" definition "fps_cos (c::'a::field_char_0) = Abs_fps (\n. if even n then (- 1) ^ (n div 2) * c^n / (of_nat (fact n)) else 0)" lemma fps_sin_0 [simp]: "fps_sin 0 = 0" by (intro fps_ext) (auto simp: fps_sin_def elim!: oddE) lemma fps_cos_0 [simp]: "fps_cos 0 = 1" by (intro fps_ext) (simp add: fps_cos_def) lemma fps_sin_deriv: "fps_deriv (fps_sin c) = fps_const c * fps_cos c" (is "?lhs = ?rhs") proof (rule fps_ext) fix n :: nat show "?lhs $ n = ?rhs $ n" proof (cases "even n") case True have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp also have "\ = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))" using True by (simp add: fps_sin_def) also have "\ = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))" unfolding fact_Suc of_nat_mult by (simp add: field_simps del: of_nat_add of_nat_Suc) also have "\ = (- 1)^(n div 2) * c^Suc n / of_nat (fact n)" by (simp add: field_simps del: of_nat_add of_nat_Suc) finally show ?thesis using True by (simp add: fps_cos_def field_simps) next case False then show ?thesis by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def) qed qed lemma fps_cos_deriv: "fps_deriv (fps_cos c) = fps_const (- c)* (fps_sin c)" (is "?lhs = ?rhs") proof (rule fps_ext) have th0: "- ((- 1::'a) ^ n) = (- 1)^Suc n" for n by simp show "?lhs $ n = ?rhs $ n" for n proof (cases "even n") case False then have n0: "n \ 0" by presburger from False have th1: "Suc ((n - 1) div 2) = Suc n div 2" by (cases n) simp_all have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp also have "\ = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))" using False by (simp add: fps_cos_def) also have "\ = (- 1)^((n + 1) div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))" unfolding fact_Suc of_nat_mult by (simp add: field_simps del: of_nat_add of_nat_Suc) also have "\ = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)" by (simp add: field_simps del: of_nat_add of_nat_Suc) also have "\ = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)" unfolding th0 unfolding th1 by simp finally show ?thesis using False by (simp add: fps_sin_def field_simps) next case True then show ?thesis by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def) qed qed lemma fps_sin_cos_sum_of_squares: "(fps_cos c)\<^sup>2 + (fps_sin c)\<^sup>2 = 1" (is "?lhs = _") proof - have "fps_deriv ?lhs = 0" by (simp add: fps_deriv_power fps_sin_deriv fps_cos_deriv field_simps flip: fps_const_neg) then have "?lhs = fps_const (?lhs $ 0)" unfolding fps_deriv_eq_0_iff . also have "\ = 1" by (simp add: fps_eq_iff numeral_2_eq_2 fps_mult_nth fps_cos_def fps_sin_def) finally show ?thesis . qed lemma fps_sin_nth_0 [simp]: "fps_sin c $ 0 = 0" unfolding fps_sin_def by simp lemma fps_sin_nth_1 [simp]: "fps_sin c $ Suc 0 = c" unfolding fps_sin_def by simp lemma fps_sin_nth_add_2: "fps_sin c $ (n + 2) = - (c * c * fps_sin c $ n / (of_nat (n + 1) * of_nat (n + 2)))" proof (cases n) case (Suc n') then show ?thesis unfolding fps_sin_def by (simp add: field_simps) qed (auto simp: fps_sin_def) lemma fps_cos_nth_0 [simp]: "fps_cos c $ 0 = 1" unfolding fps_cos_def by simp lemma fps_cos_nth_1 [simp]: "fps_cos c $ Suc 0 = 0" unfolding fps_cos_def by simp lemma fps_cos_nth_add_2: "fps_cos c $ (n + 2) = - (c * c * fps_cos c $ n / (of_nat (n + 1) * of_nat (n + 2)))" proof (cases n) case (Suc n') then show ?thesis unfolding fps_cos_def by (simp add: field_simps) qed (auto simp: fps_cos_def) lemma nat_add_1_add_1: "(n::nat) + 1 + 1 = n + 2" by simp lemma eq_fps_sin: assumes a0: "a $ 0 = 0" and a1: "a $ 1 = c" and a2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)" shows "fps_sin c = a" proof (rule fps_ext) fix n show "fps_sin c $ n = a $ n" proof (induction n rule: nat_induct2) case (step n) then have "of_nat (n + 1) * (of_nat (n + 2) * a $ (n + 2)) = - (c * c * fps_sin c $ n)" using a2 by (metis fps_const_mult fps_deriv_nth fps_mult_left_const_nth fps_neg_nth nat_add_1_add_1) with step show ?case by (metis (no_types, lifting) a0 add.commute add.inverse_inverse fps_sin_nth_0 fps_sin_nth_add_2 mult_divide_mult_cancel_left_if mult_minus_right nonzero_mult_div_cancel_left not_less_zero of_nat_eq_0_iff plus_1_eq_Suc zero_less_Suc) qed (use assms in auto) qed lemma eq_fps_cos: assumes a0: "a $ 0 = 1" and a1: "a $ 1 = 0" and a2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)" shows "fps_cos c = a" proof (rule fps_ext) fix n show "fps_cos c $ n = a $ n" proof (induction n rule: nat_induct2) case (step n) then have "of_nat (n + 1) * (of_nat (n + 2) * a $ (n + 2)) = - (c * c * fps_cos c $ n)" using a2 by (metis fps_const_mult fps_deriv_nth fps_mult_left_const_nth fps_neg_nth nat_add_1_add_1) with step show ?case by (metis (no_types, lifting) a0 add.commute add.inverse_inverse fps_cos_nth_0 fps_cos_nth_add_2 mult_divide_mult_cancel_left_if mult_minus_right nonzero_mult_div_cancel_left not_less_zero of_nat_eq_0_iff plus_1_eq_Suc zero_less_Suc) qed (use assms in auto) qed lemma fps_sin_add: "fps_sin (a + b) = fps_sin a * fps_cos b + fps_cos a * fps_sin b" proof - have "fps_deriv (fps_deriv (fps_sin a * fps_cos b + fps_cos a * fps_sin b)) = - (fps_const (a + b) * fps_const (a + b) * (fps_sin a * fps_cos b + fps_cos a * fps_sin b))" by (simp flip: fps_const_neg fps_const_add fps_const_mult add: fps_sin_deriv fps_cos_deriv algebra_simps) then show ?thesis by (auto intro: eq_fps_sin) qed lemma fps_cos_add: "fps_cos (a + b) = fps_cos a * fps_cos b - fps_sin a * fps_sin b" proof - have "fps_deriv (fps_deriv (fps_cos a * fps_cos b - fps_sin a * fps_sin b)) = - (fps_const (a + b) * fps_const (a + b) * (fps_cos a * fps_cos b - fps_sin a * fps_sin b))" by (simp flip: fps_const_neg fps_const_add fps_const_mult add: fps_sin_deriv fps_cos_deriv algebra_simps) then show ?thesis by (auto intro: eq_fps_cos) qed lemma fps_sin_even: "fps_sin (- c) = - fps_sin c" by (simp add: fps_eq_iff fps_sin_def) lemma fps_cos_odd: "fps_cos (- c) = fps_cos c" by (simp add: fps_eq_iff fps_cos_def) definition "fps_tan c = fps_sin c / fps_cos c" lemma fps_tan_0 [simp]: "fps_tan 0 = 0" by (simp add: fps_tan_def) lemma fps_tan_deriv: "fps_deriv (fps_tan c) = fps_const c / (fps_cos c)\<^sup>2" proof - have th0: "fps_cos c $ 0 \ 0" by (simp add: fps_cos_def) from this have "fps_cos c \ 0" by (intro notI) simp hence "fps_deriv (fps_tan c) = fps_const c * (fps_cos c^2 + fps_sin c^2) / (fps_cos c^2)" by (simp add: fps_tan_def fps_divide_deriv power2_eq_square algebra_simps fps_sin_deriv fps_cos_deriv fps_const_neg[symmetric] div_mult_swap del: fps_const_neg) also note fps_sin_cos_sum_of_squares finally show ?thesis by simp qed text \Connection to @{const "fps_exp"} over the complex numbers --- Euler and de Moivre.\ lemma fps_exp_ii_sin_cos: "fps_exp (\ * c) = fps_cos c + fps_const \ * fps_sin c" (is "?l = ?r") proof - have "?l $ n = ?r $ n" for n proof (cases "even n") case True then obtain m where m: "n = 2 * m" .. show ?thesis by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"]) next case False then obtain m where m: "n = 2 * m + 1" .. show ?thesis by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"]) qed then show ?thesis by (simp add: fps_eq_iff) qed lemma fps_exp_minus_ii_sin_cos: "fps_exp (- (\ * c)) = fps_cos c - fps_const \ * fps_sin c" unfolding minus_mult_right fps_exp_ii_sin_cos by (simp add: fps_sin_even fps_cos_odd) lemma fps_cos_fps_exp_ii: "fps_cos c = (fps_exp (\ * c) + fps_exp (- \ * c)) / fps_const 2" proof - have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2" by (simp add: numeral_fps_const) show ?thesis unfolding fps_exp_ii_sin_cos minus_mult_commute by (simp add: fps_sin_even fps_cos_odd numeral_fps_const fps_divide_unit fps_const_inverse th) qed lemma fps_sin_fps_exp_ii: "fps_sin c = (fps_exp (\ * c) - fps_exp (- \ * c)) / fps_const (2*\)" proof - have th: "fps_const \ * fps_sin c + fps_const \ * fps_sin c = fps_sin c * fps_const (2 * \)" by (simp add: fps_eq_iff numeral_fps_const) show ?thesis unfolding fps_exp_ii_sin_cos minus_mult_commute by (simp add: fps_sin_even fps_cos_odd fps_divide_unit fps_const_inverse th) qed lemma fps_tan_fps_exp_ii: "fps_tan c = (fps_exp (\ * c) - fps_exp (- \ * c)) / (fps_const \ * (fps_exp (\ * c) + fps_exp (- \ * c)))" unfolding fps_tan_def fps_sin_fps_exp_ii fps_cos_fps_exp_ii by (simp add: fps_divide_unit fps_inverse_mult fps_const_inverse) lemma fps_demoivre: "(fps_cos a + fps_const \ * fps_sin a)^n = fps_cos (of_nat n * a) + fps_const \ * fps_sin (of_nat n * a)" unfolding fps_exp_ii_sin_cos[symmetric] fps_exp_power_mult by (simp add: ac_simps) subsection \Hypergeometric series\ definition "fps_hypergeo as bs (c::'a::field_char_0) = Abs_fps (\n. (foldl (\r a. r* pochhammer a n) 1 as * c^n) / (foldl (\r b. r * pochhammer b n) 1 bs * of_nat (fact n)))" lemma fps_hypergeo_nth[simp]: "fps_hypergeo as bs c $ n = (foldl (\r a. r* pochhammer a n) 1 as * c^n) / (foldl (\r b. r * pochhammer b n) 1 bs * of_nat (fact n))" by (simp add: fps_hypergeo_def) lemma foldl_mult_start: fixes v :: "'a::comm_ring_1" shows "foldl (\r x. r * f x) v as * x = foldl (\r x. r * f x) (v * x) as " by (induct as arbitrary: x v) (auto simp add: algebra_simps) lemma foldr_mult_foldl: fixes v :: "'a::comm_ring_1" shows "foldr (\x r. r * f x) as v = foldl (\r x. r * f x) v as" by (induct as arbitrary: v) (simp_all add: foldl_mult_start) lemma fps_hypergeo_nth_alt: "fps_hypergeo as bs c $ n = foldr (\a r. r * pochhammer a n) as (c ^ n) / foldr (\b r. r * pochhammer b n) bs (of_nat (fact n))" by (simp add: foldl_mult_start foldr_mult_foldl) lemma fps_hypergeo_fps_exp[simp]: "fps_hypergeo [] [] c = fps_exp c" by (simp add: fps_eq_iff) lemma fps_hypergeo_1_0[simp]: "fps_hypergeo [1] [] c = 1/(1 - fps_const c * fps_X)" proof - let ?a = "(Abs_fps (\n. 1)) oo (fps_const c * fps_X)" have th0: "(fps_const c * fps_X) $ 0 = 0" by simp show ?thesis unfolding gp[OF th0, symmetric] by (simp add: fps_eq_iff pochhammer_fact[symmetric] fps_compose_nth power_mult_distrib if_distrib cong del: if_weak_cong) qed lemma fps_hypergeo_B[simp]: "fps_hypergeo [-a] [] (- 1) = fps_binomial a" by (simp add: fps_eq_iff gbinomial_pochhammer algebra_simps) lemma fps_hypergeo_0[simp]: "fps_hypergeo as bs c $ 0 = 1" proof - have "foldl (\(r::'a) (a::'a). r) 1 as = 1" for as by (induction as) auto then show ?thesis by auto qed lemma foldl_prod_prod: "foldl (\(r::'b::comm_ring_1) (x::'a::comm_ring_1). r * f x) v as * foldl (\r x. r * g x) w as = foldl (\r x. r * f x * g x) (v * w) as" by (induct as arbitrary: v w) (simp_all add: algebra_simps) lemma fps_hypergeo_rec: "fps_hypergeo as bs c $ Suc n = ((foldl (\r a. r* (a + of_nat n)) c as) / (foldl (\r b. r * (b + of_nat n)) (of_nat (Suc n)) bs )) * fps_hypergeo as bs c $ n" apply (simp add: foldl_mult_start del: of_nat_Suc of_nat_add fact_Suc) unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc by (simp add: algebra_simps) lemma fps_XD_nth[simp]: "fps_XD a $ n = of_nat n * a$n" by (simp add: fps_XD_def) lemma fps_XD_0th[simp]: "fps_XD a $ 0 = 0" by simp lemma fps_XD_Suc[simp]:" fps_XD a $ Suc n = of_nat (Suc n) * a $ Suc n" by simp definition "fps_XDp c a = fps_XD a + fps_const c * a" lemma fps_XDp_nth[simp]: "fps_XDp c a $ n = (c + of_nat n) * a$n" by (simp add: fps_XDp_def algebra_simps) lemma fps_XDp_commute: "fps_XDp b \ fps_XDp (c::'a::comm_ring_1) = fps_XDp c \ fps_XDp b" by (simp add: fps_XDp_def fun_eq_iff fps_eq_iff algebra_simps) lemma fps_XDp0 [simp]: "fps_XDp 0 = fps_XD" by (simp add: fun_eq_iff fps_eq_iff) lemma fps_XDp_fps_integral [simp]: fixes a :: "'a::{division_ring,ring_char_0} fps" shows "fps_XDp 0 (fps_integral a c) = fps_X * a" using fps_deriv_fps_integral[of a c] by (simp add: fps_XD_def) lemma fps_hypergeo_minus_nat: "fps_hypergeo [- of_nat n] [- of_nat (n + m)] (c::'a::field_char_0) $ k = (if k \ n then pochhammer (- of_nat n) k * c ^ k / (pochhammer (- of_nat (n + m)) k * of_nat (fact k)) else 0)" "fps_hypergeo [- of_nat m] [- of_nat (m + n)] (c::'a::field_char_0) $ k = (if k \ m then pochhammer (- of_nat m) k * c ^ k / (pochhammer (- of_nat (m + n)) k * of_nat (fact k)) else 0)" by (simp_all add: pochhammer_eq_0_iff) lemma pochhammer_rec_if: "pochhammer a n = (if n = 0 then 1 else a * pochhammer (a + 1) (n - 1))" by (cases n) (simp_all add: pochhammer_rec) lemma fps_XDp_foldr_nth [simp]: "foldr (\c r. fps_XDp c \ r) cs (\c. fps_XDp c a) c0 $ n = foldr (\c r. (c + of_nat n) * r) cs (c0 + of_nat n) * a$n" by (induct cs arbitrary: c0) (simp_all add: algebra_simps) lemma genric_fps_XDp_foldr_nth: assumes f: "\n c a. f c a $ n = (of_nat n + k c) * a$n" shows "foldr (\c r. f c \ r) cs (\c. g c a) c0 $ n = foldr (\c r. (k c + of_nat n) * r) cs (g c0 a $ n)" by (induct cs arbitrary: c0) (simp_all add: algebra_simps f) lemma dist_less_imp_nth_equal: assumes "dist f g < inverse (2 ^ i)" and"j \ i" shows "f $ j = g $ j" proof (rule ccontr) assume "f $ j \ g $ j" hence "f \ g" by auto with assms have "i < subdegree (f - g)" by (simp add: if_split_asm dist_fps_def) also have "\ \ j" using \f $ j \ g $ j\ by (intro subdegree_leI) simp_all finally show False using \j \ i\ by simp qed lemma nth_equal_imp_dist_less: assumes "\j. j \ i \ f $ j = g $ j" shows "dist f g < inverse (2 ^ i)" proof (cases "f = g") case True then show ?thesis by simp next case False with assms have "dist f g = inverse (2 ^ subdegree (f - g))" by (simp add: if_split_asm dist_fps_def) moreover from assms and False have "i < subdegree (f - g)" by (intro subdegree_greaterI) simp_all ultimately show ?thesis by simp qed lemma dist_less_eq_nth_equal: "dist f g < inverse (2 ^ i) \ (\j \ i. f $ j = g $ j)" using dist_less_imp_nth_equal nth_equal_imp_dist_less by blast instance fps :: (comm_ring_1) complete_space proof fix fps_X :: "nat \ 'a fps" assume "Cauchy fps_X" obtain M where M: "\i. \m \ M i. \j \ i. fps_X (M i) $ j = fps_X m $ j" proof - have "\M. \m \ M. \j\i. fps_X M $ j = fps_X m $ j" for i proof - have "0 < inverse ((2::real)^i)" by simp from metric_CauchyD[OF \Cauchy fps_X\ this] dist_less_imp_nth_equal show ?thesis by blast qed then show ?thesis using that by metis qed show "convergent fps_X" proof (rule convergentI) show "fps_X \ Abs_fps (\i. fps_X (M i) $ i)" unfolding tendsto_iff proof safe fix e::real assume e: "0 < e" have "(\n. inverse (2 ^ n) :: real) \ 0" by (rule LIMSEQ_inverse_realpow_zero) simp_all from this and e have "eventually (\i. inverse (2 ^ i) < e) sequentially" by (rule order_tendstoD) then obtain i where "inverse (2 ^ i) < e" by (auto simp: eventually_sequentially) have "eventually (\x. M i \ x) sequentially" by (auto simp: eventually_sequentially) then show "eventually (\x. dist (fps_X x) (Abs_fps (\i. fps_X (M i) $ i)) < e) sequentially" proof eventually_elim fix x assume x: "M i \ x" have "fps_X (M i) $ j = fps_X (M j) $ j" if "j \ i" for j using M that by (metis nat_le_linear) with x have "dist (fps_X x) (Abs_fps (\j. fps_X (M j) $ j)) < inverse (2 ^ i)" using M by (force simp: dist_less_eq_nth_equal) also note \inverse (2 ^ i) < e\ finally show "dist (fps_X x) (Abs_fps (\j. fps_X (M j) $ j)) < e" . qed qed qed qed (* TODO: Figure out better notation for this thing *) no_notation fps_nth (infixl "$" 75) bundle fps_notation begin notation fps_nth (infixl "$" 75) end end diff --git a/src/HOL/Computational_Algebra/Normalized_Fraction.thy b/src/HOL/Computational_Algebra/Normalized_Fraction.thy --- a/src/HOL/Computational_Algebra/Normalized_Fraction.thy +++ b/src/HOL/Computational_Algebra/Normalized_Fraction.thy @@ -1,364 +1,416 @@ (* Title: HOL/Computational_Algebra/Normalized_Fraction.thy Author: Manuel Eberl *) theory Normalized_Fraction imports Main Euclidean_Algorithm Fraction_Field begin definition quot_to_fract :: "'a :: {idom} \ 'a \ 'a fract" where "quot_to_fract = (\(a,b). Fraction_Field.Fract a b)" definition normalize_quot :: "'a :: {ring_gcd,idom_divide,semiring_gcd_mult_normalize} \ 'a \ 'a \ 'a" where "normalize_quot = (\(a,b). if b = 0 then (0,1) else let d = gcd a b * unit_factor b in (a div d, b div d))" lemma normalize_quot_zero [simp]: "normalize_quot (a, 0) = (0, 1)" by (simp add: normalize_quot_def) lemma normalize_quot_proj: "fst (normalize_quot (a, b)) = a div (gcd a b * unit_factor b)" "snd (normalize_quot (a, b)) = normalize b div gcd a b" if "b \ 0" using that by (simp_all add: normalize_quot_def Let_def mult.commute [of _ "unit_factor b"] dvd_div_mult2_eq mult_unit_dvd_iff') definition normalized_fracts :: "('a :: {ring_gcd,idom_divide} \ 'a) set" where "normalized_fracts = {(a,b). coprime a b \ unit_factor b = 1}" lemma not_normalized_fracts_0_denom [simp]: "(a, 0) \ normalized_fracts" by (auto simp: normalized_fracts_def) lemma unit_factor_snd_normalize_quot [simp]: "unit_factor (snd (normalize_quot x)) = 1" by (simp add: normalize_quot_def case_prod_unfold Let_def dvd_unit_factor_div mult_unit_dvd_iff unit_factor_mult unit_factor_gcd) lemma snd_normalize_quot_nonzero [simp]: "snd (normalize_quot x) \ 0" using unit_factor_snd_normalize_quot[of x] by (auto simp del: unit_factor_snd_normalize_quot) lemma normalize_quot_aux: fixes a b assumes "b \ 0" defines "d \ gcd a b * unit_factor b" shows "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d dvd a" "d dvd b" "d \ 0" proof - from assms show "d dvd a" "d dvd b" by (simp_all add: d_def mult_unit_dvd_iff) thus "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d \ 0" by (auto simp: normalize_quot_def Let_def d_def \b \ 0\) qed lemma normalize_quotE: assumes "b \ 0" obtains d where "a = fst (normalize_quot (a,b)) * d" "b = snd (normalize_quot (a,b)) * d" "d dvd a" "d dvd b" "d \ 0" using that[OF normalize_quot_aux[OF assms]] . lemma normalize_quotE': assumes "snd x \ 0" obtains d where "fst x = fst (normalize_quot x) * d" "snd x = snd (normalize_quot x) * d" "d dvd fst x" "d dvd snd x" "d \ 0" proof - - from normalize_quotE[OF assms, of "fst x"] guess d . - from this show ?thesis unfolding prod.collapse by (intro that[of d]) + from normalize_quotE[OF assms, of "fst x"] obtain d where + "fst x = fst (normalize_quot (fst x, snd x)) * d" + "snd x = snd (normalize_quot (fst x, snd x)) * d" + "d dvd fst x" + "d dvd snd x" + "d \ 0" . + then show ?thesis unfolding prod.collapse by (intro that[of d]) qed lemma coprime_normalize_quot: "coprime (fst (normalize_quot x)) (snd (normalize_quot x))" by (simp add: normalize_quot_def case_prod_unfold div_mult_unit2) (metis coprime_mult_self_right_iff div_gcd_coprime unit_div_mult_self unit_factor_is_unit) lemma normalize_quot_in_normalized_fracts [simp]: "normalize_quot x \ normalized_fracts" by (simp add: normalized_fracts_def coprime_normalize_quot case_prod_unfold) lemma normalize_quot_eq_iff: assumes "b \ 0" "d \ 0" shows "normalize_quot (a,b) = normalize_quot (c,d) \ a * d = b * c" proof - define x y where "x = normalize_quot (a,b)" and "y = normalize_quot (c,d)" from normalize_quotE[OF assms(1), of a] normalize_quotE[OF assms(2), of c] obtain d1 d2 where "a = fst x * d1" "b = snd x * d1" "c = fst y * d2" "d = snd y * d2" "d1 \ 0" "d2 \ 0" unfolding x_def y_def by metis hence "a * d = b * c \ fst x * snd y = snd x * fst y" by simp also have "\ \ fst x = fst y \ snd x = snd y" by (intro coprime_crossproduct') (simp_all add: x_def y_def coprime_normalize_quot) also have "\ \ x = y" using prod_eqI by blast finally show "x = y \ a * d = b * c" .. qed lemma normalize_quot_eq_iff': assumes "snd x \ 0" "snd y \ 0" shows "normalize_quot x = normalize_quot y \ fst x * snd y = snd x * fst y" using assms by (cases x, cases y, hypsubst) (subst normalize_quot_eq_iff, simp_all) lemma normalize_quot_id: "x \ normalized_fracts \ normalize_quot x = x" by (auto simp: normalized_fracts_def normalize_quot_def case_prod_unfold) lemma normalize_quot_idem [simp]: "normalize_quot (normalize_quot x) = normalize_quot x" by (rule normalize_quot_id) simp_all lemma fractrel_iff_normalize_quot_eq: "fractrel x y \ normalize_quot x = normalize_quot y \ snd x \ 0 \ snd y \ 0" by (cases x, cases y) (auto simp: fractrel_def normalize_quot_eq_iff) lemma fractrel_normalize_quot_left: assumes "snd x \ 0" shows "fractrel (normalize_quot x) y \ fractrel x y" using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto lemma fractrel_normalize_quot_right: assumes "snd x \ 0" shows "fractrel y (normalize_quot x) \ fractrel y x" using assms by (subst (1 2) fractrel_iff_normalize_quot_eq) auto lift_definition quot_of_fract :: "'a :: {ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract \ 'a \ 'a" is normalize_quot by (subst (asm) fractrel_iff_normalize_quot_eq) simp_all lemma quot_to_fract_quot_of_fract [simp]: "quot_to_fract (quot_of_fract x) = x" unfolding quot_to_fract_def proof transfer fix x :: "'a \ 'a" assume rel: "fractrel x x" define x' where "x' = normalize_quot x" obtain a b where [simp]: "x = (a, b)" by (cases x) from rel have "b \ 0" by simp - from normalize_quotE[OF this, of a] guess d . + from normalize_quotE[OF this, of a] obtain d + where + "a = fst (normalize_quot (a, b)) * d" + "b = snd (normalize_quot (a, b)) * d" + "d dvd a" + "d dvd b" + "d \ 0" . hence "a = fst x' * d" "b = snd x' * d" "d \ 0" "snd x' \ 0" by (simp_all add: x'_def) thus "fractrel (case x' of (a, b) \ if b = 0 then (0, 1) else (a, b)) x" by (auto simp add: case_prod_unfold) qed lemma quot_of_fract_quot_to_fract: "quot_of_fract (quot_to_fract x) = normalize_quot x" proof (cases "snd x = 0") case True thus ?thesis unfolding quot_to_fract_def by transfer (simp add: case_prod_unfold normalize_quot_def) next case False thus ?thesis unfolding quot_to_fract_def by transfer (simp add: case_prod_unfold) qed lemma quot_of_fract_quot_to_fract': "x \ normalized_fracts \ quot_of_fract (quot_to_fract x) = x" unfolding quot_to_fract_def by transfer (auto simp: normalize_quot_id) lemma quot_of_fract_in_normalized_fracts [simp]: "quot_of_fract x \ normalized_fracts" by transfer simp lemma normalize_quotI: assumes "a * d = b * c" "b \ 0" "(c, d) \ normalized_fracts" shows "normalize_quot (a, b) = (c, d)" proof - from assms have "normalize_quot (a, b) = normalize_quot (c, d)" by (subst normalize_quot_eq_iff) auto also have "\ = (c, d)" by (intro normalize_quot_id) fact finally show ?thesis . qed lemma td_normalized_fract: "type_definition quot_of_fract quot_to_fract normalized_fracts" by standard (simp_all add: quot_of_fract_quot_to_fract') lemma quot_of_fract_add_aux: assumes "snd x \ 0" "snd y \ 0" shows "(fst x * snd y + fst y * snd x) * (snd (normalize_quot x) * snd (normalize_quot y)) = snd x * snd y * (fst (normalize_quot x) * snd (normalize_quot y) + snd (normalize_quot x) * fst (normalize_quot y))" proof - - from normalize_quotE'[OF assms(1)] guess d . note d = this - from normalize_quotE'[OF assms(2)] guess e . note e = this + from normalize_quotE'[OF assms(1)] obtain d + where d: + "fst x = fst (normalize_quot x) * d" + "snd x = snd (normalize_quot x) * d" + "d dvd fst x" + "d dvd snd x" + "d \ 0" . + from normalize_quotE'[OF assms(2)] obtain e + where e: + "fst y = fst (normalize_quot y) * e" + "snd y = snd (normalize_quot y) * e" + "e dvd fst y" + "e dvd snd y" + "e \ 0" . show ?thesis by (simp_all add: d e algebra_simps) qed locale fract_as_normalized_quot begin setup_lifting td_normalized_fract end lemma quot_of_fract_add: "quot_of_fract (x + y) = (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y in normalize_quot (a * d + b * c, b * d))" by transfer (insert quot_of_fract_add_aux, simp_all add: Let_def case_prod_unfold normalize_quot_eq_iff) lemma quot_of_fract_uminus: "quot_of_fract (-x) = (let (a,b) = quot_of_fract x in (-a, b))" by transfer (auto simp: case_prod_unfold Let_def normalize_quot_def dvd_neg_div mult_unit_dvd_iff) lemma quot_of_fract_diff: "quot_of_fract (x - y) = (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y in normalize_quot (a * d - b * c, b * d))" (is "_ = ?rhs") proof - have "x - y = x + -y" by simp also have "quot_of_fract \ = ?rhs" by (simp only: quot_of_fract_add quot_of_fract_uminus Let_def case_prod_unfold) simp_all finally show ?thesis . qed lemma normalize_quot_mult_coprime: assumes "coprime a b" "coprime c d" "unit_factor b = 1" "unit_factor d = 1" defines "e \ fst (normalize_quot (a, d))" and "f \ snd (normalize_quot (a, d))" and "g \ fst (normalize_quot (c, b))" and "h \ snd (normalize_quot (c, b))" shows "normalize_quot (a * c, b * d) = (e * g, f * h)" proof (rule normalize_quotI) from assms have "gcd a b = 1" "gcd c d = 1" by simp_all from assms have "b \ 0" "d \ 0" by auto with assms have "normalize b = b" "normalize d = d" by (auto intro: normalize_unit_factor_eqI) - from normalize_quotE [OF \b \ 0\, of c] guess k . + from normalize_quotE [OF \b \ 0\, of c] obtain k + where + "c = fst (normalize_quot (c, b)) * k" + "b = snd (normalize_quot (c, b)) * k" + "k dvd c" "k dvd b" "k \ 0" . note k = this [folded \gcd a b = 1\ \gcd c d = 1\ assms(3) assms(4)] - from normalize_quotE [OF \d \ 0\, of a] guess l . + from normalize_quotE [OF \d \ 0\, of a] obtain l + where "a = fst (normalize_quot (a, d)) * l" + "d = snd (normalize_quot (a, d)) * l" + "l dvd a" "l dvd d" "l \ 0" . note l = this [folded \gcd a b = 1\ \gcd c d = 1\ assms(3) assms(4)] from k l show "a * c * (f * h) = b * d * (e * g)" by (metis e_def f_def g_def h_def mult.commute mult.left_commute) from assms have [simp]: "unit_factor f = 1" "unit_factor h = 1" by simp_all from assms have "coprime e f" "coprime g h" by (simp_all add: coprime_normalize_quot) with k l assms(1,2) \b \ 0\ \d \ 0\ \unit_factor b = 1\ \unit_factor d = 1\ \normalize b = b\ \normalize d = d\ show "(e * g, f * h) \ normalized_fracts" by (simp add: normalized_fracts_def unit_factor_mult e_def f_def g_def h_def coprime_normalize_quot dvd_unit_factor_div unit_factor_gcd) (metis coprime_mult_left_iff coprime_mult_right_iff) qed (insert assms(3,4), auto) lemma normalize_quot_mult: assumes "snd x \ 0" "snd y \ 0" shows "normalize_quot (fst x * fst y, snd x * snd y) = normalize_quot (fst (normalize_quot x) * fst (normalize_quot y), snd (normalize_quot x) * snd (normalize_quot y))" proof - - from normalize_quotE'[OF assms(1)] guess d . note d = this - from normalize_quotE'[OF assms(2)] guess e . note e = this + from normalize_quotE'[OF assms(1)] obtain d where d: + "fst x = fst (normalize_quot x) * d" + "snd x = snd (normalize_quot x) * d" + "d dvd fst x" + "d dvd snd x" + "d \ 0" . + from normalize_quotE'[OF assms(2)] obtain e where e: + "fst y = fst (normalize_quot y) * e" + "snd y = snd (normalize_quot y) * e" + "e dvd fst y" + "e dvd snd y" + "e \ 0" . show ?thesis by (simp_all add: d e algebra_simps normalize_quot_eq_iff) qed lemma quot_of_fract_mult: "quot_of_fract (x * y) = (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y; (e,f) = normalize_quot (a,d); (g,h) = normalize_quot (c,b) in (e*g, f*h))" by transfer (simp add: split_def Let_def coprime_normalize_quot normalize_quot_mult normalize_quot_mult_coprime) lemma normalize_quot_0 [simp]: "normalize_quot (0, x) = (0, 1)" "normalize_quot (x, 0) = (0, 1)" by (simp_all add: normalize_quot_def) lemma normalize_quot_eq_0_iff [simp]: "fst (normalize_quot x) = 0 \ fst x = 0 \ snd x = 0" by (auto simp: normalize_quot_def case_prod_unfold Let_def div_mult_unit2 dvd_div_eq_0_iff) lemma fst_quot_of_fract_0_imp: "fst (quot_of_fract x) = 0 \ snd (quot_of_fract x) = 1" by transfer auto lemma normalize_quot_swap: assumes "a \ 0" "b \ 0" defines "a' \ fst (normalize_quot (a, b))" and "b' \ snd (normalize_quot (a, b))" shows "normalize_quot (b, a) = (b' div unit_factor a', a' div unit_factor a')" proof (rule normalize_quotI) - from normalize_quotE[OF assms(2), of a] guess d . note d = this [folded assms(3,4)] + from normalize_quotE[OF assms(2), of a] obtain d where + "a = fst (normalize_quot (a, b)) * d" + "b = snd (normalize_quot (a, b)) * d" + "d dvd a" "d dvd b" "d \ 0" . + note d = this [folded assms(3,4)] show "b * (a' div unit_factor a') = a * (b' div unit_factor a')" using assms(1,2) d by (simp add: div_unit_factor [symmetric] unit_div_mult_swap mult_ac del: div_unit_factor) have "coprime a' b'" by (simp add: a'_def b'_def coprime_normalize_quot) thus "(b' div unit_factor a', a' div unit_factor a') \ normalized_fracts" using assms(1,2) d by (auto simp add: normalized_fracts_def ac_simps dvd_div_unit_iff elim: coprime_imp_coprime) qed fact+ lemma quot_of_fract_inverse: "quot_of_fract (inverse x) = (let (a,b) = quot_of_fract x; d = unit_factor a in if d = 0 then (0, 1) else (b div d, a div d))" proof (transfer, goal_cases) case (1 x) from normalize_quot_swap[of "fst x" "snd x"] show ?case by (auto simp: Let_def case_prod_unfold) qed lemma normalize_quot_div_unit_left: fixes x y u assumes "is_unit u" defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (x div u, y) = (x' div u, y')" proof (cases "y = 0") case False define v where "v = 1 div u" with \is_unit u\ have "is_unit v" and u: "\a. a div u = a * v" by simp_all from \is_unit v\ have "coprime v = top" by (simp add: fun_eq_iff is_unit_left_imp_coprime) - from normalize_quotE[OF False, of x] guess d . + from normalize_quotE[OF False, of x] obtain d where + "x = fst (normalize_quot (x, y)) * d" + "y = snd (normalize_quot (x, y)) * d" + "d dvd x" "d dvd y" "d \ 0" . note d = this[folded assms(2,3)] from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot) with d \coprime v = top\ have "normalize_quot (x * v, y) = (x' * v, y')" by (auto simp: normalized_fracts_def intro: normalize_quotI) then show ?thesis by (simp add: u) qed (simp_all add: assms) lemma normalize_quot_div_unit_right: fixes x y u assumes "is_unit u" defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (x, y div u) = (x' * u, y')" proof (cases "y = 0") case False - from normalize_quotE[OF this, of x] guess d . note d = this[folded assms(2,3)] + from normalize_quotE[OF this, of x] + obtain d where d: + "x = fst (normalize_quot (x, y)) * d" + "y = snd (normalize_quot (x, y)) * d" + "d dvd x" "d dvd y" "d \ 0" . + note d = this[folded assms(2,3)] from assms have "coprime x' y'" "unit_factor y' = 1" by (simp_all add: coprime_normalize_quot) with d \is_unit u\ show ?thesis by (auto simp add: normalized_fracts_def is_unit_left_imp_coprime unit_div_eq_0_iff intro: normalize_quotI) qed (simp_all add: assms) lemma normalize_quot_normalize_left: fixes x y u defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (normalize x, y) = (x' div unit_factor x, y')" using normalize_quot_div_unit_left[of "unit_factor x" x y] by (cases "x = 0") (simp_all add: assms) lemma normalize_quot_normalize_right: fixes x y u defines "x' \ fst (normalize_quot (x, y))" and "y' \ snd (normalize_quot (x, y))" shows "normalize_quot (x, normalize y) = (x' * unit_factor y, y')" using normalize_quot_div_unit_right[of "unit_factor y" x y] by (cases "y = 0") (simp_all add: assms) lemma quot_of_fract_0 [simp]: "quot_of_fract 0 = (0, 1)" by transfer auto lemma quot_of_fract_1 [simp]: "quot_of_fract 1 = (1, 1)" by transfer (rule normalize_quotI, simp_all add: normalized_fracts_def) lemma quot_of_fract_divide: "quot_of_fract (x / y) = (if y = 0 then (0, 1) else (let (a,b) = quot_of_fract x; (c,d) = quot_of_fract y; (e,f) = normalize_quot (a,c); (g,h) = normalize_quot (d,b) in (e * g, f * h)))" (is "_ = ?rhs") proof (cases "y = 0") case False hence A: "fst (quot_of_fract y) \ 0" by transfer auto have "x / y = x * inverse y" by (simp add: divide_inverse) also from False A have "quot_of_fract \ = ?rhs" by (simp only: quot_of_fract_mult quot_of_fract_inverse) (simp_all add: Let_def case_prod_unfold fst_quot_of_fract_0_imp normalize_quot_div_unit_left normalize_quot_div_unit_right normalize_quot_normalize_right normalize_quot_normalize_left) finally show ?thesis . qed simp_all end diff --git a/src/HOL/Computational_Algebra/Polynomial.thy b/src/HOL/Computational_Algebra/Polynomial.thy --- a/src/HOL/Computational_Algebra/Polynomial.thy +++ b/src/HOL/Computational_Algebra/Polynomial.thy @@ -1,5019 +1,5022 @@ (* Title: HOL/Computational_Algebra/Polynomial.thy Author: Brian Huffman Author: Clemens Ballarin Author: Amine Chaieb Author: Florian Haftmann *) section \Polynomials as type over a ring structure\ theory Polynomial imports Complex_Main "HOL-Library.More_List" "HOL-Library.Infinite_Set" Factorial_Ring begin subsection \Auxiliary: operations for lists (later) representing coefficients\ definition cCons :: "'a::zero \ 'a list \ 'a list" (infixr "##" 65) where "x ## xs = (if xs = [] \ x = 0 then [] else x # xs)" lemma cCons_0_Nil_eq [simp]: "0 ## [] = []" by (simp add: cCons_def) lemma cCons_Cons_eq [simp]: "x ## y # ys = x # y # ys" by (simp add: cCons_def) lemma cCons_append_Cons_eq [simp]: "x ## xs @ y # ys = x # xs @ y # ys" by (simp add: cCons_def) lemma cCons_not_0_eq [simp]: "x \ 0 \ x ## xs = x # xs" by (simp add: cCons_def) lemma strip_while_not_0_Cons_eq [simp]: "strip_while (\x. x = 0) (x # xs) = x ## strip_while (\x. x = 0) xs" proof (cases "x = 0") case False then show ?thesis by simp next case True show ?thesis proof (induct xs rule: rev_induct) case Nil with True show ?case by simp next case (snoc y ys) then show ?case by (cases "y = 0") (simp_all add: append_Cons [symmetric] del: append_Cons) qed qed lemma tl_cCons [simp]: "tl (x ## xs) = xs" by (simp add: cCons_def) subsection \Definition of type \poly\\ typedef (overloaded) 'a poly = "{f :: nat \ 'a::zero. \\<^sub>\ n. f n = 0}" morphisms coeff Abs_poly by (auto intro!: ALL_MOST) setup_lifting type_definition_poly lemma poly_eq_iff: "p = q \ (\n. coeff p n = coeff q n)" by (simp add: coeff_inject [symmetric] fun_eq_iff) lemma poly_eqI: "(\n. coeff p n = coeff q n) \ p = q" by (simp add: poly_eq_iff) lemma MOST_coeff_eq_0: "\\<^sub>\ n. coeff p n = 0" using coeff [of p] by simp subsection \Degree of a polynomial\ definition degree :: "'a::zero poly \ nat" where "degree p = (LEAST n. \i>n. coeff p i = 0)" lemma coeff_eq_0: assumes "degree p < n" shows "coeff p n = 0" proof - have "\n. \i>n. coeff p i = 0" using MOST_coeff_eq_0 by (simp add: MOST_nat) then have "\i>degree p. coeff p i = 0" unfolding degree_def by (rule LeastI_ex) with assms show ?thesis by simp qed lemma le_degree: "coeff p n \ 0 \ n \ degree p" by (erule contrapos_np, rule coeff_eq_0, simp) lemma degree_le: "\i>n. coeff p i = 0 \ degree p \ n" unfolding degree_def by (erule Least_le) lemma less_degree_imp: "n < degree p \ \i>n. coeff p i \ 0" unfolding degree_def by (drule not_less_Least, simp) subsection \The zero polynomial\ instantiation poly :: (zero) zero begin lift_definition zero_poly :: "'a poly" is "\_. 0" by (rule MOST_I) simp instance .. end lemma coeff_0 [simp]: "coeff 0 n = 0" by transfer rule lemma degree_0 [simp]: "degree 0 = 0" by (rule order_antisym [OF degree_le le0]) simp lemma leading_coeff_neq_0: assumes "p \ 0" shows "coeff p (degree p) \ 0" proof (cases "degree p") case 0 from \p \ 0\ obtain n where "coeff p n \ 0" by (auto simp add: poly_eq_iff) then have "n \ degree p" by (rule le_degree) with \coeff p n \ 0\ and \degree p = 0\ show "coeff p (degree p) \ 0" by simp next case (Suc n) from \degree p = Suc n\ have "n < degree p" by simp then have "\i>n. coeff p i \ 0" by (rule less_degree_imp) then obtain i where "n < i" and "coeff p i \ 0" by blast from \degree p = Suc n\ and \n < i\ have "degree p \ i" by simp also from \coeff p i \ 0\ have "i \ degree p" by (rule le_degree) finally have "degree p = i" . with \coeff p i \ 0\ show "coeff p (degree p) \ 0" by simp qed lemma leading_coeff_0_iff [simp]: "coeff p (degree p) = 0 \ p = 0" by (cases "p = 0") (simp_all add: leading_coeff_neq_0) lemma eq_zero_or_degree_less: assumes "degree p \ n" and "coeff p n = 0" shows "p = 0 \ degree p < n" proof (cases n) case 0 with \degree p \ n\ and \coeff p n = 0\ have "coeff p (degree p) = 0" by simp then have "p = 0" by simp then show ?thesis .. next case (Suc m) from \degree p \ n\ have "\i>n. coeff p i = 0" by (simp add: coeff_eq_0) with \coeff p n = 0\ have "\i\n. coeff p i = 0" by (simp add: le_less) with \n = Suc m\ have "\i>m. coeff p i = 0" by (simp add: less_eq_Suc_le) then have "degree p \ m" by (rule degree_le) with \n = Suc m\ have "degree p < n" by (simp add: less_Suc_eq_le) then show ?thesis .. qed lemma coeff_0_degree_minus_1: "coeff rrr dr = 0 \ degree rrr \ dr \ degree rrr \ dr - 1" using eq_zero_or_degree_less by fastforce subsection \List-style constructor for polynomials\ lift_definition pCons :: "'a::zero \ 'a poly \ 'a poly" is "\a p. case_nat a (coeff p)" by (rule MOST_SucD) (simp add: MOST_coeff_eq_0) lemmas coeff_pCons = pCons.rep_eq lemma coeff_pCons_0 [simp]: "coeff (pCons a p) 0 = a" by transfer simp lemma coeff_pCons_Suc [simp]: "coeff (pCons a p) (Suc n) = coeff p n" by (simp add: coeff_pCons) lemma degree_pCons_le: "degree (pCons a p) \ Suc (degree p)" by (rule degree_le) (simp add: coeff_eq_0 coeff_pCons split: nat.split) lemma degree_pCons_eq: "p \ 0 \ degree (pCons a p) = Suc (degree p)" by (simp add: degree_pCons_le le_antisym le_degree) lemma degree_pCons_0: "degree (pCons a 0) = 0" proof - have "degree (pCons a 0) \ Suc 0" by (metis (no_types) degree_0 degree_pCons_le) then show ?thesis by (metis coeff_0 coeff_pCons_Suc degree_0 eq_zero_or_degree_less less_Suc0) qed lemma degree_pCons_eq_if [simp]: "degree (pCons a p) = (if p = 0 then 0 else Suc (degree p))" by (simp add: degree_pCons_0 degree_pCons_eq) lemma pCons_0_0 [simp]: "pCons 0 0 = 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma pCons_eq_iff [simp]: "pCons a p = pCons b q \ a = b \ p = q" proof safe assume "pCons a p = pCons b q" then have "coeff (pCons a p) 0 = coeff (pCons b q) 0" by simp then show "a = b" by simp next assume "pCons a p = pCons b q" then have "coeff (pCons a p) (Suc n) = coeff (pCons b q) (Suc n)" for n by simp then show "p = q" by (simp add: poly_eq_iff) qed lemma pCons_eq_0_iff [simp]: "pCons a p = 0 \ a = 0 \ p = 0" using pCons_eq_iff [of a p 0 0] by simp lemma pCons_cases [cases type: poly]: obtains (pCons) a q where "p = pCons a q" proof show "p = pCons (coeff p 0) (Abs_poly (\n. coeff p (Suc n)))" by transfer (simp_all add: MOST_inj[where f=Suc and P="\n. p n = 0" for p] fun_eq_iff Abs_poly_inverse split: nat.split) qed lemma pCons_induct [case_names 0 pCons, induct type: poly]: assumes zero: "P 0" assumes pCons: "\a p. a \ 0 \ p \ 0 \ P p \ P (pCons a p)" shows "P p" proof (induct p rule: measure_induct_rule [where f=degree]) case (less p) obtain a q where "p = pCons a q" by (rule pCons_cases) have "P q" proof (cases "q = 0") case True then show "P q" by (simp add: zero) next case False then have "degree (pCons a q) = Suc (degree q)" by (rule degree_pCons_eq) with \p = pCons a q\ have "degree q < degree p" by simp then show "P q" by (rule less.hyps) qed have "P (pCons a q)" proof (cases "a \ 0 \ q \ 0") case True with \P q\ show ?thesis by (auto intro: pCons) next case False with zero show ?thesis by simp qed with \p = pCons a q\ show ?case by simp qed lemma degree_eq_zeroE: fixes p :: "'a::zero poly" assumes "degree p = 0" obtains a where "p = pCons a 0" proof - obtain a q where p: "p = pCons a q" by (cases p) with assms have "q = 0" by (cases "q = 0") simp_all with p have "p = pCons a 0" by simp then show thesis .. qed subsection \Quickcheck generator for polynomials\ quickcheck_generator poly constructors: "0 :: _ poly", pCons subsection \List-style syntax for polynomials\ syntax "_poly" :: "args \ 'a poly" ("[:(_):]") translations "[:x, xs:]" \ "CONST pCons x [:xs:]" "[:x:]" \ "CONST pCons x 0" "[:x:]" \ "CONST pCons x (_constrain 0 t)" subsection \Representation of polynomials by lists of coefficients\ primrec Poly :: "'a::zero list \ 'a poly" where [code_post]: "Poly [] = 0" | [code_post]: "Poly (a # as) = pCons a (Poly as)" lemma Poly_replicate_0 [simp]: "Poly (replicate n 0) = 0" by (induct n) simp_all lemma Poly_eq_0: "Poly as = 0 \ (\n. as = replicate n 0)" by (induct as) (auto simp add: Cons_replicate_eq) lemma Poly_append_replicate_zero [simp]: "Poly (as @ replicate n 0) = Poly as" by (induct as) simp_all lemma Poly_snoc_zero [simp]: "Poly (as @ [0]) = Poly as" using Poly_append_replicate_zero [of as 1] by simp lemma Poly_cCons_eq_pCons_Poly [simp]: "Poly (a ## p) = pCons a (Poly p)" by (simp add: cCons_def) lemma Poly_on_rev_starting_with_0 [simp]: "hd as = 0 \ Poly (rev (tl as)) = Poly (rev as)" by (cases as) simp_all lemma degree_Poly: "degree (Poly xs) \ length xs" by (induct xs) simp_all lemma coeff_Poly_eq [simp]: "coeff (Poly xs) = nth_default 0 xs" by (induct xs) (simp_all add: fun_eq_iff coeff_pCons split: nat.splits) definition coeffs :: "'a poly \ 'a::zero list" where "coeffs p = (if p = 0 then [] else map (\i. coeff p i) [0 ..< Suc (degree p)])" lemma coeffs_eq_Nil [simp]: "coeffs p = [] \ p = 0" by (simp add: coeffs_def) lemma not_0_coeffs_not_Nil: "p \ 0 \ coeffs p \ []" by simp lemma coeffs_0_eq_Nil [simp]: "coeffs 0 = []" by simp lemma coeffs_pCons_eq_cCons [simp]: "coeffs (pCons a p) = a ## coeffs p" proof - have *: "\m\set ms. m > 0 \ map (case_nat x f) ms = map f (map (\n. n - 1) ms)" for ms :: "nat list" and f :: "nat \ 'a" and x :: "'a" by (induct ms) (auto split: nat.split) show ?thesis by (simp add: * coeffs_def upt_conv_Cons coeff_pCons map_decr_upt del: upt_Suc) qed lemma length_coeffs: "p \ 0 \ length (coeffs p) = degree p + 1" by (simp add: coeffs_def) lemma coeffs_nth: "p \ 0 \ n \ degree p \ coeffs p ! n = coeff p n" by (auto simp: coeffs_def simp del: upt_Suc) lemma coeff_in_coeffs: "p \ 0 \ n \ degree p \ coeff p n \ set (coeffs p)" using coeffs_nth [of p n, symmetric] by (simp add: length_coeffs) lemma not_0_cCons_eq [simp]: "p \ 0 \ a ## coeffs p = a # coeffs p" by (simp add: cCons_def) lemma Poly_coeffs [simp, code abstype]: "Poly (coeffs p) = p" by (induct p) auto lemma coeffs_Poly [simp]: "coeffs (Poly as) = strip_while (HOL.eq 0) as" proof (induct as) case Nil then show ?case by simp next case (Cons a as) from replicate_length_same [of as 0] have "(\n. as \ replicate n 0) \ (\a\set as. a \ 0)" by (auto dest: sym [of _ as]) with Cons show ?case by auto qed lemma no_trailing_coeffs [simp]: "no_trailing (HOL.eq 0) (coeffs p)" by (induct p) auto lemma strip_while_coeffs [simp]: "strip_while (HOL.eq 0) (coeffs p) = coeffs p" by simp lemma coeffs_eq_iff: "p = q \ coeffs p = coeffs q" (is "?P \ ?Q") proof assume ?P then show ?Q by simp next assume ?Q then have "Poly (coeffs p) = Poly (coeffs q)" by simp then show ?P by simp qed lemma nth_default_coeffs_eq: "nth_default 0 (coeffs p) = coeff p" by (simp add: fun_eq_iff coeff_Poly_eq [symmetric]) lemma [code]: "coeff p = nth_default 0 (coeffs p)" by (simp add: nth_default_coeffs_eq) lemma coeffs_eqI: assumes coeff: "\n. coeff p n = nth_default 0 xs n" assumes zero: "no_trailing (HOL.eq 0) xs" shows "coeffs p = xs" proof - from coeff have "p = Poly xs" by (simp add: poly_eq_iff) with zero show ?thesis by simp qed lemma degree_eq_length_coeffs [code]: "degree p = length (coeffs p) - 1" by (simp add: coeffs_def) lemma length_coeffs_degree: "p \ 0 \ length (coeffs p) = Suc (degree p)" by (induct p) (auto simp: cCons_def) lemma [code abstract]: "coeffs 0 = []" by (fact coeffs_0_eq_Nil) lemma [code abstract]: "coeffs (pCons a p) = a ## coeffs p" by (fact coeffs_pCons_eq_cCons) lemma set_coeffs_subset_singleton_0_iff [simp]: "set (coeffs p) \ {0} \ p = 0" by (auto simp add: coeffs_def intro: classical) lemma set_coeffs_not_only_0 [simp]: "set (coeffs p) \ {0}" by (auto simp add: set_eq_subset) lemma forall_coeffs_conv: "(\n. P (coeff p n)) \ (\c \ set (coeffs p). P c)" if "P 0" using that by (auto simp add: coeffs_def) (metis atLeastLessThan_iff coeff_eq_0 not_less_iff_gr_or_eq zero_le) instantiation poly :: ("{zero, equal}") equal begin definition [code]: "HOL.equal (p::'a poly) q \ HOL.equal (coeffs p) (coeffs q)" instance by standard (simp add: equal equal_poly_def coeffs_eq_iff) end lemma [code nbe]: "HOL.equal (p :: _ poly) p \ True" by (fact equal_refl) definition is_zero :: "'a::zero poly \ bool" where [code]: "is_zero p \ List.null (coeffs p)" lemma is_zero_null [code_abbrev]: "is_zero p \ p = 0" by (simp add: is_zero_def null_def) subsubsection \Reconstructing the polynomial from the list\ \ \contributed by Sebastiaan J.C. Joosten and René Thiemann\ definition poly_of_list :: "'a::comm_monoid_add list \ 'a poly" where [simp]: "poly_of_list = Poly" lemma poly_of_list_impl [code abstract]: "coeffs (poly_of_list as) = strip_while (HOL.eq 0) as" by simp subsection \Fold combinator for polynomials\ definition fold_coeffs :: "('a::zero \ 'b \ 'b) \ 'a poly \ 'b \ 'b" where "fold_coeffs f p = foldr f (coeffs p)" lemma fold_coeffs_0_eq [simp]: "fold_coeffs f 0 = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_eq [simp]: "f 0 = id \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def cCons_def fun_eq_iff) lemma fold_coeffs_pCons_0_0_eq [simp]: "fold_coeffs f (pCons 0 0) = id" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_coeff_not_0_eq [simp]: "a \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) lemma fold_coeffs_pCons_not_0_0_eq [simp]: "p \ 0 \ fold_coeffs f (pCons a p) = f a \ fold_coeffs f p" by (simp add: fold_coeffs_def) subsection \Canonical morphism on polynomials -- evaluation\ definition poly :: \'a::comm_semiring_0 poly \ 'a \ 'a\ where \poly p a = horner_sum id a (coeffs p)\ lemma poly_eq_fold_coeffs: \poly p = fold_coeffs (\a f x. a + x * f x) p (\x. 0)\ by (induction p) (auto simp add: fun_eq_iff poly_def) lemma poly_0 [simp]: "poly 0 x = 0" by (simp add: poly_def) lemma poly_pCons [simp]: "poly (pCons a p) x = a + x * poly p x" by (cases "p = 0 \ a = 0") (auto simp add: poly_def) lemma poly_altdef: "poly p x = (\i\degree p. coeff p i * x ^ i)" for x :: "'a::{comm_semiring_0,semiring_1}" proof (induction p rule: pCons_induct) case 0 then show ?case by simp next case (pCons a p) show ?case proof (cases "p = 0") case True then show ?thesis by simp next case False let ?p' = "pCons a p" note poly_pCons[of a p x] also note pCons.IH also have "a + x * (\i\degree p. coeff p i * x ^ i) = coeff ?p' 0 * x^0 + (\i\degree p. coeff ?p' (Suc i) * x^Suc i)" by (simp add: field_simps sum_distrib_left coeff_pCons) also note sum.atMost_Suc_shift[symmetric] also note degree_pCons_eq[OF \p \ 0\, of a, symmetric] finally show ?thesis . qed qed lemma poly_0_coeff_0: "poly p 0 = coeff p 0" by (cases p) (auto simp: poly_altdef) subsection \Monomials\ lift_definition monom :: "'a \ nat \ 'a::zero poly" is "\a m n. if m = n then a else 0" by (simp add: MOST_iff_cofinite) lemma coeff_monom [simp]: "coeff (monom a m) n = (if m = n then a else 0)" by transfer rule lemma monom_0: "monom a 0 = pCons a 0" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_Suc: "monom a (Suc n) = pCons 0 (monom a n)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma monom_eq_0 [simp]: "monom 0 n = 0" by (rule poly_eqI) simp lemma monom_eq_0_iff [simp]: "monom a n = 0 \ a = 0" by (simp add: poly_eq_iff) lemma monom_eq_iff [simp]: "monom a n = monom b n \ a = b" by (simp add: poly_eq_iff) lemma degree_monom_le: "degree (monom a n) \ n" by (rule degree_le, simp) lemma degree_monom_eq: "a \ 0 \ degree (monom a n) = n" by (metis coeff_monom leading_coeff_0_iff) lemma coeffs_monom [code abstract]: "coeffs (monom a n) = (if a = 0 then [] else replicate n 0 @ [a])" by (induct n) (simp_all add: monom_0 monom_Suc) lemma fold_coeffs_monom [simp]: "a \ 0 \ fold_coeffs f (monom a n) = f 0 ^^ n \ f a" by (simp add: fold_coeffs_def coeffs_monom fun_eq_iff) lemma poly_monom: "poly (monom a n) x = a * x ^ n" for a x :: "'a::comm_semiring_1" by (cases "a = 0", simp_all) (induct n, simp_all add: mult.left_commute poly_eq_fold_coeffs) lemma monom_eq_iff': "monom c n = monom d m \ c = d \ (c = 0 \ n = m)" by (auto simp: poly_eq_iff) lemma monom_eq_const_iff: "monom c n = [:d:] \ c = d \ (c = 0 \ n = 0)" using monom_eq_iff'[of c n d 0] by (simp add: monom_0) subsection \Leading coefficient\ abbreviation lead_coeff:: "'a::zero poly \ 'a" where "lead_coeff p \ coeff p (degree p)" lemma lead_coeff_pCons[simp]: "p \ 0 \ lead_coeff (pCons a p) = lead_coeff p" "p = 0 \ lead_coeff (pCons a p) = a" by auto lemma lead_coeff_monom [simp]: "lead_coeff (monom c n) = c" by (cases "c = 0") (simp_all add: degree_monom_eq) lemma last_coeffs_eq_coeff_degree: "last (coeffs p) = lead_coeff p" if "p \ 0" using that by (simp add: coeffs_def) subsection \Addition and subtraction\ instantiation poly :: (comm_monoid_add) comm_monoid_add begin lift_definition plus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n + coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n + coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_add [simp]: "coeff (p + q) n = coeff p n + coeff q n" by (simp add: plus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "(p + q) + r = p + (q + r)" by (simp add: poly_eq_iff add.assoc) show "p + q = q + p" by (simp add: poly_eq_iff add.commute) show "0 + p = p" by (simp add: poly_eq_iff) qed end instantiation poly :: (cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly :: "'a poly \ 'a poly \ 'a poly" is "\p q n. coeff p n - coeff q n" proof - fix q p :: "'a poly" show "\\<^sub>\n. coeff p n - coeff q n = 0" using MOST_coeff_eq_0[of p] MOST_coeff_eq_0[of q] by eventually_elim simp qed lemma coeff_diff [simp]: "coeff (p - q) n = coeff p n - coeff q n" by (simp add: minus_poly.rep_eq) instance proof fix p q r :: "'a poly" show "p + q - p = q" by (simp add: poly_eq_iff) show "p - q - r = p - (q + r)" by (simp add: poly_eq_iff diff_diff_eq) qed end instantiation poly :: (ab_group_add) ab_group_add begin lift_definition uminus_poly :: "'a poly \ 'a poly" is "\p n. - coeff p n" proof - fix p :: "'a poly" show "\\<^sub>\n. - coeff p n = 0" using MOST_coeff_eq_0 by simp qed lemma coeff_minus [simp]: "coeff (- p) n = - coeff p n" by (simp add: uminus_poly.rep_eq) instance proof fix p q :: "'a poly" show "- p + p = 0" by (simp add: poly_eq_iff) show "p - q = p + - q" by (simp add: poly_eq_iff) qed end lemma add_pCons [simp]: "pCons a p + pCons b q = pCons (a + b) (p + q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma minus_pCons [simp]: "- pCons a p = pCons (- a) (- p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma diff_pCons [simp]: "pCons a p - pCons b q = pCons (a - b) (p - q)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma degree_add_le_max: "degree (p + q) \ max (degree p) (degree q)" by (rule degree_le) (auto simp add: coeff_eq_0) lemma degree_add_le: "degree p \ n \ degree q \ n \ degree (p + q) \ n" by (auto intro: order_trans degree_add_le_max) lemma degree_add_less: "degree p < n \ degree q < n \ degree (p + q) < n" by (auto intro: le_less_trans degree_add_le_max) lemma degree_add_eq_right: assumes "degree p < degree q" shows "degree (p + q) = degree q" proof (cases "q = 0") case False show ?thesis proof (rule order_antisym) show "degree (p + q) \ degree q" by (simp add: assms degree_add_le order.strict_implies_order) show "degree q \ degree (p + q)" by (simp add: False assms coeff_eq_0 le_degree) qed qed (use assms in auto) lemma degree_add_eq_left: "degree q < degree p \ degree (p + q) = degree p" using degree_add_eq_right [of q p] by (simp add: add.commute) lemma degree_minus [simp]: "degree (- p) = degree p" by (simp add: degree_def) lemma lead_coeff_add_le: "degree p < degree q \ lead_coeff (p + q) = lead_coeff q" by (metis coeff_add coeff_eq_0 monoid_add_class.add.left_neutral degree_add_eq_right) lemma lead_coeff_minus: "lead_coeff (- p) = - lead_coeff p" by (metis coeff_minus degree_minus) lemma degree_diff_le_max: "degree (p - q) \ max (degree p) (degree q)" for p q :: "'a::ab_group_add poly" using degree_add_le [where p=p and q="-q"] by simp lemma degree_diff_le: "degree p \ n \ degree q \ n \ degree (p - q) \ n" for p q :: "'a::ab_group_add poly" using degree_add_le [of p n "- q"] by simp lemma degree_diff_less: "degree p < n \ degree q < n \ degree (p - q) < n" for p q :: "'a::ab_group_add poly" using degree_add_less [of p n "- q"] by simp lemma add_monom: "monom a n + monom b n = monom (a + b) n" by (rule poly_eqI) simp lemma diff_monom: "monom a n - monom b n = monom (a - b) n" by (rule poly_eqI) simp lemma minus_monom: "- monom a n = monom (- a) n" by (rule poly_eqI) simp lemma coeff_sum: "coeff (\x\A. p x) i = (\x\A. coeff (p x) i)" by (induct A rule: infinite_finite_induct) simp_all lemma monom_sum: "monom (\x\A. a x) n = (\x\A. monom (a x) n)" by (rule poly_eqI) (simp add: coeff_sum) fun plus_coeffs :: "'a::comm_monoid_add list \ 'a list \ 'a list" where "plus_coeffs xs [] = xs" | "plus_coeffs [] ys = ys" | "plus_coeffs (x # xs) (y # ys) = (x + y) ## plus_coeffs xs ys" lemma coeffs_plus_eq_plus_coeffs [code abstract]: "coeffs (p + q) = plus_coeffs (coeffs p) (coeffs q)" proof - have *: "nth_default 0 (plus_coeffs xs ys) n = nth_default 0 xs n + nth_default 0 ys n" for xs ys :: "'a list" and n proof (induct xs ys arbitrary: n rule: plus_coeffs.induct) case (3 x xs y ys n) then show ?case by (cases n) (auto simp add: cCons_def) qed simp_all have **: "no_trailing (HOL.eq 0) (plus_coeffs xs ys)" if "no_trailing (HOL.eq 0) xs" and "no_trailing (HOL.eq 0) ys" for xs ys :: "'a list" using that by (induct xs ys rule: plus_coeffs.induct) (simp_all add: cCons_def) show ?thesis by (rule coeffs_eqI) (auto simp add: * nth_default_coeffs_eq intro: **) qed lemma coeffs_uminus [code abstract]: "coeffs (- p) = map uminus (coeffs p)" proof - have eq_0: "HOL.eq 0 \ uminus = HOL.eq (0::'a)" by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (simp_all add: nth_default_map_eq nth_default_coeffs_eq no_trailing_map eq_0) qed lemma [code]: "p - q = p + - q" for p q :: "'a::ab_group_add poly" by (fact diff_conv_add_uminus) lemma poly_add [simp]: "poly (p + q) x = poly p x + poly q x" proof (induction p arbitrary: q) case (pCons a p) then show ?case by (cases q) (simp add: algebra_simps) qed auto lemma poly_minus [simp]: "poly (- p) x = - poly p x" for x :: "'a::comm_ring" by (induct p) simp_all lemma poly_diff [simp]: "poly (p - q) x = poly p x - poly q x" for x :: "'a::comm_ring" using poly_add [of p "- q" x] by simp lemma poly_sum: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_sum_le: "finite S \ (\p. p \ S \ degree (f p) \ n) \ degree (sum f S) \ n" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert p S) then have "degree (sum f S) \ n" "degree (f p) \ n" by auto then show ?case unfolding sum.insert[OF insert(1-2)] by (metis degree_add_le) qed lemma poly_as_sum_of_monoms': assumes "degree p \ n" shows "(\i\n. monom (coeff p i) i) = p" proof - have eq: "\i. {..n} \ {i} = (if i \ n then {i} else {})" by auto from assms show ?thesis by (simp add: poly_eq_iff coeff_sum coeff_eq_0 sum.If_cases eq if_distrib[where f="\x. x * a" for a]) qed lemma poly_as_sum_of_monoms: "(\i\degree p. monom (coeff p i) i) = p" by (intro poly_as_sum_of_monoms' order_refl) lemma Poly_snoc: "Poly (xs @ [x]) = Poly xs + monom x (length xs)" by (induct xs) (simp_all add: monom_0 monom_Suc) subsection \Multiplication by a constant, polynomial multiplication and the unit polynomial\ lift_definition smult :: "'a::comm_semiring_0 \ 'a poly \ 'a poly" is "\a p n. a * coeff p n" proof - fix a :: 'a and p :: "'a poly" show "\\<^sub>\ i. a * coeff p i = 0" using MOST_coeff_eq_0[of p] by eventually_elim simp qed lemma coeff_smult [simp]: "coeff (smult a p) n = a * coeff p n" by (simp add: smult.rep_eq) lemma degree_smult_le: "degree (smult a p) \ degree p" by (rule degree_le) (simp add: coeff_eq_0) lemma smult_smult [simp]: "smult a (smult b p) = smult (a * b) p" by (rule poly_eqI) (simp add: mult.assoc) lemma smult_0_right [simp]: "smult a 0 = 0" by (rule poly_eqI) simp lemma smult_0_left [simp]: "smult 0 p = 0" by (rule poly_eqI) simp lemma smult_1_left [simp]: "smult (1::'a::comm_semiring_1) p = p" by (rule poly_eqI) simp lemma smult_add_right: "smult a (p + q) = smult a p + smult a q" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_add_left: "smult (a + b) p = smult a p + smult b p" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_minus_right [simp]: "smult a (- p) = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_minus_left [simp]: "smult (- a) p = - smult a p" for a :: "'a::comm_ring" by (rule poly_eqI) simp lemma smult_diff_right: "smult a (p - q) = smult a p - smult a q" for a :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemma smult_diff_left: "smult (a - b) p = smult a p - smult b p" for a b :: "'a::comm_ring" by (rule poly_eqI) (simp add: algebra_simps) lemmas smult_distribs = smult_add_left smult_add_right smult_diff_left smult_diff_right lemma smult_pCons [simp]: "smult a (pCons b p) = pCons (a * b) (smult a p)" by (rule poly_eqI) (simp add: coeff_pCons split: nat.split) lemma smult_monom: "smult a (monom b n) = monom (a * b) n" by (induct n) (simp_all add: monom_0 monom_Suc) lemma smult_Poly: "smult c (Poly xs) = Poly (map ((*) c) xs)" by (auto simp: poly_eq_iff nth_default_def) lemma degree_smult_eq [simp]: "degree (smult a p) = (if a = 0 then 0 else degree p)" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (cases "a = 0") (simp_all add: degree_def) lemma smult_eq_0_iff [simp]: "smult a p = 0 \ a = 0 \ p = 0" for a :: "'a::{comm_semiring_0,semiring_no_zero_divisors}" by (simp add: poly_eq_iff) lemma coeffs_smult [code abstract]: "coeffs (smult a p) = (if a = 0 then [] else map (Groups.times a) (coeffs p))" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have eq_0: "HOL.eq 0 \ times a = HOL.eq (0::'a)" if "a \ 0" using that by (simp add: fun_eq_iff) show ?thesis by (rule coeffs_eqI) (auto simp add: no_trailing_map nth_default_map_eq nth_default_coeffs_eq eq_0) qed lemma smult_eq_iff: fixes b :: "'a :: field" assumes "b \ 0" shows "smult a p = smult b q \ smult (a / b) p = q" (is "?lhs \ ?rhs") proof assume ?lhs also from assms have "smult (inverse b) \ = q" by simp finally show ?rhs by (simp add: field_simps) next assume ?rhs with assms show ?lhs by auto qed instantiation poly :: (comm_semiring_0) comm_semiring_0 begin definition "p * q = fold_coeffs (\a p. smult a q + pCons 0 p) p 0" lemma mult_poly_0_left: "(0::'a poly) * q = 0" by (simp add: times_poly_def) lemma mult_pCons_left [simp]: "pCons a p * q = smult a q + pCons 0 (p * q)" by (cases "p = 0 \ a = 0") (auto simp add: times_poly_def) lemma mult_poly_0_right: "p * (0::'a poly) = 0" by (induct p) (simp_all add: mult_poly_0_left) lemma mult_pCons_right [simp]: "p * pCons a q = smult a p + pCons 0 (p * q)" by (induct p) (simp_all add: mult_poly_0_left algebra_simps) lemmas mult_poly_0 = mult_poly_0_left mult_poly_0_right lemma mult_smult_left [simp]: "smult a p * q = smult a (p * q)" by (induct p) (simp_all add: mult_poly_0 smult_add_right) lemma mult_smult_right [simp]: "p * smult a q = smult a (p * q)" by (induct q) (simp_all add: mult_poly_0 smult_add_right) lemma mult_poly_add_left: "(p + q) * r = p * r + q * r" for p q r :: "'a poly" by (induct r) (simp_all add: mult_poly_0 smult_distribs algebra_simps) instance proof fix p q r :: "'a poly" show 0: "0 * p = 0" by (rule mult_poly_0_left) show "p * 0 = 0" by (rule mult_poly_0_right) show "(p + q) * r = p * r + q * r" by (rule mult_poly_add_left) show "(p * q) * r = p * (q * r)" by (induct p) (simp_all add: mult_poly_0 mult_poly_add_left) show "p * q = q * p" by (induct p) (simp_all add: mult_poly_0) qed end lemma coeff_mult_degree_sum: "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (induct p) (simp_all add: coeff_eq_0) instance poly :: ("{comm_semiring_0,semiring_no_zero_divisors}") semiring_no_zero_divisors proof fix p q :: "'a poly" assume "p \ 0" and "q \ 0" have "coeff (p * q) (degree p + degree q) = coeff p (degree p) * coeff q (degree q)" by (rule coeff_mult_degree_sum) also from \p \ 0\ \q \ 0\ have "coeff p (degree p) * coeff q (degree q) \ 0" by simp finally have "\n. coeff (p * q) n \ 0" .. then show "p * q \ 0" by (simp add: poly_eq_iff) qed instance poly :: (comm_semiring_0_cancel) comm_semiring_0_cancel .. lemma coeff_mult: "coeff (p * q) n = (\i\n. coeff p i * coeff q (n-i))" proof (induct p arbitrary: n) case 0 show ?case by simp next case (pCons a p n) then show ?case by (cases n) (simp_all add: sum.atMost_Suc_shift del: sum.atMost_Suc) qed lemma degree_mult_le: "degree (p * q) \ degree p + degree q" proof (rule degree_le) show "\i>degree p + degree q. coeff (p * q) i = 0" by (induct p) (simp_all add: coeff_eq_0 coeff_pCons split: nat.split) qed lemma mult_monom: "monom a m * monom b n = monom (a * b) (m + n)" by (induct m) (simp add: monom_0 smult_monom, simp add: monom_Suc) instantiation poly :: (comm_semiring_1) comm_semiring_1 begin lift_definition one_poly :: "'a poly" is "\n. of_bool (n = 0)" by (rule MOST_SucD) simp lemma coeff_1 [simp]: "coeff 1 n = of_bool (n = 0)" by (simp add: one_poly.rep_eq) lemma one_pCons: "1 = [:1:]" by (simp add: poly_eq_iff coeff_pCons split: nat.splits) lemma pCons_one: "[:1:] = 1" by (simp add: one_pCons) instance by standard (simp_all add: one_pCons) end lemma poly_1 [simp]: "poly 1 x = 1" by (simp add: one_pCons) lemma one_poly_eq_simps [simp]: "1 = [:1:] \ True" "[:1:] = 1 \ True" by (simp_all add: one_pCons) lemma degree_1 [simp]: "degree 1 = 0" by (simp add: one_pCons) lemma coeffs_1_eq [simp, code abstract]: "coeffs 1 = [1]" by (simp add: one_pCons) lemma smult_one [simp]: "smult c 1 = [:c:]" by (simp add: one_pCons) lemma monom_eq_1 [simp]: "monom 1 0 = 1" by (simp add: monom_0 one_pCons) lemma monom_eq_1_iff: "monom c n = 1 \ c = 1 \ n = 0" using monom_eq_const_iff [of c n 1] by auto lemma monom_altdef: "monom c n = smult c ([:0, 1:] ^ n)" by (induct n) (simp_all add: monom_0 monom_Suc) instance poly :: ("{comm_semiring_1,semiring_1_no_zero_divisors}") semiring_1_no_zero_divisors .. instance poly :: (comm_ring) comm_ring .. instance poly :: (comm_ring_1) comm_ring_1 .. instance poly :: (comm_ring_1) comm_semiring_1_cancel .. lemma degree_power_le: "degree (p ^ n) \ degree p * n" by (induct n) (auto intro: order_trans degree_mult_le) lemma coeff_0_power: "coeff (p ^ n) 0 = coeff p 0 ^ n" by (induct n) (simp_all add: coeff_mult) lemma poly_smult [simp]: "poly (smult a p) x = a * poly p x" by (induct p) (simp_all add: algebra_simps) lemma poly_mult [simp]: "poly (p * q) x = poly p x * poly q x" by (induct p) (simp_all add: algebra_simps) lemma poly_power [simp]: "poly (p ^ n) x = poly p x ^ n" for p :: "'a::comm_semiring_1 poly" by (induct n) simp_all lemma poly_prod: "poly (\k\A. p k) x = (\k\A. poly (p k) x)" by (induct A rule: infinite_finite_induct) simp_all lemma degree_prod_sum_le: "finite S \ degree (prod f S) \ sum (degree \ f) S" proof (induct S rule: finite_induct) case empty then show ?case by simp next case (insert a S) show ?case unfolding prod.insert[OF insert(1-2)] sum.insert[OF insert(1-2)] by (rule le_trans[OF degree_mult_le]) (use insert in auto) qed lemma coeff_0_prod_list: "coeff (prod_list xs) 0 = prod_list (map (\p. coeff p 0) xs)" by (induct xs) (simp_all add: coeff_mult) lemma coeff_monom_mult: "coeff (monom c n * p) k = (if k < n then 0 else c * coeff p (k - n))" proof - have "coeff (monom c n * p) k = (\i\k. (if n = i then c else 0) * coeff p (k - i))" by (simp add: coeff_mult) also have "\ = (\i\k. (if n = i then c * coeff p (k - i) else 0))" by (intro sum.cong) simp_all also have "\ = (if k < n then 0 else c * coeff p (k - n))" by simp finally show ?thesis . qed lemma monom_1_dvd_iff': "monom 1 n dvd p \ (\kkkk. coeff p (k + n))" have "\\<^sub>\k. coeff p (k + n) = 0" by (subst cofinite_eq_sequentially, subst eventually_sequentially_seg, subst cofinite_eq_sequentially [symmetric]) transfer then have coeff_r [simp]: "coeff r k = coeff p (k + n)" for k unfolding r_def by (subst poly.Abs_poly_inverse) simp_all have "p = monom 1 n * r" by (rule poly_eqI, subst coeff_monom_mult) (simp_all add: zero) then show "monom 1 n dvd p" by simp qed subsection \Mapping polynomials\ definition map_poly :: "('a :: zero \ 'b :: zero) \ 'a poly \ 'b poly" where "map_poly f p = Poly (map f (coeffs p))" lemma map_poly_0 [simp]: "map_poly f 0 = 0" by (simp add: map_poly_def) lemma map_poly_1: "map_poly f 1 = [:f 1:]" by (simp add: map_poly_def) lemma map_poly_1' [simp]: "f 1 = 1 \ map_poly f 1 = 1" by (simp add: map_poly_def one_pCons) lemma coeff_map_poly: assumes "f 0 = 0" shows "coeff (map_poly f p) n = f (coeff p n)" by (auto simp: assms map_poly_def nth_default_def coeffs_def not_less Suc_le_eq coeff_eq_0 simp del: upt_Suc) lemma coeffs_map_poly [code abstract]: "coeffs (map_poly f p) = strip_while ((=) 0) (map f (coeffs p))" by (simp add: map_poly_def) lemma coeffs_map_poly': assumes "\x. x \ 0 \ f x \ 0" shows "coeffs (map_poly f p) = map f (coeffs p)" using assms by (auto simp add: coeffs_map_poly strip_while_idem_iff last_coeffs_eq_coeff_degree no_trailing_unfold last_map) lemma set_coeffs_map_poly: "(\x. f x = 0 \ x = 0) \ set (coeffs (map_poly f p)) = f ` set (coeffs p)" by (simp add: coeffs_map_poly') lemma degree_map_poly: assumes "\x. x \ 0 \ f x \ 0" shows "degree (map_poly f p) = degree p" by (simp add: degree_eq_length_coeffs coeffs_map_poly' assms) lemma map_poly_eq_0_iff: assumes "f 0 = 0" "\x. x \ set (coeffs p) \ x \ 0 \ f x \ 0" shows "map_poly f p = 0 \ p = 0" proof - have "(coeff (map_poly f p) n = 0) = (coeff p n = 0)" for n proof - have "coeff (map_poly f p) n = f (coeff p n)" by (simp add: coeff_map_poly assms) also have "\ = 0 \ coeff p n = 0" proof (cases "n < length (coeffs p)") case True then have "coeff p n \ set (coeffs p)" by (auto simp: coeffs_def simp del: upt_Suc) with assms show "f (coeff p n) = 0 \ coeff p n = 0" by auto next case False then show ?thesis by (auto simp: assms length_coeffs nth_default_coeffs_eq [symmetric] nth_default_def) qed finally show ?thesis . qed then show ?thesis by (auto simp: poly_eq_iff) qed lemma map_poly_smult: assumes "f 0 = 0""\c x. f (c * x) = f c * f x" shows "map_poly f (smult c p) = smult (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly) lemma map_poly_pCons: assumes "f 0 = 0" shows "map_poly f (pCons c p) = pCons (f c) (map_poly f p)" by (intro poly_eqI) (simp_all add: assms coeff_map_poly coeff_pCons split: nat.splits) lemma map_poly_map_poly: assumes "f 0 = 0" "g 0 = 0" shows "map_poly f (map_poly g p) = map_poly (f \ g) p" by (intro poly_eqI) (simp add: coeff_map_poly assms) lemma map_poly_id [simp]: "map_poly id p = p" by (simp add: map_poly_def) lemma map_poly_id' [simp]: "map_poly (\x. x) p = p" by (simp add: map_poly_def) lemma map_poly_cong: assumes "(\x. x \ set (coeffs p) \ f x = g x)" shows "map_poly f p = map_poly g p" proof - from assms have "map f (coeffs p) = map g (coeffs p)" by (intro map_cong) simp_all then show ?thesis by (simp only: coeffs_eq_iff coeffs_map_poly) qed lemma map_poly_monom: "f 0 = 0 \ map_poly f (monom c n) = monom (f c) n" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma map_poly_idI: assumes "\x. x \ set (coeffs p) \ f x = x" shows "map_poly f p = p" using map_poly_cong[OF assms, of _ id] by simp lemma map_poly_idI': assumes "\x. x \ set (coeffs p) \ f x = x" shows "p = map_poly f p" using map_poly_cong[OF assms, of _ id] by simp lemma smult_conv_map_poly: "smult c p = map_poly (\x. c * x) p" by (intro poly_eqI) (simp_all add: coeff_map_poly) subsection \Conversions\ lemma of_nat_poly: "of_nat n = [:of_nat n:]" by (induct n) (simp_all add: one_pCons) lemma of_nat_monom: "of_nat n = monom (of_nat n) 0" by (simp add: of_nat_poly monom_0) lemma degree_of_nat [simp]: "degree (of_nat n) = 0" by (simp add: of_nat_poly) lemma lead_coeff_of_nat [simp]: "lead_coeff (of_nat n) = of_nat n" by (simp add: of_nat_poly) lemma of_int_poly: "of_int k = [:of_int k:]" by (simp only: of_int_of_nat of_nat_poly) simp lemma of_int_monom: "of_int k = monom (of_int k) 0" by (simp add: of_int_poly monom_0) lemma degree_of_int [simp]: "degree (of_int k) = 0" by (simp add: of_int_poly) lemma lead_coeff_of_int [simp]: "lead_coeff (of_int k) = of_int k" by (simp add: of_int_poly) lemma numeral_poly: "numeral n = [:numeral n:]" proof - have "numeral n = of_nat (numeral n)" by simp also have "\ = [:of_nat (numeral n):]" by (simp add: of_nat_poly) finally show ?thesis by simp qed lemma numeral_monom: "numeral n = monom (numeral n) 0" by (simp add: numeral_poly monom_0) lemma degree_numeral [simp]: "degree (numeral n) = 0" by (simp add: numeral_poly) lemma lead_coeff_numeral [simp]: "lead_coeff (numeral n) = numeral n" by (simp add: numeral_poly) subsection \Lemmas about divisibility\ lemma dvd_smult: assumes "p dvd q" shows "p dvd smult a q" proof - from assms obtain k where "q = p * k" .. then have "smult a q = p * smult a k" by simp then show "p dvd smult a q" .. qed lemma dvd_smult_cancel: "p dvd smult a q \ a \ 0 \ p dvd q" for a :: "'a::field" by (drule dvd_smult [where a="inverse a"]) simp lemma dvd_smult_iff: "a \ 0 \ p dvd smult a q \ p dvd q" for a :: "'a::field" by (safe elim!: dvd_smult dvd_smult_cancel) lemma smult_dvd_cancel: assumes "smult a p dvd q" shows "p dvd q" proof - from assms obtain k where "q = smult a p * k" .. then have "q = p * smult a k" by simp then show "p dvd q" .. qed lemma smult_dvd: "p dvd q \ a \ 0 \ smult a p dvd q" for a :: "'a::field" by (rule smult_dvd_cancel [where a="inverse a"]) simp lemma smult_dvd_iff: "smult a p dvd q \ (if a = 0 then q = 0 else p dvd q)" for a :: "'a::field" by (auto elim: smult_dvd smult_dvd_cancel) lemma is_unit_smult_iff: "smult c p dvd 1 \ c dvd 1 \ p dvd 1" proof - have "smult c p = [:c:] * p" by simp also have "\ dvd 1 \ c dvd 1 \ p dvd 1" proof safe assume *: "[:c:] * p dvd 1" then show "p dvd 1" by (rule dvd_mult_right) from * obtain q where q: "1 = [:c:] * p * q" by (rule dvdE) have "c dvd c * (coeff p 0 * coeff q 0)" by simp also have "\ = coeff ([:c:] * p * q) 0" by (simp add: mult.assoc coeff_mult) also note q [symmetric] finally have "c dvd coeff 1 0" . then show "c dvd 1" by simp next assume "c dvd 1" "p dvd 1" from this(1) obtain d where "1 = c * d" by (rule dvdE) then have "1 = [:c:] * [:d:]" by (simp add: one_pCons ac_simps) then have "[:c:] dvd 1" by (rule dvdI) from mult_dvd_mono[OF this \p dvd 1\] show "[:c:] * p dvd 1" by simp qed finally show ?thesis . qed subsection \Polynomials form an integral domain\ instance poly :: (idom) idom .. instance poly :: ("{ring_char_0, comm_ring_1}") ring_char_0 by standard (auto simp add: of_nat_poly intro: injI) lemma degree_mult_eq: "p \ 0 \ q \ 0 \ degree (p * q) = degree p + degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (rule order_antisym [OF degree_mult_le le_degree]) (simp add: coeff_mult_degree_sum) lemma degree_prod_eq_sum_degree: fixes A :: "'a set" and f :: "'a \ 'b::idom poly" assumes f0: "\i\A. f i \ 0" shows "degree (\i\A. (f i)) = (\i\A. degree (f i))" using assms by (induction A rule: infinite_finite_induct) (auto simp: degree_mult_eq) lemma degree_mult_eq_0: "degree (p * q) = 0 \ p = 0 \ q = 0 \ (p \ 0 \ q \ 0 \ degree p = 0 \ degree q = 0)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (auto simp: degree_mult_eq) lemma degree_power_eq: "p \ 0 \ degree ((p :: 'a :: idom poly) ^ n) = n * degree p" by (induction n) (simp_all add: degree_mult_eq) lemma degree_mult_right_le: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "q \ 0" shows "degree p \ degree (p * q)" using assms by (cases "p = 0") (simp_all add: degree_mult_eq) lemma coeff_degree_mult: "coeff (p * q) (degree (p * q)) = coeff q (degree q) * coeff p (degree p)" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: degree_mult_eq coeff_mult_degree_sum mult_ac) lemma dvd_imp_degree_le: "p dvd q \ q \ 0 \ degree p \ degree q" for p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (erule dvdE, hypsubst, subst degree_mult_eq) auto lemma divides_degree: fixes p q :: "'a ::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd q" shows "degree p \ degree q \ q = 0" by (metis dvd_imp_degree_le assms) lemma const_poly_dvd_iff: fixes c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" shows "[:c:] dvd p \ (\n. c dvd coeff p n)" proof (cases "c = 0 \ p = 0") case True then show ?thesis by (auto intro!: poly_eqI) next case False show ?thesis proof assume "[:c:] dvd p" then show "\n. c dvd coeff p n" by (auto elim!: dvdE simp: coeffs_def) next assume *: "\n. c dvd coeff p n" define mydiv where "mydiv x y = (SOME z. x = y * z)" for x y :: 'a have mydiv: "x = y * mydiv x y" if "y dvd x" for x y using that unfolding mydiv_def dvd_def by (rule someI_ex) define q where "q = Poly (map (\a. mydiv a c) (coeffs p))" from False * have "p = q * [:c:]" by (intro poly_eqI) (auto simp: q_def nth_default_def not_less length_coeffs_degree coeffs_nth intro!: coeff_eq_0 mydiv) then show "[:c:] dvd p" by (simp only: dvd_triv_right) qed qed lemma const_poly_dvd_const_poly_iff [simp]: "[:a:] dvd [:b:] \ a dvd b" for a b :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (subst const_poly_dvd_iff) (auto simp: coeff_pCons split: nat.splits) lemma lead_coeff_mult: "lead_coeff (p * q) = lead_coeff p * lead_coeff q" for p q :: "'a::{comm_semiring_0, semiring_no_zero_divisors} poly" by (cases "p = 0 \ q = 0") (auto simp: coeff_mult_degree_sum degree_mult_eq) lemma lead_coeff_prod: "lead_coeff (prod f A) = (\x\A. lead_coeff (f x))" for f :: "'a \ 'b::{comm_semiring_1, semiring_no_zero_divisors} poly" by (induction A rule: infinite_finite_induct) (auto simp: lead_coeff_mult) lemma lead_coeff_smult: "lead_coeff (smult c p) = c * lead_coeff p" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof - have "smult c p = [:c:] * p" by simp also have "lead_coeff \ = c * lead_coeff p" by (subst lead_coeff_mult) simp_all finally show ?thesis . qed lemma lead_coeff_1 [simp]: "lead_coeff 1 = 1" by simp lemma lead_coeff_power: "lead_coeff (p ^ n) = lead_coeff p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: lead_coeff_mult) subsection \Polynomials form an ordered integral domain\ definition pos_poly :: "'a::linordered_semidom poly \ bool" where "pos_poly p \ 0 < coeff p (degree p)" lemma pos_poly_pCons: "pos_poly (pCons a p) \ pos_poly p \ (p = 0 \ 0 < a)" by (simp add: pos_poly_def) lemma not_pos_poly_0 [simp]: "\ pos_poly 0" by (simp add: pos_poly_def) lemma pos_poly_add: "pos_poly p \ pos_poly q \ pos_poly (p + q)" proof (induction p arbitrary: q) case (pCons a p) then show ?case by (cases q; force simp add: pos_poly_pCons add_pos_pos) qed auto lemma pos_poly_mult: "pos_poly p \ pos_poly q \ pos_poly (p * q)" by (simp add: pos_poly_def coeff_degree_mult) lemma pos_poly_total: "p = 0 \ pos_poly p \ pos_poly (- p)" for p :: "'a::linordered_idom poly" by (induct p) (auto simp: pos_poly_pCons) lemma pos_poly_coeffs [code]: "pos_poly p \ (let as = coeffs p in as \ [] \ last as > 0)" (is "?lhs \ ?rhs") proof assume ?rhs then show ?lhs by (auto simp add: pos_poly_def last_coeffs_eq_coeff_degree) next assume ?lhs then have *: "0 < coeff p (degree p)" by (simp add: pos_poly_def) then have "p \ 0" by auto with * show ?rhs by (simp add: last_coeffs_eq_coeff_degree) qed instantiation poly :: (linordered_idom) linordered_idom begin definition "x < y \ pos_poly (y - x)" definition "x \ y \ x = y \ pos_poly (y - x)" definition "\x::'a poly\ = (if x < 0 then - x else x)" definition "sgn (x::'a poly) = (if x = 0 then 0 else if 0 < x then 1 else - 1)" instance proof fix x y z :: "'a poly" show "x < y \ x \ y \ \ y \ x" unfolding less_eq_poly_def less_poly_def using pos_poly_add by force then show "x \ y \ y \ x \ x = y" using less_eq_poly_def less_poly_def by force show "x \ x" by (simp add: less_eq_poly_def) show "x \ y \ y \ z \ x \ z" using less_eq_poly_def pos_poly_add by fastforce show "x \ y \ z + x \ z + y" by (simp add: less_eq_poly_def) show "x \ y \ y \ x" unfolding less_eq_poly_def using pos_poly_total [of "x - y"] by auto show "x < y \ 0 < z \ z * x < z * y" by (simp add: less_poly_def right_diff_distrib [symmetric] pos_poly_mult) show "\x\ = (if x < 0 then - x else x)" by (rule abs_poly_def) show "sgn x = (if x = 0 then 0 else if 0 < x then 1 else - 1)" by (rule sgn_poly_def) qed end text \TODO: Simplification rules for comparisons\ subsection \Synthetic division and polynomial roots\ subsubsection \Synthetic division\ text \Synthetic division is simply division by the linear polynomial \<^term>\x - c\.\ definition synthetic_divmod :: "'a::comm_semiring_0 poly \ 'a \ 'a poly \ 'a" where "synthetic_divmod p c = fold_coeffs (\a (q, r). (pCons r q, a + c * r)) p (0, 0)" definition synthetic_div :: "'a::comm_semiring_0 poly \ 'a \ 'a poly" where "synthetic_div p c = fst (synthetic_divmod p c)" lemma synthetic_divmod_0 [simp]: "synthetic_divmod 0 c = (0, 0)" by (simp add: synthetic_divmod_def) lemma synthetic_divmod_pCons [simp]: "synthetic_divmod (pCons a p) c = (\(q, r). (pCons r q, a + c * r)) (synthetic_divmod p c)" by (cases "p = 0 \ a = 0") (auto simp add: synthetic_divmod_def) lemma synthetic_div_0 [simp]: "synthetic_div 0 c = 0" by (simp add: synthetic_div_def) lemma synthetic_div_unique_lemma: "smult c p = pCons a p \ p = 0" by (induct p arbitrary: a) simp_all lemma snd_synthetic_divmod: "snd (synthetic_divmod p c) = poly p c" by (induct p) (simp_all add: split_def) lemma synthetic_div_pCons [simp]: "synthetic_div (pCons a p) c = pCons (poly p c) (synthetic_div p c)" by (simp add: synthetic_div_def split_def snd_synthetic_divmod) lemma synthetic_div_eq_0_iff: "synthetic_div p c = 0 \ degree p = 0" proof (induct p) case 0 then show ?case by simp next case (pCons a p) then show ?case by (cases p) simp qed lemma degree_synthetic_div: "degree (synthetic_div p c) = degree p - 1" by (induct p) (simp_all add: synthetic_div_eq_0_iff) lemma synthetic_div_correct: "p + smult c (synthetic_div p c) = pCons (poly p c) (synthetic_div p c)" by (induct p) simp_all lemma synthetic_div_unique: "p + smult c q = pCons r q \ r = poly p c \ q = synthetic_div p c" proof (induction p arbitrary: q r) case 0 then show ?case using synthetic_div_unique_lemma by fastforce next case (pCons a p) then show ?case by (cases q; force) qed lemma synthetic_div_correct': "[:-c, 1:] * synthetic_div p c + [:poly p c:] = p" for c :: "'a::comm_ring_1" using synthetic_div_correct [of p c] by (simp add: algebra_simps) subsubsection \Polynomial roots\ lemma poly_eq_0_iff_dvd: "poly p c = 0 \ [:- c, 1:] dvd p" (is "?lhs \ ?rhs") for c :: "'a::comm_ring_1" proof assume ?lhs with synthetic_div_correct' [of c p] have "p = [:-c, 1:] * synthetic_div p c" by simp then show ?rhs .. next assume ?rhs then obtain k where "p = [:-c, 1:] * k" by (rule dvdE) then show ?lhs by simp qed lemma dvd_iff_poly_eq_0: "[:c, 1:] dvd p \ poly p (- c) = 0" for c :: "'a::comm_ring_1" by (simp add: poly_eq_0_iff_dvd) lemma poly_roots_finite: "p \ 0 \ finite {x. poly p x = 0}" for p :: "'a::{comm_ring_1,ring_no_zero_divisors} poly" proof (induct n \ "degree p" arbitrary: p) case 0 then obtain a where "a \ 0" and "p = [:a:]" by (cases p) (simp split: if_splits) then show "finite {x. poly p x = 0}" by simp next case (Suc n) show "finite {x. poly p x = 0}" proof (cases "\x. poly p x = 0") case False then show "finite {x. poly p x = 0}" by simp next case True then obtain a where "poly p a = 0" .. then have "[:-a, 1:] dvd p" by (simp only: poly_eq_0_iff_dvd) then obtain k where k: "p = [:-a, 1:] * k" .. with \p \ 0\ have "k \ 0" by auto with k have "degree p = Suc (degree k)" by (simp add: degree_mult_eq del: mult_pCons_left) with \Suc n = degree p\ have "n = degree k" by simp from this \k \ 0\ have "finite {x. poly k x = 0}" by (rule Suc.hyps) then have "finite (insert a {x. poly k x = 0})" by simp then show "finite {x. poly p x = 0}" by (simp add: k Collect_disj_eq del: mult_pCons_left) qed qed lemma poly_eq_poly_eq_iff: "poly p = poly q \ p = q" (is "?lhs \ ?rhs") for p q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly" proof assume ?rhs then show ?lhs by simp next assume ?lhs have "poly p = poly 0 \ p = 0" for p :: "'a poly" proof (cases "p = 0") case False then show ?thesis by (auto simp add: infinite_UNIV_char_0 dest: poly_roots_finite) qed auto from \?lhs\ and this [of "p - q"] show ?rhs by auto qed lemma poly_all_0_iff_0: "(\x. poly p x = 0) \ p = 0" for p :: "'a::{ring_char_0,comm_ring_1,ring_no_zero_divisors} poly" by (auto simp add: poly_eq_poly_eq_iff [symmetric]) subsubsection \Order of polynomial roots\ definition order :: "'a::idom \ 'a poly \ nat" where "order a p = (LEAST n. \ [:-a, 1:] ^ Suc n dvd p)" lemma coeff_linear_power: "coeff ([:a, 1:] ^ n) n = 1" for a :: "'a::comm_semiring_1" proof (induct n) case (Suc n) have "degree ([:a, 1:] ^ n) \ 1 * n" by (metis One_nat_def degree_pCons_eq_if degree_power_le one_neq_zero one_pCons) then have "coeff ([:a, 1:] ^ n) (Suc n) = 0" by (simp add: coeff_eq_0) then show ?case using Suc.hyps by fastforce qed auto lemma degree_linear_power: "degree ([:a, 1:] ^ n) = n" for a :: "'a::comm_semiring_1" proof (rule order_antisym) show "degree ([:a, 1:] ^ n) \ n" by (metis One_nat_def degree_pCons_eq_if degree_power_le mult.left_neutral one_neq_zero one_pCons) qed (simp add: coeff_linear_power le_degree) lemma order_1: "[:-a, 1:] ^ order a p dvd p" proof (cases "p = 0") case False show ?thesis proof (cases "order a p") case (Suc n) then show ?thesis by (metis lessI not_less_Least order_def) qed auto qed auto lemma order_2: assumes "p \ 0" shows "\ [:-a, 1:] ^ Suc (order a p) dvd p" proof - have False if "[:- a, 1:] ^ Suc (degree p) dvd p" using dvd_imp_degree_le [OF that] by (metis Suc_n_not_le_n assms degree_linear_power) then show ?thesis unfolding order_def by (metis (no_types, lifting) LeastI) qed lemma order: "p \ 0 \ [:-a, 1:] ^ order a p dvd p \ \ [:-a, 1:] ^ Suc (order a p) dvd p" by (rule conjI [OF order_1 order_2]) lemma order_degree: assumes p: "p \ 0" shows "order a p \ degree p" proof - have "order a p = degree ([:-a, 1:] ^ order a p)" by (simp only: degree_linear_power) also from order_1 p have "\ \ degree p" by (rule dvd_imp_degree_le) finally show ?thesis . qed lemma order_root: "poly p a = 0 \ p = 0 \ order a p \ 0" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (metis One_nat_def order_2 poly_eq_0_iff_dvd power_one_right) show "?rhs \ ?lhs" by (meson dvd_power dvd_trans neq0_conv order_1 poly_0 poly_eq_0_iff_dvd) qed lemma order_0I: "poly p a \ 0 \ order a p = 0" by (subst (asm) order_root) auto lemma order_unique_lemma: fixes p :: "'a::idom poly" assumes "[:-a, 1:] ^ n dvd p" "\ [:-a, 1:] ^ Suc n dvd p" shows "order a p = n" unfolding Polynomial.order_def by (metis (mono_tags, lifting) Least_equality assms not_less_eq_eq power_le_dvd) lemma order_mult: assumes "p * q \ 0" shows "order a (p * q) = order a p + order a q" proof - define i where "i \ order a p" define j where "j \ order a q" define t where "t \ [:-a, 1:]" have t_dvd_iff: "\u. t dvd u \ poly u a = 0" by (simp add: t_def dvd_iff_poly_eq_0) have dvd: "t ^ i dvd p" "t ^ j dvd q" and "\ t ^ Suc i dvd p" "\ t ^ Suc j dvd q" using assms i_def j_def order_1 order_2 t_def by auto then have "\ t ^ Suc(i + j) dvd p * q" by (elim dvdE) (simp add: power_add t_dvd_iff) moreover have "t ^ (i + j) dvd p * q" using dvd by (simp add: mult_dvd_mono power_add) ultimately show "order a (p * q) = i + j" using order_unique_lemma t_def by blast qed lemma order_smult: assumes "c \ 0" shows "order x (smult c p) = order x p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "smult c p = [:c:] * p" by simp also from assms False have "order x \ = order x [:c:] + order x p" by (subst order_mult) simp_all also have "order x [:c:] = 0" by (rule order_0I) (use assms in auto) finally show ?thesis by simp qed text \Next three lemmas contributed by Wenda Li\ lemma order_1_eq_0 [simp]:"order x 1 = 0" by (metis order_root poly_1 zero_neq_one) lemma order_uminus[simp]: "order x (-p) = order x p" by (metis neg_equal_0_iff_equal order_smult smult_1_left smult_minus_left) lemma order_power_n_n: "order a ([:-a,1:]^n)=n" proof (induct n) (*might be proved more concisely using nat_less_induct*) case 0 then show ?case by (metis order_root poly_1 power_0 zero_neq_one) next case (Suc n) have "order a ([:- a, 1:] ^ Suc n) = order a ([:- a, 1:] ^ n) + order a [:-a,1:]" by (metis (no_types, opaque_lifting) One_nat_def add_Suc_right monoid_add_class.add.right_neutral one_neq_zero order_mult pCons_eq_0_iff power_add power_eq_0_iff power_one_right) moreover have "order a [:-a,1:] = 1" unfolding order_def proof (rule Least_equality, rule notI) assume "[:- a, 1:] ^ Suc 1 dvd [:- a, 1:]" then have "degree ([:- a, 1:] ^ Suc 1) \ degree ([:- a, 1:])" by (rule dvd_imp_degree_le) auto then show False by auto next fix y assume *: "\ [:- a, 1:] ^ Suc y dvd [:- a, 1:]" show "1 \ y" proof (rule ccontr) assume "\ 1 \ y" then have "y = 0" by auto then have "[:- a, 1:] ^ Suc y dvd [:- a, 1:]" by auto with * show False by auto qed qed ultimately show ?case using Suc by auto qed lemma order_0_monom [simp]: "c \ 0 \ order 0 (monom c n) = n" using order_power_n_n[of 0 n] by (simp add: monom_altdef order_smult) lemma dvd_imp_order_le: "q \ 0 \ p dvd q \ Polynomial.order a p \ Polynomial.order a q" by (auto simp: order_mult elim: dvdE) text \Now justify the standard squarefree decomposition, i.e. \f / gcd f f'\.\ lemma order_divides: "[:-a, 1:] ^ n dvd p \ p = 0 \ n \ order a p" by (meson dvd_0_right not_less_eq_eq order_1 order_2 power_le_dvd) lemma order_decomp: assumes "p \ 0" shows "\q. p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" proof - from assms have *: "[:- a, 1:] ^ order a p dvd p" and **: "\ [:- a, 1:] ^ Suc (order a p) dvd p" by (auto dest: order) from * obtain q where q: "p = [:- a, 1:] ^ order a p * q" .. with ** have "\ [:- a, 1:] ^ Suc (order a p) dvd [:- a, 1:] ^ order a p * q" by simp then have "\ [:- a, 1:] ^ order a p * [:- a, 1:] dvd [:- a, 1:] ^ order a p * q" by simp with idom_class.dvd_mult_cancel_left [of "[:- a, 1:] ^ order a p" "[:- a, 1:]" q] have "\ [:- a, 1:] dvd q" by auto with q show ?thesis by blast qed lemma monom_1_dvd_iff: "p \ 0 \ monom 1 n dvd p \ n \ order 0 p" using order_divides[of 0 n p] by (simp add: monom_altdef) subsection \Additional induction rules on polynomials\ text \ An induction rule for induction over the roots of a polynomial with a certain property. (e.g. all positive roots) \ lemma poly_root_induct [case_names 0 no_roots root]: fixes p :: "'a :: idom poly" assumes "Q 0" and "\p. (\a. P a \ poly p a \ 0) \ Q p" and "\a p. P a \ Q p \ Q ([:a, -1:] * p)" shows "Q p" proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "p = 0") case True with assms(1) show ?thesis by simp next case False show ?thesis proof (cases "\a. P a \ poly p a = 0") case False then show ?thesis by (intro assms(2)) blast next case True then obtain a where a: "P a" "poly p a = 0" by blast then have "-[:-a, 1:] dvd p" by (subst minus_dvd_iff) (simp add: poly_eq_0_iff_dvd) then obtain q where q: "p = [:a, -1:] * q" by (elim dvdE) simp with False have "q \ 0" by auto have "degree p = Suc (degree q)" by (subst q, subst degree_mult_eq) (simp_all add: \q \ 0\) then have "Q q" by (intro less) simp with a(1) have "Q ([:a, -1:] * q)" by (rule assms(3)) with q show ?thesis by simp qed qed qed lemma dropWhile_replicate_append: "dropWhile ((=) a) (replicate n a @ ys) = dropWhile ((=) a) ys" by (induct n) simp_all lemma Poly_append_replicate_0: "Poly (xs @ replicate n 0) = Poly xs" by (subst coeffs_eq_iff) (simp_all add: strip_while_def dropWhile_replicate_append) text \ An induction rule for simultaneous induction over two polynomials, prepending one coefficient in each step. \ lemma poly_induct2 [case_names 0 pCons]: assumes "P 0 0" "\a p b q. P p q \ P (pCons a p) (pCons b q)" shows "P p q" proof - define n where "n = max (length (coeffs p)) (length (coeffs q))" define xs where "xs = coeffs p @ (replicate (n - length (coeffs p)) 0)" define ys where "ys = coeffs q @ (replicate (n - length (coeffs q)) 0)" have "length xs = length ys" by (simp add: xs_def ys_def n_def) then have "P (Poly xs) (Poly ys)" by (induct rule: list_induct2) (simp_all add: assms) also have "Poly xs = p" by (simp add: xs_def Poly_append_replicate_0) also have "Poly ys = q" by (simp add: ys_def Poly_append_replicate_0) finally show ?thesis . qed subsection \Composition of polynomials\ (* Several lemmas contributed by René Thiemann and Akihisa Yamada *) definition pcompose :: "'a::comm_semiring_0 poly \ 'a poly \ 'a poly" where "pcompose p q = fold_coeffs (\a c. [:a:] + q * c) p 0" notation pcompose (infixl "\\<^sub>p" 71) lemma pcompose_0 [simp]: "pcompose 0 q = 0" by (simp add: pcompose_def) lemma pcompose_pCons: "pcompose (pCons a p) q = [:a:] + q * pcompose p q" by (cases "p = 0 \ a = 0") (auto simp add: pcompose_def) lemma pcompose_1: "pcompose 1 p = 1" for p :: "'a::comm_semiring_1 poly" by (auto simp: one_pCons pcompose_pCons) lemma poly_pcompose: "poly (pcompose p q) x = poly p (poly q x)" by (induct p) (simp_all add: pcompose_pCons) lemma degree_pcompose_le: "degree (pcompose p q) \ degree p * degree q" proof (induction p) case (pCons a p) then show ?case proof (clarsimp simp add: pcompose_pCons) assume "degree (p \\<^sub>p q) \ degree p * degree q" "p \ 0" then have "degree (q * p \\<^sub>p q) \ degree q + degree p * degree q" by (meson add_le_cancel_left degree_mult_le dual_order.trans pCons.IH) then show "degree ([:a:] + q * p \\<^sub>p q) \ degree q + degree p * degree q" by (simp add: degree_add_le) qed qed auto lemma pcompose_add: "pcompose (p + q) r = pcompose p r + pcompose q r" for p q r :: "'a::{comm_semiring_0, ab_semigroup_add} poly" proof (induction p q rule: poly_induct2) case 0 then show ?case by simp next case (pCons a p b q) have "pcompose (pCons a p + pCons b q) r = [:a + b:] + r * pcompose p r + r * pcompose q r" by (simp_all add: pcompose_pCons pCons.IH algebra_simps) also have "[:a + b:] = [:a:] + [:b:]" by simp also have "\ + r * pcompose p r + r * pcompose q r = pcompose (pCons a p) r + pcompose (pCons b q) r" by (simp only: pcompose_pCons add_ac) finally show ?case . qed lemma pcompose_uminus: "pcompose (-p) r = -pcompose p r" for p r :: "'a::comm_ring poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_diff: "pcompose (p - q) r = pcompose p r - pcompose q r" for p q r :: "'a::comm_ring poly" using pcompose_add[of p "-q"] by (simp add: pcompose_uminus) lemma pcompose_smult: "pcompose (smult a p) r = smult a (pcompose p r)" for p r :: "'a::comm_semiring_0 poly" by (induct p) (simp_all add: pcompose_pCons pcompose_add smult_add_right) lemma pcompose_mult: "pcompose (p * q) r = pcompose p r * pcompose q r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_add pcompose_smult pcompose_pCons algebra_simps) lemma pcompose_assoc: "pcompose p (pcompose q r) = pcompose (pcompose p q) r" for p q r :: "'a::comm_semiring_0 poly" by (induct p arbitrary: q) (simp_all add: pcompose_pCons pcompose_add pcompose_mult) lemma pcompose_idR[simp]: "pcompose p [: 0, 1 :] = p" for p :: "'a::comm_semiring_1 poly" by (induct p) (simp_all add: pcompose_pCons) lemma pcompose_sum: "pcompose (sum f A) p = sum (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_add) lemma pcompose_prod: "pcompose (prod f A) p = prod (\i. pcompose (f i) p) A" by (induct A rule: infinite_finite_induct) (simp_all add: pcompose_1 pcompose_mult) lemma pcompose_const [simp]: "pcompose [:a:] q = [:a:]" by (subst pcompose_pCons) simp lemma pcompose_0': "pcompose p 0 = [:coeff p 0:]" by (induct p) (auto simp add: pcompose_pCons) lemma degree_pcompose: "degree (pcompose p q) = degree p * degree q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 show ?thesis proof (cases "p = 0") case True then show ?thesis by auto next case False from prems have "degree q = 0 \ pcompose p q = 0" by (auto simp add: degree_mult_eq_0) moreover have False if "pcompose p q = 0" "degree q \ 0" proof - from pCons.hyps(2) that have "degree p = 0" by auto then obtain a1 where "p = [:a1:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) with \pcompose p q = 0\ \p \ 0\ show False by auto qed ultimately have "degree (pCons a p) * degree q = 0" by auto moreover have "degree (pcompose (pCons a p) q) = 0" proof - from prems have "0 = max (degree [:a:]) (degree (q * pcompose p q))" by simp also have "\ \ degree ([:a:] + q * pcompose p q)" by (rule degree_add_le_max) finally show ?thesis by (auto simp add: pcompose_pCons) qed ultimately show ?thesis by simp qed next case prems: 2 then have "p \ 0" "q \ 0" "pcompose p q \ 0" by auto from prems degree_add_eq_right [of "[:a:]"] have "degree (pcompose (pCons a p) q) = degree (q * pcompose p q)" by (auto simp: pcompose_pCons) with pCons.hyps(2) degree_mult_eq[OF \q\0\ \pcompose p q\0\] show ?thesis by auto qed qed lemma pcompose_eq_0: fixes p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" assumes "pcompose p q = 0" "degree q > 0" shows "p = 0" proof - from assms degree_pcompose [of p q] have "degree p = 0" by auto then obtain a where "p = [:a:]" by (metis degree_pCons_eq_if gr0_conv_Suc neq0_conv pCons_cases) with assms(1) have "a = 0" by auto with \p = [:a:]\ show ?thesis by simp qed lemma lead_coeff_comp: fixes p q :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" assumes "degree q > 0" shows "lead_coeff (pcompose p q) = lead_coeff p * lead_coeff q ^ (degree p)" proof (induct p) case 0 then show ?case by auto next case (pCons a p) consider "degree (q * pcompose p q) = 0" | "degree (q * pcompose p q) > 0" by blast then show ?case proof cases case prems: 1 then have "pcompose p q = 0" by (metis assms degree_0 degree_mult_eq_0 neq0_conv) with pcompose_eq_0[OF _ \degree q > 0\] have "p = 0" by simp then show ?thesis by auto next case prems: 2 then have "degree [:a:] < degree (q * pcompose p q)" by simp then have "lead_coeff ([:a:] + q * p \\<^sub>p q) = lead_coeff (q * p \\<^sub>p q)" by (rule lead_coeff_add_le) then have "lead_coeff (pcompose (pCons a p) q) = lead_coeff (q * pcompose p q)" by (simp add: pcompose_pCons) also have "\ = lead_coeff q * (lead_coeff p * lead_coeff q ^ degree p)" using pCons.hyps(2) lead_coeff_mult[of q "pcompose p q"] by simp also have "\ = lead_coeff p * lead_coeff q ^ (degree p + 1)" by (auto simp: mult_ac) finally show ?thesis by auto qed qed subsection \Closure properties of coefficients\ context fixes R :: "'a :: comm_semiring_1 set" assumes R_0: "0 \ R" assumes R_plus: "\x y. x \ R \ y \ R \ x + y \ R" assumes R_mult: "\x y. x \ R \ y \ R \ x * y \ R" begin lemma coeff_mult_semiring_closed: assumes "\i. coeff p i \ R" "\i. coeff q i \ R" shows "coeff (p * q) i \ R" proof - have R_sum: "sum f A \ R" if "\x. x \ A \ f x \ R" for A and f :: "nat \ 'a" using that by (induction A rule: infinite_finite_induct) (auto intro: R_0 R_plus) show ?thesis unfolding coeff_mult by (auto intro!: R_sum R_mult assms) qed lemma coeff_pcompose_semiring_closed: assumes "\i. coeff p i \ R" "\i. coeff q i \ R" shows "coeff (pcompose p q) i \ R" using assms(1) proof (induction p arbitrary: i) case (pCons a p i) have [simp]: "a \ R" using pCons.prems[of 0] by auto have "coeff p i \ R" for i using pCons.prems[of "Suc i"] by auto hence "coeff (p \\<^sub>p q) i \ R" for i using pCons.prems by (intro pCons.IH) thus ?case by (auto simp: pcompose_pCons coeff_pCons split: nat.splits intro!: assms R_plus coeff_mult_semiring_closed) qed auto end subsection \Shifting polynomials\ definition poly_shift :: "nat \ 'a::zero poly \ 'a poly" where "poly_shift n p = Abs_poly (\i. coeff p (i + n))" lemma nth_default_drop: "nth_default x (drop n xs) m = nth_default x xs (m + n)" by (auto simp add: nth_default_def add_ac) lemma nth_default_take: "nth_default x (take n xs) m = (if m < n then nth_default x xs m else x)" by (auto simp add: nth_default_def add_ac) lemma coeff_poly_shift: "coeff (poly_shift n p) i = coeff p (i + n)" proof - from MOST_coeff_eq_0[of p] obtain m where "\k>m. coeff p k = 0" by (auto simp: MOST_nat) then have "\k>m. coeff p (k + n) = 0" by auto then have "\\<^sub>\k. coeff p (k + n) = 0" by (auto simp: MOST_nat) then show ?thesis by (simp add: poly_shift_def poly.Abs_poly_inverse) qed lemma poly_shift_id [simp]: "poly_shift 0 = (\x. x)" by (simp add: poly_eq_iff fun_eq_iff coeff_poly_shift) lemma poly_shift_0 [simp]: "poly_shift n 0 = 0" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_1: "poly_shift n 1 = (if n = 0 then 1 else 0)" by (simp add: poly_eq_iff coeff_poly_shift) lemma poly_shift_monom: "poly_shift n (monom c m) = (if m \ n then monom c (m - n) else 0)" by (auto simp add: poly_eq_iff coeff_poly_shift) lemma coeffs_shift_poly [code abstract]: "coeffs (poly_shift n p) = drop n (coeffs p)" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_shift nth_default_drop nth_default_coeffs_eq) qed subsection \Truncating polynomials\ definition poly_cutoff where "poly_cutoff n p = Abs_poly (\k. if k < n then coeff p k else 0)" lemma coeff_poly_cutoff: "coeff (poly_cutoff n p) k = (if k < n then coeff p k else 0)" unfolding poly_cutoff_def by (subst poly.Abs_poly_inverse) (auto simp: MOST_nat intro: exI[of _ n]) lemma poly_cutoff_0 [simp]: "poly_cutoff n 0 = 0" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma poly_cutoff_1 [simp]: "poly_cutoff n 1 = (if n = 0 then 0 else 1)" by (simp add: poly_eq_iff coeff_poly_cutoff) lemma coeffs_poly_cutoff [code abstract]: "coeffs (poly_cutoff n p) = strip_while ((=) 0) (take n (coeffs p))" proof (cases "strip_while ((=) 0) (take n (coeffs p)) = []") case True then have "coeff (poly_cutoff n p) k = 0" for k unfolding coeff_poly_cutoff by (auto simp: nth_default_coeffs_eq [symmetric] nth_default_def set_conv_nth) then have "poly_cutoff n p = 0" by (simp add: poly_eq_iff) then show ?thesis by (subst True) simp_all next case False have "no_trailing ((=) 0) (strip_while ((=) 0) (take n (coeffs p)))" by simp with False have "last (strip_while ((=) 0) (take n (coeffs p))) \ 0" unfolding no_trailing_unfold by auto then show ?thesis by (intro coeffs_eqI) (simp_all add: coeff_poly_cutoff nth_default_take nth_default_coeffs_eq) qed subsection \Reflecting polynomials\ definition reflect_poly :: "'a::zero poly \ 'a poly" where "reflect_poly p = Poly (rev (coeffs p))" lemma coeffs_reflect_poly [code abstract]: "coeffs (reflect_poly p) = rev (dropWhile ((=) 0) (coeffs p))" by (simp add: reflect_poly_def) lemma reflect_poly_0 [simp]: "reflect_poly 0 = 0" by (simp add: reflect_poly_def) lemma reflect_poly_1 [simp]: "reflect_poly 1 = 1" by (simp add: reflect_poly_def one_pCons) lemma coeff_reflect_poly: "coeff (reflect_poly p) n = (if n > degree p then 0 else coeff p (degree p - n))" by (cases "p = 0") (auto simp add: reflect_poly_def nth_default_def rev_nth degree_eq_length_coeffs coeffs_nth not_less dest: le_imp_less_Suc) lemma coeff_0_reflect_poly_0_iff [simp]: "coeff (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly) lemma reflect_poly_at_0_eq_0_iff [simp]: "poly (reflect_poly p) 0 = 0 \ p = 0" by (simp add: coeff_reflect_poly poly_0_coeff_0) lemma reflect_poly_pCons': "p \ 0 \ reflect_poly (pCons c p) = reflect_poly p + monom c (Suc (degree p))" by (intro poly_eqI) (auto simp: coeff_reflect_poly coeff_pCons not_less Suc_diff_le split: nat.split) lemma reflect_poly_const [simp]: "reflect_poly [:a:] = [:a:]" by (cases "a = 0") (simp_all add: reflect_poly_def) lemma poly_reflect_poly_nz: "x \ 0 \ poly (reflect_poly p) x = x ^ degree p * poly p (inverse x)" for x :: "'a::field" by (induct rule: pCons_induct) (simp_all add: field_simps reflect_poly_pCons' poly_monom) lemma coeff_0_reflect_poly [simp]: "coeff (reflect_poly p) 0 = lead_coeff p" by (simp add: coeff_reflect_poly) lemma poly_reflect_poly_0 [simp]: "poly (reflect_poly p) 0 = lead_coeff p" by (simp add: poly_0_coeff_0) lemma reflect_poly_reflect_poly [simp]: "coeff p 0 \ 0 \ reflect_poly (reflect_poly p) = p" by (cases p rule: pCons_cases) (simp add: reflect_poly_def ) lemma degree_reflect_poly_le: "degree (reflect_poly p) \ degree p" by (simp add: degree_eq_length_coeffs coeffs_reflect_poly length_dropWhile_le diff_le_mono) lemma reflect_poly_pCons: "a \ 0 \ reflect_poly (pCons a p) = Poly (rev (a # coeffs p))" by (subst coeffs_eq_iff) (simp add: coeffs_reflect_poly) lemma degree_reflect_poly_eq [simp]: "coeff p 0 \ 0 \ degree (reflect_poly p) = degree p" by (cases p rule: pCons_cases) (simp add: reflect_poly_pCons degree_eq_length_coeffs) (* TODO: does this work with zero divisors as well? Probably not. *) lemma reflect_poly_mult: "reflect_poly (p * q) = reflect_poly p * reflect_poly q" for p q :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" proof (cases "p = 0 \ q = 0") case False then have [simp]: "p \ 0" "q \ 0" by auto show ?thesis proof (rule poly_eqI) show "coeff (reflect_poly (p * q)) i = coeff (reflect_poly p * reflect_poly q) i" for i proof (cases "i \ degree (p * q)") case True define A where "A = {..i} \ {i - degree q..degree p}" define B where "B = {..degree p} \ {degree p - i..degree (p*q) - i}" let ?f = "\j. degree p - j" from True have "coeff (reflect_poly (p * q)) i = coeff (p * q) (degree (p * q) - i)" by (simp add: coeff_reflect_poly) also have "\ = (\j\degree (p * q) - i. coeff p j * coeff q (degree (p * q) - i - j))" by (simp add: coeff_mult) also have "\ = (\j\B. coeff p j * coeff q (degree (p * q) - i - j))" by (intro sum.mono_neutral_right) (auto simp: B_def degree_mult_eq not_le coeff_eq_0) also from True have "\ = (\j\A. coeff p (degree p - j) * coeff q (degree q - (i - j)))" by (intro sum.reindex_bij_witness[of _ ?f ?f]) (auto simp: A_def B_def degree_mult_eq add_ac) also have "\ = (\j\i. if j \ {i - degree q..degree p} then coeff p (degree p - j) * coeff q (degree q - (i - j)) else 0)" by (subst sum.inter_restrict [symmetric]) (simp_all add: A_def) also have "\ = coeff (reflect_poly p * reflect_poly q) i" by (fastforce simp: coeff_mult coeff_reflect_poly intro!: sum.cong) finally show ?thesis . qed (auto simp: coeff_mult coeff_reflect_poly coeff_eq_0 degree_mult_eq intro!: sum.neutral) qed qed auto lemma reflect_poly_smult: "reflect_poly (smult c p) = smult c (reflect_poly p)" for p :: "'a::{comm_semiring_0,semiring_no_zero_divisors} poly" using reflect_poly_mult[of "[:c:]" p] by simp lemma reflect_poly_power: "reflect_poly (p ^ n) = reflect_poly p ^ n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (induct n) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod: "reflect_poly (prod f A) = prod (\x. reflect_poly (f x)) A" for f :: "_ \ _::{comm_semiring_0,semiring_no_zero_divisors} poly" by (induct A rule: infinite_finite_induct) (simp_all add: reflect_poly_mult) lemma reflect_poly_prod_list: "reflect_poly (prod_list xs) = prod_list (map reflect_poly xs)" for xs :: "_::{comm_semiring_0,semiring_no_zero_divisors} poly list" by (induct xs) (simp_all add: reflect_poly_mult) lemma reflect_poly_Poly_nz: "no_trailing (HOL.eq 0) xs \ reflect_poly (Poly xs) = Poly (rev xs)" by (simp add: reflect_poly_def) lemmas reflect_poly_simps = reflect_poly_0 reflect_poly_1 reflect_poly_const reflect_poly_smult reflect_poly_mult reflect_poly_power reflect_poly_prod reflect_poly_prod_list subsection \Derivatives\ function pderiv :: "('a :: {comm_semiring_1,semiring_no_zero_divisors}) poly \ 'a poly" where "pderiv (pCons a p) = (if p = 0 then 0 else p + pCons 0 (pderiv p))" by (auto intro: pCons_cases) termination pderiv by (relation "measure degree") simp_all declare pderiv.simps[simp del] lemma pderiv_0 [simp]: "pderiv 0 = 0" using pderiv.simps [of 0 0] by simp lemma pderiv_pCons: "pderiv (pCons a p) = p + pCons 0 (pderiv p)" by (simp add: pderiv.simps) lemma pderiv_1 [simp]: "pderiv 1 = 0" by (simp add: one_pCons pderiv_pCons) lemma pderiv_of_nat [simp]: "pderiv (of_nat n) = 0" and pderiv_numeral [simp]: "pderiv (numeral m) = 0" by (simp_all add: of_nat_poly numeral_poly pderiv_pCons) lemma coeff_pderiv: "coeff (pderiv p) n = of_nat (Suc n) * coeff p (Suc n)" by (induct p arbitrary: n) (auto simp add: pderiv_pCons coeff_pCons algebra_simps split: nat.split) fun pderiv_coeffs_code :: "'a::{comm_semiring_1,semiring_no_zero_divisors} \ 'a list \ 'a list" where "pderiv_coeffs_code f (x # xs) = cCons (f * x) (pderiv_coeffs_code (f+1) xs)" | "pderiv_coeffs_code f [] = []" definition pderiv_coeffs :: "'a::{comm_semiring_1,semiring_no_zero_divisors} list \ 'a list" where "pderiv_coeffs xs = pderiv_coeffs_code 1 (tl xs)" (* Efficient code for pderiv contributed by René Thiemann and Akihisa Yamada *) lemma pderiv_coeffs_code: "nth_default 0 (pderiv_coeffs_code f xs) n = (f + of_nat n) * nth_default 0 xs n" proof (induct xs arbitrary: f n) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases n) case 0 then show ?thesis by (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") (auto simp: cCons_def) next case n: (Suc m) show ?thesis proof (cases "pderiv_coeffs_code (f + 1) xs = [] \ f * x = 0") case False then have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = nth_default 0 (pderiv_coeffs_code (f + 1) xs) m" by (auto simp: cCons_def n) also have "\ = (f + of_nat n) * nth_default 0 xs m" by (simp add: Cons n add_ac) finally show ?thesis by (simp add: n) next case True have empty: "pderiv_coeffs_code g xs = [] \ g + of_nat m = 0 \ nth_default 0 xs m = 0" for g proof (induct xs arbitrary: g m) case Nil then show ?case by simp next case (Cons x xs) from Cons(2) have empty: "pderiv_coeffs_code (g + 1) xs = []" and g: "g = 0 \ x = 0" by (auto simp: cCons_def split: if_splits) note IH = Cons(1)[OF empty] from IH[of m] IH[of "m - 1"] g show ?case by (cases m) (auto simp: field_simps) qed from True have "nth_default 0 (pderiv_coeffs_code f (x # xs)) n = 0" by (auto simp: cCons_def n) moreover from True have "(f + of_nat n) * nth_default 0 (x # xs) n = 0" by (simp add: n) (use empty[of "f+1"] in \auto simp: field_simps\) ultimately show ?thesis by simp qed qed qed lemma coeffs_pderiv_code [code abstract]: "coeffs (pderiv p) = pderiv_coeffs (coeffs p)" unfolding pderiv_coeffs_def proof (rule coeffs_eqI, unfold pderiv_coeffs_code coeff_pderiv, goal_cases) case (1 n) have id: "coeff p (Suc n) = nth_default 0 (map (\i. coeff p (Suc i)) [0.. degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" proof (cases "degree p") case 0 then show ?thesis by (metis degree_eq_zeroE pderiv.simps) next case (Suc n) then show ?thesis using coeff_0 coeff_pderiv degree_0 leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff by (metis coeff_0 coeff_pderiv degree_0 leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff) qed lemma degree_pderiv: "degree (pderiv p) = degree p - 1" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" proof - have "degree p - 1 \ degree (pderiv p)" proof (cases "degree p") case (Suc n) then show ?thesis by (metis coeff_pderiv degree_0 diff_Suc_1 le_degree leading_coeff_0_iff mult_eq_0_iff nat.distinct(1) of_nat_eq_0_iff) qed auto moreover have "\i>degree p - 1. coeff (pderiv p) i = 0" by (simp add: coeff_eq_0 coeff_pderiv) ultimately show ?thesis using order_antisym [OF degree_le] by blast qed lemma not_dvd_pderiv: fixes p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" assumes "degree p \ 0" shows "\ p dvd pderiv p" proof assume dvd: "p dvd pderiv p" then obtain q where p: "pderiv p = p * q" unfolding dvd_def by auto from dvd have le: "degree p \ degree (pderiv p)" by (simp add: assms dvd_imp_degree_le pderiv_eq_0_iff) from assms and this [unfolded degree_pderiv] show False by auto qed lemma dvd_pderiv_iff [simp]: "p dvd pderiv p \ degree p = 0" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" using not_dvd_pderiv[of p] by (auto simp: pderiv_eq_0_iff [symmetric]) lemma pderiv_singleton [simp]: "pderiv [:a:] = 0" by (simp add: pderiv_pCons) lemma pderiv_add: "pderiv (p + q) = pderiv p + pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_minus: "pderiv (- p :: 'a :: idom poly) = - pderiv p" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_diff: "pderiv ((p :: _ :: idom poly) - q) = pderiv p - pderiv q" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_smult: "pderiv (smult a p) = smult a (pderiv p)" by (rule poly_eqI) (simp add: coeff_pderiv algebra_simps) lemma pderiv_mult: "pderiv (p * q) = p * pderiv q + q * pderiv p" by (induct p) (auto simp: pderiv_add pderiv_smult pderiv_pCons algebra_simps) lemma pderiv_power_Suc: "pderiv (p ^ Suc n) = smult (of_nat (Suc n)) (p ^ n) * pderiv p" proof (induction n) case (Suc n) then show ?case by (simp add: pderiv_mult smult_add_left algebra_simps) qed auto lemma pderiv_pcompose: "pderiv (pcompose p q) = pcompose (pderiv p) q * pderiv q" by (induction p rule: pCons_induct) (auto simp: pcompose_pCons pderiv_add pderiv_mult pderiv_pCons pcompose_add algebra_simps) lemma pderiv_prod: "pderiv (prod f (as)) = (\a\as. prod f (as - {a}) * pderiv (f a))" proof (induct as rule: infinite_finite_induct) case (insert a as) then have id: "prod f (insert a as) = f a * prod f as" "\g. sum g (insert a as) = g a + sum g as" "insert a as - {a} = as" by auto have "prod f (insert a as - {b}) = f a * prod f (as - {b})" if "b \ as" for b proof - from \a \ as\ that have *: "insert a as - {b} = insert a (as - {b})" by auto show ?thesis unfolding * by (subst prod.insert) (use insert in auto) qed then show ?case unfolding id pderiv_mult insert(3) sum_distrib_left by (auto simp add: ac_simps intro!: sum.cong) qed auto lemma DERIV_pow2: "DERIV (\x. x ^ Suc n) x :> real (Suc n) * (x ^ n)" by (rule DERIV_cong, rule DERIV_pow) simp declare DERIV_pow2 [simp] DERIV_pow [simp] lemma DERIV_add_const: "DERIV f x :> D \ DERIV (\x. a + f x :: 'a::real_normed_field) x :> D" by (rule DERIV_cong, rule DERIV_add) auto lemma poly_DERIV [simp]: "DERIV (\x. poly p x) x :> poly (pderiv p) x" by (induct p) (auto intro!: derivative_eq_intros simp add: pderiv_pCons) lemma poly_isCont[simp]: fixes x::"'a::real_normed_field" shows "isCont (\x. poly p x) x" by (rule poly_DERIV [THEN DERIV_isCont]) lemma tendsto_poly [tendsto_intros]: "(f \ a) F \ ((\x. poly p (f x)) \ poly p a) F" for f :: "_ \ 'a::real_normed_field" by (rule isCont_tendsto_compose [OF poly_isCont]) lemma continuous_within_poly: "continuous (at z within s) (poly p)" for z :: "'a::{real_normed_field}" by (simp add: continuous_within tendsto_poly) lemma continuous_poly [continuous_intros]: "continuous F f \ continuous F (\x. poly p (f x))" for f :: "_ \ 'a::real_normed_field" unfolding continuous_def by (rule tendsto_poly) lemma continuous_on_poly [continuous_intros]: fixes p :: "'a :: {real_normed_field} poly" assumes "continuous_on A f" shows "continuous_on A (\x. poly p (f x))" by (metis DERIV_continuous_on assms continuous_on_compose2 poly_DERIV subset_UNIV) text \Consequences of the derivative theorem above.\ lemma poly_differentiable[simp]: "(\x. poly p x) differentiable (at x)" for x :: real by (simp add: real_differentiable_def) (blast intro: poly_DERIV) lemma poly_IVT_pos: "a < b \ poly p a < 0 \ 0 < poly p b \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using IVT [of "poly p" a 0 b] by (auto simp add: order_le_less) lemma poly_IVT_neg: "a < b \ 0 < poly p a \ poly p b < 0 \ \x. a < x \ x < b \ poly p x = 0" for a b :: real using poly_IVT_pos [where p = "- p"] by simp lemma poly_IVT: "a < b \ poly p a * poly p b < 0 \ \x>a. x < b \ poly p x = 0" for p :: "real poly" by (metis less_not_sym mult_less_0_iff poly_IVT_neg poly_IVT_pos) lemma poly_MVT: "a < b \ \x. a < x \ x < b \ poly p b - poly p a = (b - a) * poly (pderiv p) x" for a b :: real by (simp add: MVT2) lemma poly_MVT': fixes a b :: real assumes "{min a b..max a b} \ A" shows "\x\A. poly p b - poly p a = (b - a) * poly (pderiv p) x" proof (cases a b rule: linorder_cases) case less - from poly_MVT[OF less, of p] guess x by (elim exE conjE) + from poly_MVT[OF less, of p] obtain x + where "a < x" "x < b" "poly p b - poly p a = (b - a) * poly (pderiv p) x" + by auto then show ?thesis by (intro bexI[of _ x]) (auto intro!: subsetD[OF assms]) next case greater - from poly_MVT[OF greater, of p] guess x by (elim exE conjE) + from poly_MVT[OF greater, of p] obtain x + where "b < x" "x < a" "poly p a - poly p b = (a - b) * poly (pderiv p) x" by auto then show ?thesis by (intro bexI[of _ x]) (auto simp: algebra_simps intro!: subsetD[OF assms]) qed (use assms in auto) lemma poly_pinfty_gt_lc: fixes p :: "real poly" assumes "lead_coeff p > 0" shows "\n. \ x \ n. poly p x \ lead_coeff p" using assms proof (induct p) case 0 then show ?case by auto next case (pCons a p) from this(1) consider "a \ 0" "p = 0" | "p \ 0" by auto then show ?case proof cases case 1 then show ?thesis by auto next case 2 with pCons obtain n1 where gte_lcoeff: "\x\n1. lead_coeff p \ poly p x" by auto from pCons(3) \p \ 0\ have gt_0: "lead_coeff p > 0" by auto define n where "n = max n1 (1 + \a\ / lead_coeff p)" have "lead_coeff (pCons a p) \ poly (pCons a p) x" if "n \ x" for x proof - from gte_lcoeff that have "lead_coeff p \ poly p x" by (auto simp: n_def) with gt_0 have "\a\ / lead_coeff p \ \a\ / poly p x" and "poly p x > 0" by (auto intro: frac_le) with \n \ x\[unfolded n_def] have "x \ 1 + \a\ / poly p x" by auto with \lead_coeff p \ poly p x\ \poly p x > 0\ \p \ 0\ show "lead_coeff (pCons a p) \ poly (pCons a p) x" by (auto simp: field_simps) qed then show ?thesis by blast qed qed lemma lemma_order_pderiv1: "pderiv ([:- a, 1:] ^ Suc n * q) = [:- a, 1:] ^ Suc n * pderiv q + smult (of_nat (Suc n)) (q * [:- a, 1:] ^ n)" by (simp only: pderiv_mult pderiv_power_Suc) (simp del: power_Suc of_nat_Suc add: pderiv_pCons) lemma lemma_order_pderiv: fixes p :: "'a :: field_char_0 poly" assumes n: "0 < n" and pd: "pderiv p \ 0" and pe: "p = [:- a, 1:] ^ n * q" and nd: "\ [:- a, 1:] dvd q" shows "n = Suc (order a (pderiv p))" proof - from assms have "pderiv ([:- a, 1:] ^ n * q) \ 0" by auto from assms obtain n' where "n = Suc n'" "0 < Suc n'" "pderiv ([:- a, 1:] ^ Suc n' * q) \ 0" by (cases n) auto have "order a (pderiv ([:- a, 1:] ^ Suc n' * q)) = n'" proof (rule order_unique_lemma) show "[:- a, 1:] ^ n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" unfolding lemma_order_pderiv1 proof (rule dvd_add) show "[:- a, 1:] ^ n' dvd [:- a, 1:] ^ Suc n' * pderiv q" by (metis dvdI dvd_mult2 power_Suc2) show "[:- a, 1:] ^ n' dvd smult (of_nat (Suc n')) (q * [:- a, 1:] ^ n')" by (metis dvd_smult dvd_triv_right) qed have "k dvd k * pderiv q + smult (of_nat (Suc n')) l \ k dvd l" for k l by (auto simp del: of_nat_Suc simp: dvd_add_right_iff dvd_smult_iff) then show "\ [:- a, 1:] ^ Suc n' dvd pderiv ([:- a, 1:] ^ Suc n' * q)" unfolding lemma_order_pderiv1 by (metis nd dvd_mult_cancel_right power_not_zero pCons_eq_0_iff power_Suc zero_neq_one) qed then show ?thesis by (metis \n = Suc n'\ pe) qed lemma order_pderiv: "order a p = Suc (order a (pderiv p))" if "pderiv p \ 0" "order a p \ 0" for p :: "'a::field_char_0 poly" proof (cases "p = 0") case False obtain q where "p = [:- a, 1:] ^ order a p * q \ \ [:- a, 1:] dvd q" using False order_decomp by blast then show ?thesis using lemma_order_pderiv that by blast qed (use that in auto) lemma poly_squarefree_decomp_order: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and p: "p = q * d" and p': "pderiv p = e * d" and d: "d = r * p + s * pderiv p" shows "order a q = (if order a p = 0 then 0 else 1)" proof (rule classical) assume 1: "\ ?thesis" from \pderiv p \ 0\ have "p \ 0" by auto with p have "order a p = order a q + order a d" by (simp add: order_mult) with 1 have "order a p \ 0" by (auto split: if_splits) from \pderiv p \ 0\ \pderiv p = e * d\ have oapp: "order a (pderiv p) = order a e + order a d" by (simp add: order_mult) from \pderiv p \ 0\ \order a p \ 0\ have oap: "order a p = Suc (order a (pderiv p))" by (rule order_pderiv) from \p \ 0\ \p = q * d\ have "d \ 0" by simp have "[:- a, 1:] ^ order a (pderiv p) dvd r * p" by (metis dvd_trans dvd_triv_right oap order_1 power_Suc) then have "([:-a, 1:] ^ (order a (pderiv p))) dvd d" by (simp add: d order_1) with \d \ 0\ have "order a (pderiv p) \ order a d" by (simp add: order_divides) show ?thesis using \order a p = order a q + order a d\ and oapp oap and \order a (pderiv p) \ order a d\ by auto qed lemma poly_squarefree_decomp_order2: "pderiv p \ 0 \ p = q * d \ pderiv p = e * d \ d = r * p + s * pderiv p \ \a. order a q = (if order a p = 0 then 0 else 1)" for p :: "'a::field_char_0 poly" by (blast intro: poly_squarefree_decomp_order) lemma order_pderiv2: "pderiv p \ 0 \ order a p \ 0 \ order a (pderiv p) = n \ order a p = Suc n" for p :: "'a::field_char_0 poly" by (auto dest: order_pderiv) definition rsquarefree :: "'a::idom poly \ bool" where "rsquarefree p \ p \ 0 \ (\a. order a p = 0 \ order a p = 1)" lemma pderiv_iszero: "pderiv p = 0 \ \h. p = [:h:]" for p :: "'a::{semidom,semiring_char_0} poly" by (cases p) (auto simp: pderiv_eq_0_iff split: if_splits) lemma rsquarefree_roots: "rsquarefree p \ (\a. \ (poly p a = 0 \ poly (pderiv p) a = 0))" for p :: "'a::field_char_0 poly" proof (cases "p = 0") case False show ?thesis proof (cases "pderiv p = 0") case True with \p \ 0\ pderiv_iszero show ?thesis by (force simp add: order_0I rsquarefree_def) next case False with \p \ 0\ order_pderiv2 show ?thesis by (force simp add: rsquarefree_def order_root) qed qed (simp add: rsquarefree_def) lemma poly_squarefree_decomp: fixes p :: "'a::field_char_0 poly" assumes "pderiv p \ 0" and "p = q * d" and "pderiv p = e * d" and "d = r * p + s * pderiv p" shows "rsquarefree q \ (\a. poly q a = 0 \ poly p a = 0)" proof - from \pderiv p \ 0\ have "p \ 0" by auto with \p = q * d\ have "q \ 0" by simp from assms have "\a. order a q = (if order a p = 0 then 0 else 1)" by (rule poly_squarefree_decomp_order2) with \p \ 0\ \q \ 0\ show ?thesis by (simp add: rsquarefree_def order_root) qed subsection \Algebraic numbers\ text \ Algebraic numbers can be defined in two equivalent ways: all real numbers that are roots of rational polynomials or of integer polynomials. The Algebraic-Numbers AFP entry uses the rational definition, but we need the integer definition. The equivalence is obvious since any rational polynomial can be multiplied with the LCM of its coefficients, yielding an integer polynomial with the same roots. \ definition algebraic :: "'a :: field_char_0 \ bool" where "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" lemma algebraicI: "(\i. coeff p i \ \) \ p \ 0 \ poly p x = 0 \ algebraic x" unfolding algebraic_def by blast lemma algebraicE: assumes "algebraic x" obtains p where "\i. coeff p i \ \" "p \ 0" "poly p x = 0" using assms unfolding algebraic_def by blast lemma algebraic_altdef: "algebraic x \ (\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0)" for p :: "'a::field_char_0 poly" proof safe fix p assume rat: "\i. coeff p i \ \" and root: "poly p x = 0" and nz: "p \ 0" define cs where "cs = coeffs p" from rat have "\c\range (coeff p). \c'. c = of_rat c'" unfolding Rats_def by blast then obtain f where f: "coeff p i = of_rat (f (coeff p i))" for i by (subst (asm) bchoice_iff) blast define cs' where "cs' = map (quotient_of \ f) (coeffs p)" define d where "d = Lcm (set (map snd cs'))" define p' where "p' = smult (of_int d) p" have "coeff p' n \ \" for n proof (cases "n \ degree p") case True define c where "c = coeff p n" define a where "a = fst (quotient_of (f (coeff p n)))" define b where "b = snd (quotient_of (f (coeff p n)))" have b_pos: "b > 0" unfolding b_def using quotient_of_denom_pos' by simp have "coeff p' n = of_int d * coeff p n" by (simp add: p'_def) also have "coeff p n = of_rat (of_int a / of_int b)" unfolding a_def b_def by (subst quotient_of_div [of "f (coeff p n)", symmetric]) (simp_all add: f [symmetric]) also have "of_int d * \ = of_rat (of_int (a*d) / of_int b)" by (simp add: of_rat_mult of_rat_divide) also from nz True have "b \ snd ` set cs'" by (force simp: cs'_def o_def b_def coeffs_def simp del: upt_Suc) then have "b dvd (a * d)" by (simp add: d_def) then have "of_int (a * d) / of_int b \ (\ :: rat set)" by (rule of_int_divide_in_Ints) then have "of_rat (of_int (a * d) / of_int b) \ \" by (elim Ints_cases) auto finally show ?thesis . next case False then show ?thesis by (auto simp: p'_def not_le coeff_eq_0) qed moreover have "set (map snd cs') \ {0<..}" unfolding cs'_def using quotient_of_denom_pos' by (auto simp: coeffs_def simp del: upt_Suc) then have "d \ 0" unfolding d_def by (induct cs') simp_all with nz have "p' \ 0" by (simp add: p'_def) moreover from root have "poly p' x = 0" by (simp add: p'_def) ultimately show "algebraic x" unfolding algebraic_def by blast next assume "algebraic x" then obtain p where p: "coeff p i \ \" "poly p x = 0" "p \ 0" for i by (force simp: algebraic_def) moreover have "coeff p i \ \ \ coeff p i \ \" for i by (elim Ints_cases) simp ultimately show "\p. (\i. coeff p i \ \) \ p \ 0 \ poly p x = 0" by auto qed subsection \Algebraic integers\ inductive algebraic_int :: "'a :: field \ bool" where "\lead_coeff p = 1; \i. coeff p i \ \; poly p x = 0\ \ algebraic_int x" lemma algebraic_int_altdef_ipoly: fixes x :: "'a :: field_char_0" shows "algebraic_int x \ (\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1)" proof assume "algebraic_int x" then obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto elim: algebraic_int.cases) define the_int where "the_int = (\x::'a. THE r. x = of_int r)" define p' where "p' = map_poly the_int p" have of_int_the_int: "of_int (the_int x) = x" if "x \ \" for x unfolding the_int_def by (rule sym, rule theI') (insert that, auto simp: Ints_def) have the_int_0_iff: "the_int x = 0 \ x = 0" if "x \ \" for x using of_int_the_int[OF that] by auto have [simp]: "the_int 0 = 0" by (subst the_int_0_iff) auto have "map_poly of_int p' = map_poly (of_int \ the_int) p" by (simp add: p'_def map_poly_map_poly) also from p of_int_the_int have "\ = p" by (subst poly_eq_iff) (auto simp: coeff_map_poly) finally have p_p': "map_poly of_int p' = p" . show "(\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1)" proof (intro exI conjI notI) from p show "poly (map_poly of_int p') x = 0" by (simp add: p_p') next show "lead_coeff p' = 1" using p by (simp flip: p_p' add: degree_map_poly coeff_map_poly) qed next assume "\p. poly (map_poly of_int p) x = 0 \ lead_coeff p = 1" then obtain p where p: "poly (map_poly of_int p) x = 0" "lead_coeff p = 1" by auto define p' where "p' = (map_poly of_int p :: 'a poly)" from p have "lead_coeff p' = 1" "poly p' x = 0" "\i. coeff p' i \ \" by (auto simp: p'_def coeff_map_poly degree_map_poly) thus "algebraic_int x" by (intro algebraic_int.intros) qed theorem rational_algebraic_int_is_int: assumes "algebraic_int x" and "x \ \" shows "x \ \" proof - from assms(2) obtain a b where ab: "b > 0" "Rings.coprime a b" and x_eq: "x = of_int a / of_int b" by (auto elim: Rats_cases') from \b > 0\ have [simp]: "b \ 0" by auto from assms(1) obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto simp: algebraic_int.simps) define q :: "'a poly" where "q = [:-of_int a, of_int b:]" have "poly q x = 0" "q \ 0" "\i. coeff q i \ \" by (auto simp: x_eq q_def coeff_pCons split: nat.splits) define n where "n = degree p" have "n > 0" using p by (intro Nat.gr0I) (auto simp: n_def elim!: degree_eq_zeroE) have "(\i \" using p by auto then obtain R where R: "of_int R = (\i = (\i\n. coeff p i * x ^ i * of_int b ^ n)" by (simp add: poly_altdef n_def sum_distrib_right) also have "\ = (\i\n. coeff p i * of_int (a ^ i * b ^ (n - i)))" by (intro sum.cong) (auto simp: x_eq field_simps simp flip: power_add) also have "{..n} = insert n {..n > 0\ by auto also have "(\i\\. coeff p i * of_int (a ^ i * b ^ (n - i))) = coeff p n * of_int (a ^ n) + (\iii = of_int (b * R)" by (simp add: R sum_distrib_left sum_distrib_right mult_ac) finally have "of_int (a ^ n) = (-of_int (b * R) :: 'a)" by (auto simp: add_eq_0_iff) hence "a ^ n = -b * R" by (simp flip: of_int_mult of_int_power of_int_minus) hence "b dvd a ^ n" by simp with \Rings.coprime a b\ have "b dvd 1" by (meson coprime_power_left_iff dvd_refl not_coprimeI) with x_eq and \b > 0\ show ?thesis by auto qed lemma algebraic_int_imp_algebraic [dest]: "algebraic_int x \ algebraic x" by (auto simp: algebraic_int.simps algebraic_def) lemma int_imp_algebraic_int: assumes "x \ \" shows "algebraic_int x" proof show "\i. coeff [:-x, 1:] i \ \" using assms by (auto simp: coeff_pCons split: nat.splits) qed auto lemma algebraic_int_0 [simp, intro]: "algebraic_int 0" and algebraic_int_1 [simp, intro]: "algebraic_int 1" and algebraic_int_numeral [simp, intro]: "algebraic_int (numeral n)" and algebraic_int_of_nat [simp, intro]: "algebraic_int (of_nat k)" and algebraic_int_of_int [simp, intro]: "algebraic_int (of_int m)" by (simp_all add: int_imp_algebraic_int) lemma algebraic_int_ii [simp, intro]: "algebraic_int \" proof show "poly [:1, 0, 1:] \ = 0" by simp qed (auto simp: coeff_pCons split: nat.splits) lemma algebraic_int_minus [intro]: assumes "algebraic_int x" shows "algebraic_int (-x)" proof - from assms obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto simp: algebraic_int.simps) define s where "s = (if even (degree p) then 1 else -1 :: 'a)" define q where "q = Polynomial.smult s (pcompose p [:0, -1:])" have "lead_coeff q = s * lead_coeff (pcompose p [:0, -1:])" by (simp add: q_def) also have "lead_coeff (pcompose p [:0, -1:]) = lead_coeff p * (- 1) ^ degree p" by (subst lead_coeff_comp) auto finally have "poly q (-x) = 0" and "lead_coeff q = 1" using p by (auto simp: q_def poly_pcompose s_def) moreover have "coeff q i \ \" for i proof - have "coeff (pcompose p [:0, -1:]) i \ \" using p by (intro coeff_pcompose_semiring_closed) (auto simp: coeff_pCons split: nat.splits) thus ?thesis by (simp add: q_def s_def) qed ultimately show ?thesis by (auto simp: algebraic_int.simps) qed lemma algebraic_int_minus_iff [simp]: "algebraic_int (-x) \ algebraic_int (x :: 'a :: field_char_0)" using algebraic_int_minus[of x] algebraic_int_minus[of "-x"] by auto lemma algebraic_int_inverse [intro]: assumes "poly p x = 0" and "\i. coeff p i \ \" and "coeff p 0 = 1" shows "algebraic_int (inverse x)" proof from assms have [simp]: "x \ 0" by (auto simp: poly_0_coeff_0) show "poly (reflect_poly p) (inverse x) = 0" using assms by (simp add: poly_reflect_poly_nz) qed (use assms in \auto simp: coeff_reflect_poly\) lemma algebraic_int_root: assumes "algebraic_int y" and "poly p x = y" and "\i. coeff p i \ \" and "lead_coeff p = 1" and "degree p > 0" shows "algebraic_int x" proof - from assms obtain q where q: "poly q y = 0" "\i. coeff q i \ \" "lead_coeff q = 1" by (auto simp: algebraic_int.simps) show ?thesis proof from assms q show "lead_coeff (pcompose q p) = 1" by (subst lead_coeff_comp) auto from assms q show "\i. coeff (pcompose q p) i \ \" by (intro allI coeff_pcompose_semiring_closed) auto show "poly (pcompose q p) x = 0" using assms q by (simp add: poly_pcompose) qed qed lemma algebraic_int_abs_real [simp]: "algebraic_int \x :: real\ \ algebraic_int x" by (auto simp: abs_if) lemma algebraic_int_nth_root_real [intro]: assumes "algebraic_int x" shows "algebraic_int (root n x)" proof (cases "n = 0") case False show ?thesis proof (rule algebraic_int_root) show "poly (monom 1 n) (root n x) = (if even n then \x\ else x)" using sgn_power_root[of n x] False by (auto simp add: poly_monom sgn_if split: if_splits) qed (use False assms in \auto simp: degree_monom_eq\) qed auto lemma algebraic_int_sqrt [intro]: "algebraic_int x \ algebraic_int (sqrt x)" by (auto simp: sqrt_def) lemma algebraic_int_csqrt [intro]: "algebraic_int x \ algebraic_int (csqrt x)" by (rule algebraic_int_root[where p = "monom 1 2"]) (auto simp: poly_monom degree_monom_eq) lemma poly_map_poly_cnj [simp]: "poly (map_poly cnj p) x = cnj (poly p (cnj x))" by (induction p) (auto simp: map_poly_pCons) lemma algebraic_int_cnj [intro]: assumes "algebraic_int x" shows "algebraic_int (cnj x)" proof - from assms obtain p where p: "lead_coeff p = 1" "\i. coeff p i \ \" "poly p x = 0" by (auto simp: algebraic_int.simps) show ?thesis proof show "poly (map_poly cnj p) (cnj x) = 0" using p by simp show "lead_coeff (map_poly cnj p) = 1" using p by (simp add: coeff_map_poly degree_map_poly) show "\i. coeff (map_poly cnj p) i \ \" using p by (auto simp: coeff_map_poly) qed qed lemma algebraic_int_cnj_iff [simp]: "algebraic_int (cnj x) \ algebraic_int x" using algebraic_int_cnj[of x] algebraic_int_cnj[of "cnj x"] by auto lemma algebraic_int_of_real [intro]: assumes "algebraic_int x" shows "algebraic_int (of_real x)" proof - from assms obtain p where p: "poly p x = 0" "\i. coeff p i \ \" "lead_coeff p = 1" by (auto simp: algebraic_int.simps) show "algebraic_int (of_real x :: 'a)" proof have "poly (map_poly of_real p) (of_real x) = (of_real (poly p x) :: 'a)" by (induction p) (auto simp: map_poly_pCons) thus "poly (map_poly of_real p) (of_real x) = (0 :: 'a)" using p by simp qed (use p in \auto simp: coeff_map_poly degree_map_poly\) qed lemma algebraic_int_of_real_iff [simp]: "algebraic_int (of_real x :: 'a :: {field_char_0, real_algebra_1}) \ algebraic_int x" proof assume "algebraic_int (of_real x :: 'a)" then obtain p where p: "poly (map_poly of_int p) (of_real x :: 'a) = 0" "lead_coeff p = 1" by (auto simp: algebraic_int_altdef_ipoly) show "algebraic_int x" unfolding algebraic_int_altdef_ipoly proof (intro exI[of _ p] conjI) have "of_real (poly (map_poly real_of_int p) x) = poly (map_poly of_int p) (of_real x :: 'a)" by (induction p) (auto simp: map_poly_pCons) also note p(1) finally show "poly (map_poly real_of_int p) x = 0" by simp qed (use p in auto) qed auto subsection \Division of polynomials\ subsubsection \Division in general\ instantiation poly :: (idom_divide) idom_divide begin fun divide_poly_main :: "'a \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly" where "divide_poly_main lc q r d dr (Suc n) = (let cr = coeff r dr; a = cr div lc; mon = monom a n in if False \ a * lc = cr then \ \\False \\ is only because of problem in function-package\ divide_poly_main lc (q + mon) (r - mon * d) d (dr - 1) n else 0)" | "divide_poly_main lc q r d dr 0 = q" definition divide_poly :: "'a poly \ 'a poly \ 'a poly" where "divide_poly f g = (if g = 0 then 0 else divide_poly_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)))" lemma divide_poly_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree (d * r) \ dr" "divide_poly_main lc q (d * r) d dr n = q'" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ d * r = 0" shows "q' = q + r" using assms(3-) proof (induct n arbitrary: q r dr) case (Suc n) let ?rr = "d * r" let ?a = "coeff ?rr dr" let ?qq = "?a div lc" define b where [simp]: "b = monom ?qq n" let ?rrr = "d * (r - b)" let ?qqq = "q + b" note res = Suc(3) from Suc(4) have dr: "dr = n + degree d" by auto from d have lc: "lc \ 0" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by simp next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have c2: "coeff (b * d) dr = lc * coeff b n" . have rrr: "?rrr = ?rr - b * d" by (simp add: field_simps) have c1: "coeff (d * r) dr = lc * coeff r n" proof (cases "degree r = n") case True with Suc(2) show ?thesis unfolding dr using coeff_mult_degree_sum[of d r] d by (auto simp: ac_simps) next case False from dr Suc(2) have "degree r \ n" by auto (metis add.commute add_le_cancel_left d(1) degree_0 degree_mult_eq diff_is_0_eq diff_zero le_cases) with False have r_n: "degree r < n" by auto then have right: "lc * coeff r n = 0" by (simp add: coeff_eq_0) have "coeff (d * r) dr = coeff (d * r) (degree d + n)" by (simp add: dr ac_simps) also from r_n have "\ = 0" by (metis False Suc.prems(1) add.commute add_left_imp_eq coeff_degree_mult coeff_eq_0 coeff_mult_degree_sum degree_mult_le dr le_eq_less_or_eq) finally show ?thesis by (simp only: right) qed have c0: "coeff ?rrr dr = 0" and id: "lc * (coeff (d * r) dr div lc) = coeff (d * r) dr" unfolding rrr coeff_diff c2 unfolding b_def coeff_monom coeff_smult c1 using lc by auto from res[unfolded divide_poly_main.simps[of lc q] Let_def] id have res: "divide_poly_main lc ?qqq ?rrr d (dr - 1) n = q'" by (simp del: divide_poly_main.simps add: field_simps) note IH = Suc(1)[OF _ res] from Suc(4) have dr: "dr = n + degree d" by auto from Suc(2) have deg_rr: "degree ?rr \ dr" by auto have deg_bd: "degree (b * d) \ dr" unfolding dr b_def by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" unfolding rrr by (rule degree_diff_le[OF deg_rr deg_bd]) with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp from degree_eq_zeroE[OF this] obtain a where rrr: "?rrr = [:a:]" by metis show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed from IH[OF deg_rrr this] show ?case by simp next case 0 show ?case proof (cases "r = 0") case True with 0 show ?thesis by auto next case False from d False have "degree (d * r) = degree d + degree r" by (subst degree_mult_eq) auto with 0 d show ?thesis by auto qed qed lemma divide_poly_main_0: "divide_poly_main 0 0 r d dr n = 0" proof (induct n arbitrary: r d dr) case 0 then show ?case by simp next case Suc show ?case unfolding divide_poly_main.simps[of _ _ r] Let_def by (simp add: Suc del: divide_poly_main.simps) qed lemma divide_poly: assumes g: "g \ 0" shows "(f * g) div g = (f :: 'a poly)" proof - have len: "length (coeffs f) = Suc (degree f)" if "f \ 0" for f :: "'a poly" using that unfolding degree_eq_length_coeffs by auto have "divide_poly_main (coeff g (degree g)) 0 (g * f) g (degree (g * f)) (1 + length (coeffs (g * f)) - length (coeffs g)) = (f * g) div g" by (simp add: divide_poly_def Let_def ac_simps) note main = divide_poly_main[OF g refl le_refl this] have "(f * g) div g = 0 + f" proof (rule main, goal_cases) case 1 show ?case proof (cases "f = 0") case True with g show ?thesis by (auto simp: degree_eq_length_coeffs) next case False with g have fg: "g * f \ 0" by auto show ?thesis unfolding len[OF fg] len[OF g] by auto qed qed then show ?thesis by simp qed lemma divide_poly_0: "f div 0 = 0" for f :: "'a poly" by (simp add: divide_poly_def Let_def divide_poly_main_0) instance by standard (auto simp: divide_poly divide_poly_0) end instance poly :: (idom_divide) algebraic_semidom .. lemma div_const_poly_conv_map_poly: assumes "[:c:] dvd p" shows "p div [:c:] = map_poly (\x. x div c) p" proof (cases "c = 0") case True then show ?thesis by (auto intro!: poly_eqI simp: coeff_map_poly) next case False from assms obtain q where p: "p = [:c:] * q" by (rule dvdE) moreover { have "smult c q = [:c:] * q" by simp also have "\ div [:c:] = q" by (rule nonzero_mult_div_cancel_left) (use False in auto) finally have "smult c q div [:c:] = q" . } ultimately show ?thesis by (intro poly_eqI) (auto simp: coeff_map_poly False) qed lemma is_unit_monom_0: fixes a :: "'a::field" assumes "a \ 0" shows "is_unit (monom a 0)" proof from assms show "1 = monom a 0 * monom (inverse a) 0" by (simp add: mult_monom) qed lemma is_unit_triv: "a \ 0 \ is_unit [:a:]" for a :: "'a::field" by (simp add: is_unit_monom_0 monom_0 [symmetric]) lemma is_unit_iff_degree: fixes p :: "'a::field poly" assumes "p \ 0" shows "is_unit p \ degree p = 0" (is "?lhs \ ?rhs") proof assume ?rhs then obtain a where "p = [:a:]" by (rule degree_eq_zeroE) with assms show ?lhs by (simp add: is_unit_triv) next assume ?lhs then obtain q where "q \ 0" "p * q = 1" .. then have "degree (p * q) = degree 1" by simp with \p \ 0\ \q \ 0\ have "degree p + degree q = 0" by (simp add: degree_mult_eq) then show ?rhs by simp qed lemma is_unit_pCons_iff: "is_unit (pCons a p) \ p = 0 \ a \ 0" for p :: "'a::field poly" by (cases "p = 0") (auto simp: is_unit_triv is_unit_iff_degree) lemma is_unit_monom_trivial: "is_unit p \ monom (coeff p (degree p)) 0 = p" for p :: "'a::field poly" by (cases p) (simp_all add: monom_0 is_unit_pCons_iff) lemma is_unit_const_poly_iff: "[:c:] dvd 1 \ c dvd 1" for c :: "'a::{comm_semiring_1,semiring_no_zero_divisors}" by (auto simp: one_pCons) lemma is_unit_polyE: fixes p :: "'a :: {comm_semiring_1,semiring_no_zero_divisors} poly" assumes "p dvd 1" obtains c where "p = [:c:]" "c dvd 1" proof - from assms obtain q where "1 = p * q" by (rule dvdE) then have "p \ 0" and "q \ 0" by auto from \1 = p * q\ have "degree 1 = degree (p * q)" by simp also from \p \ 0\ and \q \ 0\ have "\ = degree p + degree q" by (simp add: degree_mult_eq) finally have "degree p = 0" by simp with degree_eq_zeroE obtain c where c: "p = [:c:]" . with \p dvd 1\ have "c dvd 1" by (simp add: is_unit_const_poly_iff) with c show thesis .. qed lemma is_unit_polyE': fixes p :: "'a::field poly" assumes "is_unit p" obtains a where "p = monom a 0" and "a \ 0" proof - obtain a q where "p = pCons a q" by (cases p) with assms have "p = [:a:]" and "a \ 0" by (simp_all add: is_unit_pCons_iff) with that show thesis by (simp add: monom_0) qed lemma is_unit_poly_iff: "p dvd 1 \ (\c. p = [:c:] \ c dvd 1)" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors} poly" by (auto elim: is_unit_polyE simp add: is_unit_const_poly_iff) subsubsection \Pseudo-Division\ text \This part is by René Thiemann and Akihisa Yamada.\ fun pseudo_divmod_main :: "'a :: comm_ring_1 \ 'a poly \ 'a poly \ 'a poly \ nat \ nat \ 'a poly \ 'a poly" where "pseudo_divmod_main lc q r d dr (Suc n) = (let rr = smult lc r; qq = coeff r dr; rrr = rr - monom qq n * d; qqq = smult lc q + monom qq n in pseudo_divmod_main lc qqq rrr d (dr - 1) n)" | "pseudo_divmod_main lc q r d dr 0 = (q,r)" definition pseudo_divmod :: "'a :: comm_ring_1 poly \ 'a poly \ 'a poly \ 'a poly" where "pseudo_divmod p q \ if q = 0 then (0, p) else pseudo_divmod_main (coeff q (degree q)) 0 p q (degree p) (1 + length (coeffs p) - length (coeffs q))" lemma pseudo_divmod_main: assumes d: "d \ 0" "lc = coeff d (degree d)" and "degree r \ dr" "pseudo_divmod_main lc q r d dr n = (q',r')" and "n = 1 + dr - degree d \ dr = 0 \ n = 0 \ r = 0" shows "(r' = 0 \ degree r' < degree d) \ smult (lc^n) (d * q + r) = d * q' + r'" using assms(3-) proof (induct n arbitrary: q r dr) case 0 then show ?case by auto next case (Suc n) let ?rr = "smult lc r" let ?qq = "coeff r dr" define b where [simp]: "b = monom ?qq n" let ?rrr = "?rr - b * d" let ?qqq = "smult lc q + b" note res = Suc(3) from res[unfolded pseudo_divmod_main.simps[of lc q] Let_def] have res: "pseudo_divmod_main lc ?qqq ?rrr d (dr - 1) n = (q',r')" by (simp del: pseudo_divmod_main.simps) from Suc(4) have dr: "dr = n + degree d" by auto have "coeff (b * d) dr = coeff b n * coeff d (degree d)" proof (cases "?qq = 0") case True then show ?thesis by auto next case False then have n: "n = degree b" by (simp add: degree_monom_eq) show ?thesis unfolding n dr by (simp add: coeff_mult_degree_sum) qed also have "\ = lc * coeff b n" by (simp add: d) finally have "coeff (b * d) dr = lc * coeff b n" . moreover have "coeff ?rr dr = lc * coeff r dr" by simp ultimately have c0: "coeff ?rrr dr = 0" by auto from Suc(4) have dr: "dr = n + degree d" by auto have deg_rr: "degree ?rr \ dr" using Suc(2) degree_smult_le dual_order.trans by blast have deg_bd: "degree (b * d) \ dr" unfolding dr by (rule order.trans[OF degree_mult_le]) (auto simp: degree_monom_le) have "degree ?rrr \ dr" using degree_diff_le[OF deg_rr deg_bd] by auto with c0 have deg_rrr: "degree ?rrr \ (dr - 1)" by (rule coeff_0_degree_minus_1) have "n = 1 + (dr - 1) - degree d \ dr - 1 = 0 \ n = 0 \ ?rrr = 0" proof (cases dr) case 0 with Suc(4) have 0: "dr = 0" "n = 0" "degree d = 0" by auto with deg_rrr have "degree ?rrr = 0" by simp then have "\a. ?rrr = [:a:]" by (metis degree_pCons_eq_if old.nat.distinct(2) pCons_cases) from this obtain a where rrr: "?rrr = [:a:]" by auto show ?thesis unfolding 0 using c0 unfolding rrr 0 by simp next case _: Suc with Suc(4) show ?thesis by auto qed note IH = Suc(1)[OF deg_rrr res this] show ?case proof (intro conjI) from IH show "r' = 0 \ degree r' < degree d" by blast show "smult (lc ^ Suc n) (d * q + r) = d * q' + r'" unfolding IH[THEN conjunct2,symmetric] by (simp add: field_simps smult_add_right) qed qed lemma pseudo_divmod: assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" shows "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" (is ?A) and "r = 0 \ degree r < degree g" (is ?B) proof - from *[unfolded pseudo_divmod_def Let_def] have "pseudo_divmod_main (coeff g (degree g)) 0 f g (degree f) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (auto simp: g) note main = pseudo_divmod_main[OF _ _ _ this, OF g refl le_refl] from g have "1 + length (coeffs f) - length (coeffs g) = 1 + degree f - degree g \ degree f = 0 \ 1 + length (coeffs f) - length (coeffs g) = 0 \ f = 0" by (cases "f = 0"; cases "coeffs g") (auto simp: degree_eq_length_coeffs) note main' = main[OF this] then show "r = 0 \ degree r < degree g" by auto show "smult (coeff g (degree g) ^ (Suc (degree f) - degree g)) f = g * q + r" by (subst main'[THEN conjunct2, symmetric], simp add: degree_eq_length_coeffs, cases "f = 0"; cases "coeffs g", use g in auto) qed definition "pseudo_mod_main lc r d dr n = snd (pseudo_divmod_main lc 0 r d dr n)" lemma snd_pseudo_divmod_main: "snd (pseudo_divmod_main lc q r d dr n) = snd (pseudo_divmod_main lc q' r d dr n)" by (induct n arbitrary: q q' lc r d dr) (simp_all add: Let_def) definition pseudo_mod :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly \ 'a poly \ 'a poly" where "pseudo_mod f g = snd (pseudo_divmod f g)" lemma pseudo_mod: fixes f g :: "'a::{comm_ring_1,semiring_1_no_zero_divisors} poly" defines "r \ pseudo_mod f g" assumes g: "g \ 0" shows "\a q. a \ 0 \ smult a f = g * q + r" "r = 0 \ degree r < degree g" proof - let ?cg = "coeff g (degree g)" let ?cge = "?cg ^ (Suc (degree f) - degree g)" define a where "a = ?cge" from r_def[unfolded pseudo_mod_def] obtain q where pdm: "pseudo_divmod f g = (q, r)" by (cases "pseudo_divmod f g") auto from pseudo_divmod[OF g pdm] have id: "smult a f = g * q + r" and "r = 0 \ degree r < degree g" by (auto simp: a_def) show "r = 0 \ degree r < degree g" by fact from g have "a \ 0" by (auto simp: a_def) with id show "\a q. a \ 0 \ smult a f = g * q + r" by auto qed lemma fst_pseudo_divmod_main_as_divide_poly_main: assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "fst (pseudo_divmod_main lc q r d dr n) = divide_poly_main lc (smult (lc^n) q) (smult (lc^n) r) d dr n" proof (induct n arbitrary: q r dr) case 0 then show ?case by simp next case (Suc n) note lc0 = leading_coeff_neq_0[OF d, folded lc] then have "pseudo_divmod_main lc q r d dr (Suc n) = pseudo_divmod_main lc (smult lc q + monom (coeff r dr) n) (smult lc r - monom (coeff r dr) n * d) d (dr - 1) n" by (simp add: Let_def ac_simps) also have "fst \ = divide_poly_main lc (smult (lc^n) (smult lc q + monom (coeff r dr) n)) (smult (lc^n) (smult lc r - monom (coeff r dr) n * d)) d (dr - 1) n" by (simp only: Suc[unfolded divide_poly_main.simps Let_def]) also have "\ = divide_poly_main lc (smult (lc ^ Suc n) q) (smult (lc ^ Suc n) r) d dr (Suc n)" unfolding smult_monom smult_distribs mult_smult_left[symmetric] using lc0 by (simp add: Let_def ac_simps) finally show ?case . qed subsubsection \Division in polynomials over fields\ lemma pseudo_divmod_field: fixes g :: "'a::field poly" assumes g: "g \ 0" and *: "pseudo_divmod f g = (q,r)" defines "c \ coeff g (degree g) ^ (Suc (degree f) - degree g)" shows "f = g * smult (1/c) q + smult (1/c) r" proof - from leading_coeff_neq_0[OF g] have c0: "c \ 0" by (auto simp: c_def) from pseudo_divmod(1)[OF g *, folded c_def] have "smult c f = g * q + r" by auto also have "smult (1 / c) \ = g * smult (1 / c) q + smult (1 / c) r" by (simp add: smult_add_right) finally show ?thesis using c0 by auto qed lemma divide_poly_main_field: fixes d :: "'a::field poly" assumes d: "d \ 0" defines lc: "lc \ coeff d (degree d)" shows "divide_poly_main lc q r d dr n = fst (pseudo_divmod_main lc (smult ((1 / lc)^n) q) (smult ((1 / lc)^n) r) d dr n)" unfolding lc by (subst fst_pseudo_divmod_main_as_divide_poly_main) (auto simp: d power_one_over) lemma divide_poly_field: fixes f g :: "'a::field poly" defines "f' \ smult ((1 / coeff g (degree g)) ^ (Suc (degree f) - degree g)) f" shows "f div g = fst (pseudo_divmod f' g)" proof (cases "g = 0") case True show ?thesis unfolding divide_poly_def pseudo_divmod_def Let_def f'_def True by (simp add: divide_poly_main_0) next case False from leading_coeff_neq_0[OF False] have "degree f' = degree f" by (auto simp: f'_def) then show ?thesis using length_coeffs_degree[of f'] length_coeffs_degree[of f] unfolding divide_poly_def pseudo_divmod_def Let_def divide_poly_main_field[OF False] length_coeffs_degree[OF False] f'_def by force qed instantiation poly :: ("{semidom_divide_unit_factor,idom_divide}") normalization_semidom begin definition unit_factor_poly :: "'a poly \ 'a poly" where "unit_factor_poly p = [:unit_factor (lead_coeff p):]" definition normalize_poly :: "'a poly \ 'a poly" where "normalize p = p div [:unit_factor (lead_coeff p):]" instance proof fix p :: "'a poly" show "unit_factor p * normalize p = p" proof (cases "p = 0") case True then show ?thesis by (simp add: unit_factor_poly_def normalize_poly_def) next case False then have "lead_coeff p \ 0" by simp then have *: "unit_factor (lead_coeff p) \ 0" using unit_factor_is_unit [of "lead_coeff p"] by auto then have "unit_factor (lead_coeff p) dvd 1" by (auto intro: unit_factor_is_unit) then have **: "unit_factor (lead_coeff p) dvd c" for c by (rule dvd_trans) simp have ***: "unit_factor (lead_coeff p) * (c div unit_factor (lead_coeff p)) = c" for c proof - from ** obtain b where "c = unit_factor (lead_coeff p) * b" .. with False * show ?thesis by simp qed have "p div [:unit_factor (lead_coeff p):] = map_poly (\c. c div unit_factor (lead_coeff p)) p" by (simp add: const_poly_dvd_iff div_const_poly_conv_map_poly **) then show ?thesis by (simp add: normalize_poly_def unit_factor_poly_def smult_conv_map_poly map_poly_map_poly o_def ***) qed next fix p :: "'a poly" assume "is_unit p" then obtain c where p: "p = [:c:]" "c dvd 1" by (auto simp: is_unit_poly_iff) then show "unit_factor p = p" by (simp add: unit_factor_poly_def monom_0 is_unit_unit_factor) next fix p :: "'a poly" assume "p \ 0" then show "is_unit (unit_factor p)" by (simp add: unit_factor_poly_def monom_0 is_unit_poly_iff unit_factor_is_unit) next fix a b :: "'a poly" assume "is_unit a" thus "unit_factor (a * b) = a * unit_factor b" by (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult elim!: is_unit_polyE) qed (simp_all add: normalize_poly_def unit_factor_poly_def monom_0 lead_coeff_mult unit_factor_mult) end instance poly :: ("{semidom_divide_unit_factor,idom_divide,normalization_semidom_multiplicative}") normalization_semidom_multiplicative by intro_classes (auto simp: unit_factor_poly_def lead_coeff_mult unit_factor_mult) lemma normalize_poly_eq_map_poly: "normalize p = map_poly (\x. x div unit_factor (lead_coeff p)) p" proof - have "[:unit_factor (lead_coeff p):] dvd p" by (metis unit_factor_poly_def unit_factor_self) then show ?thesis by (simp add: normalize_poly_def div_const_poly_conv_map_poly) qed lemma coeff_normalize [simp]: "coeff (normalize p) n = coeff p n div unit_factor (lead_coeff p)" by (simp add: normalize_poly_eq_map_poly coeff_map_poly) class field_unit_factor = field + unit_factor + assumes unit_factor_field [simp]: "unit_factor = id" begin subclass semidom_divide_unit_factor proof fix a assume "a \ 0" then have "1 = a * inverse a" by simp then have "a dvd 1" .. then show "unit_factor a dvd 1" by simp qed simp_all end lemma unit_factor_pCons: "unit_factor (pCons a p) = (if p = 0 then [:unit_factor a:] else unit_factor p)" by (simp add: unit_factor_poly_def) lemma normalize_monom [simp]: "normalize (monom a n) = monom (normalize a) n" by (cases "a = 0") (simp_all add: map_poly_monom normalize_poly_eq_map_poly degree_monom_eq) lemma unit_factor_monom [simp]: "unit_factor (monom a n) = [:unit_factor a:]" by (cases "a = 0") (simp_all add: unit_factor_poly_def degree_monom_eq) lemma normalize_const_poly: "normalize [:c:] = [:normalize c:]" by (simp add: normalize_poly_eq_map_poly map_poly_pCons) lemma normalize_smult: fixes c :: "'a :: {normalization_semidom_multiplicative, idom_divide}" shows "normalize (smult c p) = smult (normalize c) (normalize p)" proof - have "smult c p = [:c:] * p" by simp also have "normalize \ = smult (normalize c) (normalize p)" by (subst normalize_mult) (simp add: normalize_const_poly) finally show ?thesis . qed inductive eucl_rel_poly :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly \ bool" where eucl_rel_poly_by0: "eucl_rel_poly x 0 (0, x)" | eucl_rel_poly_dividesI: "y \ 0 \ x = q * y \ eucl_rel_poly x y (q, 0)" | eucl_rel_poly_remainderI: "y \ 0 \ degree r < degree y \ x = q * y + r \ eucl_rel_poly x y (q, r)" lemma eucl_rel_poly_iff: "eucl_rel_poly x y (q, r) \ x = q * y + r \ (if y = 0 then q = 0 else r = 0 \ degree r < degree y)" by (auto elim: eucl_rel_poly.cases intro: eucl_rel_poly_by0 eucl_rel_poly_dividesI eucl_rel_poly_remainderI) lemma eucl_rel_poly_0: "eucl_rel_poly 0 y (0, 0)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_by_0: "eucl_rel_poly x 0 (0, x)" by (simp add: eucl_rel_poly_iff) lemma eucl_rel_poly_pCons: assumes rel: "eucl_rel_poly x y (q, r)" assumes y: "y \ 0" assumes b: "b = coeff (pCons a r) (degree y) / coeff y (degree y)" shows "eucl_rel_poly (pCons a x) y (pCons b q, pCons a r - smult b y)" (is "eucl_rel_poly ?x y (?q, ?r)") proof - from assms have x: "x = q * y + r" and r: "r = 0 \ degree r < degree y" by (simp_all add: eucl_rel_poly_iff) from b x have "?x = ?q * y + ?r" by simp moreover have "?r = 0 \ degree ?r < degree y" proof (rule eq_zero_or_degree_less) show "degree ?r \ degree y" proof (rule degree_diff_le) from r show "degree (pCons a r) \ degree y" by auto show "degree (smult b y) \ degree y" by (rule degree_smult_le) qed from \y \ 0\ show "coeff ?r (degree y) = 0" by (simp add: b) qed ultimately show ?thesis unfolding eucl_rel_poly_iff using \y \ 0\ by simp qed lemma eucl_rel_poly_exists: "\q r. eucl_rel_poly x y (q, r)" proof (cases "y = 0") case False show ?thesis proof (induction x) case 0 then show ?case using eucl_rel_poly_0 by blast next case (pCons a x) then show ?case using False eucl_rel_poly_pCons by blast qed qed (use eucl_rel_poly_by0 in blast) lemma eucl_rel_poly_unique: assumes 1: "eucl_rel_poly x y (q1, r1)" assumes 2: "eucl_rel_poly x y (q2, r2)" shows "q1 = q2 \ r1 = r2" proof (cases "y = 0") assume "y = 0" with assms show ?thesis by (simp add: eucl_rel_poly_iff) next assume [simp]: "y \ 0" from 1 have q1: "x = q1 * y + r1" and r1: "r1 = 0 \ degree r1 < degree y" unfolding eucl_rel_poly_iff by simp_all from 2 have q2: "x = q2 * y + r2" and r2: "r2 = 0 \ degree r2 < degree y" unfolding eucl_rel_poly_iff by simp_all from q1 q2 have q3: "(q1 - q2) * y = r2 - r1" by (simp add: algebra_simps) from r1 r2 have r3: "(r2 - r1) = 0 \ degree (r2 - r1) < degree y" by (auto intro: degree_diff_less) show "q1 = q2 \ r1 = r2" proof (rule classical) assume "\ ?thesis" with q3 have "q1 \ q2" and "r1 \ r2" by auto with r3 have "degree (r2 - r1) < degree y" by simp also have "degree y \ degree (q1 - q2) + degree y" by simp also from \q1 \ q2\ have "\ = degree ((q1 - q2) * y)" by (simp add: degree_mult_eq) also from q3 have "\ = degree (r2 - r1)" by simp finally have "degree (r2 - r1) < degree (r2 - r1)" . then show ?thesis by simp qed qed lemma eucl_rel_poly_0_iff: "eucl_rel_poly 0 y (q, r) \ q = 0 \ r = 0" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_0) lemma eucl_rel_poly_by_0_iff: "eucl_rel_poly x 0 (q, r) \ q = 0 \ r = x" by (auto dest: eucl_rel_poly_unique intro: eucl_rel_poly_by_0) lemmas eucl_rel_poly_unique_div = eucl_rel_poly_unique [THEN conjunct1] lemmas eucl_rel_poly_unique_mod = eucl_rel_poly_unique [THEN conjunct2] instantiation poly :: (field) semidom_modulo begin definition modulo_poly :: "'a poly \ 'a poly \ 'a poly" where mod_poly_def: "f mod g = (if g = 0 then f else pseudo_mod (smult ((1 / lead_coeff g) ^ (Suc (degree f) - degree g)) f) g)" instance proof fix x y :: "'a poly" show "x div y * y + x mod y = x" proof (cases "y = 0") case True then show ?thesis by (simp add: divide_poly_0 mod_poly_def) next case False then have "pseudo_divmod (smult ((1 / lead_coeff y) ^ (Suc (degree x) - degree y)) x) y = (x div y, x mod y)" by (simp add: divide_poly_field mod_poly_def pseudo_mod_def) with False pseudo_divmod [OF False this] show ?thesis by (simp add: power_mult_distrib [symmetric] ac_simps) qed qed end lemma eucl_rel_poly: "eucl_rel_poly x y (x div y, x mod y)" unfolding eucl_rel_poly_iff proof show "x = x div y * y + x mod y" by (simp add: div_mult_mod_eq) show "if y = 0 then x div y = 0 else x mod y = 0 \ degree (x mod y) < degree y" proof (cases "y = 0") case True then show ?thesis by auto next case False with pseudo_mod[OF this] show ?thesis by (simp add: mod_poly_def) qed qed lemma div_poly_eq: "eucl_rel_poly x y (q, r) \ x div y = q" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_div [OF eucl_rel_poly]) lemma mod_poly_eq: "eucl_rel_poly x y (q, r) \ x mod y = r" for x :: "'a::field poly" by (rule eucl_rel_poly_unique_mod [OF eucl_rel_poly]) instance poly :: (field) idom_modulo .. lemma div_pCons_eq: "pCons a p div q = (if q = 0 then 0 else pCons (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) (p div q))" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: div_poly_eq) lemma mod_pCons_eq: "pCons a p mod q = (if q = 0 then pCons a p else pCons a (p mod q) - smult (coeff (pCons a (p mod q)) (degree q) / lead_coeff q) q)" using eucl_rel_poly_pCons [OF eucl_rel_poly _ refl, of q a p] by (auto intro: mod_poly_eq) lemma div_mod_fold_coeffs: "(p div q, p mod q) = (if q = 0 then (0, p) else fold_coeffs (\a (s, r). let b = coeff (pCons a r) (degree q) / coeff q (degree q) in (pCons b s, pCons a r - smult b q)) p (0, 0))" by (rule sym, induct p) (auto simp: div_pCons_eq mod_pCons_eq Let_def) lemma degree_mod_less: "y \ 0 \ x mod y = 0 \ degree (x mod y) < degree y" using eucl_rel_poly [of x y] unfolding eucl_rel_poly_iff by simp lemma degree_mod_less': "b \ 0 \ a mod b \ 0 \ degree (a mod b) < degree b" using degree_mod_less[of b a] by auto lemma div_poly_less: fixes x :: "'a::field poly" assumes "degree x < degree y" shows "x div y = 0" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x div y = 0" by (rule div_poly_eq) qed lemma mod_poly_less: assumes "degree x < degree y" shows "x mod y = x" proof - from assms have "eucl_rel_poly x y (0, x)" by (simp add: eucl_rel_poly_iff) then show "x mod y = x" by (rule mod_poly_eq) qed lemma eucl_rel_poly_smult_left: "eucl_rel_poly x y (q, r) \ eucl_rel_poly (smult a x) y (smult a q, smult a r)" by (simp add: eucl_rel_poly_iff smult_add_right) lemma div_smult_left: "(smult a x) div y = smult a (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma mod_smult_left: "(smult a x) mod y = smult a (x mod y)" by (rule mod_poly_eq, rule eucl_rel_poly_smult_left, rule eucl_rel_poly) lemma poly_div_minus_left [simp]: "(- x) div y = - (x div y)" for x y :: "'a::field poly" using div_smult_left [of "- 1::'a"] by simp lemma poly_mod_minus_left [simp]: "(- x) mod y = - (x mod y)" for x y :: "'a::field poly" using mod_smult_left [of "- 1::'a"] by simp lemma eucl_rel_poly_add_left: assumes "eucl_rel_poly x y (q, r)" assumes "eucl_rel_poly x' y (q', r')" shows "eucl_rel_poly (x + x') y (q + q', r + r')" using assms unfolding eucl_rel_poly_iff by (auto simp: algebra_simps degree_add_less) lemma poly_div_add_left: "(x + y) div z = x div z + y div z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule div_poly_eq) lemma poly_mod_add_left: "(x + y) mod z = x mod z + y mod z" for x y z :: "'a::field poly" using eucl_rel_poly_add_left [OF eucl_rel_poly eucl_rel_poly] by (rule mod_poly_eq) lemma poly_div_diff_left: "(x - y) div z = x div z - y div z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_div_add_left poly_div_minus_left) lemma poly_mod_diff_left: "(x - y) mod z = x mod z - y mod z" for x y z :: "'a::field poly" by (simp only: diff_conv_add_uminus poly_mod_add_left poly_mod_minus_left) lemma eucl_rel_poly_smult_right: "a \ 0 \ eucl_rel_poly x y (q, r) \ eucl_rel_poly x (smult a y) (smult (inverse a) q, r)" by (simp add: eucl_rel_poly_iff) lemma div_smult_right: "a \ 0 \ x div (smult a y) = smult (inverse a) (x div y)" for x y :: "'a::field poly" by (rule div_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma mod_smult_right: "a \ 0 \ x mod (smult a y) = x mod y" by (rule mod_poly_eq, erule eucl_rel_poly_smult_right, rule eucl_rel_poly) lemma poly_div_minus_right [simp]: "x div (- y) = - (x div y)" for x y :: "'a::field poly" using div_smult_right [of "- 1::'a"] by (simp add: nonzero_inverse_minus_eq) lemma poly_mod_minus_right [simp]: "x mod (- y) = x mod y" for x y :: "'a::field poly" using mod_smult_right [of "- 1::'a"] by simp lemma eucl_rel_poly_mult: assumes "eucl_rel_poly x y (q, r)" "eucl_rel_poly q z (q', r')" shows "eucl_rel_poly x (y * z) (q', y * r' + r)" proof (cases "y = 0") case True with assms eucl_rel_poly_0_iff show ?thesis by (force simp add: eucl_rel_poly_iff) next case False show ?thesis proof (cases "r' = 0") case True with assms show ?thesis by (auto simp add: eucl_rel_poly_iff degree_mult_eq) next case False with assms \y \ 0\ show ?thesis by (auto simp add: eucl_rel_poly_iff degree_add_less degree_mult_eq field_simps) qed qed lemma poly_div_mult_right: "x div (y * z) = (x div y) div z" for x y z :: "'a::field poly" by (rule div_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma poly_mod_mult_right: "x mod (y * z) = y * (x div y mod z) + x mod y" for x y z :: "'a::field poly" by (rule mod_poly_eq, rule eucl_rel_poly_mult, (rule eucl_rel_poly)+) lemma mod_pCons: fixes a :: "'a::field" and x y :: "'a::field poly" assumes y: "y \ 0" defines "b \ coeff (pCons a (x mod y)) (degree y) / coeff y (degree y)" shows "(pCons a x) mod y = pCons a (x mod y) - smult b y" unfolding b_def by (rule mod_poly_eq, rule eucl_rel_poly_pCons [OF eucl_rel_poly y refl]) subsubsection \List-based versions for fast implementation\ (* Subsection by: Sebastiaan Joosten René Thiemann Akihisa Yamada *) fun minus_poly_rev_list :: "'a :: group_add list \ 'a list \ 'a list" where "minus_poly_rev_list (x # xs) (y # ys) = (x - y) # (minus_poly_rev_list xs ys)" | "minus_poly_rev_list xs [] = xs" | "minus_poly_rev_list [] (y # ys) = []" fun pseudo_divmod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "pseudo_divmod_main_list lc q r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; qqq = cCons a (map ((*) lc) q); rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_divmod_main_list lc qqq rrr d n)" | "pseudo_divmod_main_list lc q r d 0 = (q, r)" fun pseudo_mod_main_list :: "'a::comm_ring_1 \ 'a list \ 'a list \ nat \ 'a list" where "pseudo_mod_main_list lc r d (Suc n) = (let rr = map ((*) lc) r; a = hd r; rrr = tl (if a = 0 then rr else minus_poly_rev_list rr (map ((*) a) d)) in pseudo_mod_main_list lc rrr d n)" | "pseudo_mod_main_list lc r d 0 = r" fun divmod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ nat \ 'a list \ 'a list" where "divmod_poly_one_main_list q r d (Suc n) = (let a = hd r; qqq = cCons a q; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in divmod_poly_one_main_list qqq rr d n)" | "divmod_poly_one_main_list q r d 0 = (q, r)" fun mod_poly_one_main_list :: "'a::comm_ring_1 list \ 'a list \ nat \ 'a list" where "mod_poly_one_main_list r d (Suc n) = (let a = hd r; rr = tl (if a = 0 then r else minus_poly_rev_list r (map ((*) a) d)) in mod_poly_one_main_list rr d n)" | "mod_poly_one_main_list r d 0 = r" definition pseudo_divmod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list \ 'a list" where "pseudo_divmod_list p q = (if q = [] then ([], p) else (let rq = rev q; (qu,re) = pseudo_divmod_main_list (hd rq) [] (rev p) rq (1 + length p - length q) in (qu, rev re)))" definition pseudo_mod_list :: "'a::comm_ring_1 list \ 'a list \ 'a list" where "pseudo_mod_list p q = (if q = [] then p else (let rq = rev q; re = pseudo_mod_main_list (hd rq) (rev p) rq (1 + length p - length q) in rev re))" lemma minus_zero_does_nothing: "minus_poly_rev_list x (map ((*) 0) y) = x" for x :: "'a::ring list" by (induct x y rule: minus_poly_rev_list.induct) auto lemma length_minus_poly_rev_list [simp]: "length (minus_poly_rev_list xs ys) = length xs" by (induct xs ys rule: minus_poly_rev_list.induct) auto lemma if_0_minus_poly_rev_list: "(if a = 0 then x else minus_poly_rev_list x (map ((*) a) y)) = minus_poly_rev_list x (map ((*) a) y)" for a :: "'a::ring" by(cases "a = 0") (simp_all add: minus_zero_does_nothing) lemma Poly_append: "Poly (a @ b) = Poly a + monom 1 (length a) * Poly b" for a :: "'a::comm_semiring_1 list" by (induct a) (auto simp: monom_0 monom_Suc) lemma minus_poly_rev_list: "length p \ length q \ Poly (rev (minus_poly_rev_list (rev p) (rev q))) = Poly p - monom 1 (length p - length q) * Poly q" for p q :: "'a :: comm_ring_1 list" proof (induct "rev p" "rev q" arbitrary: p q rule: minus_poly_rev_list.induct) case (1 x xs y ys) then have "length (rev q) \ length (rev p)" by simp from this[folded 1(2,3)] have ys_xs: "length ys \ length xs" by simp then have *: "Poly (rev (minus_poly_rev_list xs ys)) = Poly (rev xs) - monom 1 (length xs - length ys) * Poly (rev ys)" by (subst "1.hyps"(1)[of "rev xs" "rev ys", unfolded rev_rev_ident length_rev]) auto have "Poly p - monom 1 (length p - length q) * Poly q = Poly (rev (rev p)) - monom 1 (length (rev (rev p)) - length (rev (rev q))) * Poly (rev (rev q))" by simp also have "\ = Poly (rev (x # xs)) - monom 1 (length (x # xs) - length (y # ys)) * Poly (rev (y # ys))" unfolding 1(2,3) by simp also from ys_xs have "\ = Poly (rev xs) + monom x (length xs) - (monom 1 (length xs - length ys) * Poly (rev ys) + monom y (length xs))" by (simp add: Poly_append distrib_left mult_monom smult_monom) also have "\ = Poly (rev (minus_poly_rev_list xs ys)) + monom (x - y) (length xs)" unfolding * diff_monom[symmetric] by simp finally show ?case by (simp add: 1(2,3)[symmetric] smult_monom Poly_append) qed auto lemma smult_monom_mult: "smult a (monom b n * f) = monom (a * b) n * f" using smult_monom [of a _ n] by (metis mult_smult_left) lemma head_minus_poly_rev_list: "length d \ length r \ d \ [] \ hd (minus_poly_rev_list (map ((*) (last d)) r) (map ((*) (hd r)) (rev d))) = 0" for d r :: "'a::comm_ring list" proof (induct r) case Nil then show ?case by simp next case (Cons a rs) then show ?case by (cases "rev d") (simp_all add: ac_simps) qed lemma Poly_map: "Poly (map ((*) a) p) = smult a (Poly p)" proof (induct p) case Nil then show ?case by simp next case (Cons x xs) then show ?case by (cases "Poly xs = 0") auto qed lemma last_coeff_is_hd: "xs \ [] \ coeff (Poly xs) (length xs - 1) = hd (rev xs)" by (simp_all add: hd_conv_nth rev_nth nth_default_nth nth_append) lemma pseudo_divmod_main_list_invar: assumes leading_nonzero: "last d \ 0" and lc: "last d = lc" and "d \ []" and "pseudo_divmod_main_list lc q (rev r) (rev d) n = (q', rev r')" and "n = 1 + length r - length d" shows "pseudo_divmod_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n = (Poly q', Poly r')" using assms(4-) proof (induct n arbitrary: r q) case (Suc n) from Suc.prems have *: "\ Suc (length r) \ length d" by simp with \d \ []\ have "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce let ?a = "(hd (rev r))" let ?rr = "map ((*) lc) (rev r)" let ?rrr = "rev (tl (minus_poly_rev_list ?rr (map ((*) ?a) (rev d))))" let ?qq = "cCons ?a (map ((*) lc) q)" from * Suc(3) have n: "n = (1 + length r - length d - 1)" by simp from * have rr_val:"(length ?rrr) = (length r - 1)" by auto with \r \ []\ * have rr_smaller: "(1 + length r - length d - 1) = (1 + length ?rrr - length d)" by auto from * have id: "Suc (length r) - length d = Suc (length r - length d)" by auto from Suc.prems * have "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) (1 + length r - length d - 1) = (q', rev r')" by (simp add: Let_def if_0_minus_poly_rev_list id) with n have v: "pseudo_divmod_main_list lc ?qq (rev ?rrr) (rev d) n = (q', rev r')" by auto from * have sucrr:"Suc (length r) - length d = Suc (length r - length d)" using Suc_diff_le not_less_eq_eq by blast from Suc(3) \r \ []\ have n_ok : "n = 1 + (length ?rrr) - length d" by simp have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ pseudo_divmod_main lc x1 x2 x3 x4 n = pseudo_divmod_main lc y1 y2 y3 y4 n" by simp have hd_rev: "coeff (Poly r) (length r - Suc 0) = hd (rev r)" using last_coeff_is_hd[OF \r \ []\] by simp show ?case unfolding Suc.hyps(1)[OF v n_ok, symmetric] pseudo_divmod_main.simps Let_def proof (rule cong[OF _ _ refl], goal_cases) case 1 show ?case by (simp add: monom_Suc hd_rev[symmetric] smult_monom Poly_map) next case 2 show ?case proof (subst Poly_on_rev_starting_with_0, goal_cases) show "hd (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))) = 0" by (fold lc, subst head_minus_poly_rev_list, insert * \d \ []\, auto) from * have "length d \ length r" by simp then show "smult lc (Poly r) - monom (coeff (Poly r) (length r - 1)) n * Poly d = Poly (rev (minus_poly_rev_list (map ((*) lc) (rev r)) (map ((*) (hd (rev r))) (rev d))))" by (fold rev_map) (auto simp add: n smult_monom_mult Poly_map hd_rev [symmetric] minus_poly_rev_list) qed qed simp qed simp lemma pseudo_divmod_impl [code]: "pseudo_divmod f g = map_prod poly_of_list poly_of_list (pseudo_divmod_list (coeffs f) (coeffs g))" for f g :: "'a::comm_ring_1 poly" proof (cases "g = 0") case False then have "last (coeffs g) \ 0" and "last (coeffs g) = lead_coeff g" and "coeffs g \ []" by (simp_all add: last_coeffs_eq_coeff_degree) moreover obtain q r where qr: "pseudo_divmod_main_list (last (coeffs g)) (rev []) (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, rev (rev r))" by force ultimately have "(Poly q, Poly (rev r)) = pseudo_divmod_main (lead_coeff g) 0 f g (length (coeffs f) - Suc 0) (Suc (length (coeffs f)) - length (coeffs g))" by (subst pseudo_divmod_main_list_invar [symmetric]) auto moreover have "pseudo_divmod_main_list (hd (rev (coeffs g))) [] (rev (coeffs f)) (rev (coeffs g)) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by (metis hd_rev qr rev.simps(1) rev_swap) ultimately show ?thesis by (simp add: degree_eq_length_coeffs pseudo_divmod_def pseudo_divmod_list_def) next case True then show ?thesis by (auto simp add: pseudo_divmod_def pseudo_divmod_list_def) qed lemma pseudo_mod_main_list: "snd (pseudo_divmod_main_list l q xs ys n) = pseudo_mod_main_list l xs ys n" by (induct n arbitrary: l q xs ys) (auto simp: Let_def) lemma pseudo_mod_impl[code]: "pseudo_mod f g = poly_of_list (pseudo_mod_list (coeffs f) (coeffs g))" proof - have snd_case: "\f g p. snd ((\(x,y). (f x, g y)) p) = g (snd p)" by auto show ?thesis unfolding pseudo_mod_def pseudo_divmod_impl pseudo_divmod_list_def pseudo_mod_list_def Let_def by (simp add: snd_case pseudo_mod_main_list) qed subsubsection \Improved Code-Equations for Polynomial (Pseudo) Division\ lemma pdivmod_pdivmodrel: "eucl_rel_poly p q (r, s) \ (p div q, p mod q) = (r, s)" by (metis eucl_rel_poly eucl_rel_poly_unique) lemma pdivmod_via_pseudo_divmod: "(f div g, f mod g) = (if g = 0 then (0, f) else let ilc = inverse (coeff g (degree g)); h = smult ilc g; (q,r) = pseudo_divmod f h in (smult ilc q, r))" (is "?l = ?r") proof (cases "g = 0") case True then show ?thesis by simp next case False define lc where "lc = inverse (coeff g (degree g))" define h where "h = smult lc g" from False have h1: "coeff h (degree h) = 1" and lc: "lc \ 0" by (auto simp: h_def lc_def) then have h0: "h \ 0" by auto obtain q r where p: "pseudo_divmod f h = (q, r)" by force from False have id: "?r = (smult lc q, r)" by (auto simp: Let_def h_def[symmetric] lc_def[symmetric] p) from pseudo_divmod[OF h0 p, unfolded h1] have f: "f = h * q + r" and r: "r = 0 \ degree r < degree h" by auto from f r h0 have "eucl_rel_poly f h (q, r)" by (auto simp: eucl_rel_poly_iff) then have "(f div h, f mod h) = (q, r)" by (simp add: pdivmod_pdivmodrel) with lc have "(f div g, f mod g) = (smult lc q, r)" by (auto simp: h_def div_smult_right[OF lc] mod_smult_right[OF lc]) with id show ?thesis by auto qed lemma pdivmod_via_pseudo_divmod_list: "(f div g, f mod g) = (let cg = coeffs g in if cg = [] then (0, f) else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; (q, r) = pseudo_divmod_main_list 1 [] (rev cf) (rev ch) (1 + length cf - length cg) in (poly_of_list (map ((*) ilc) q), poly_of_list (rev r)))" proof - note d = pdivmod_via_pseudo_divmod pseudo_divmod_impl pseudo_divmod_list_def show ?thesis proof (cases "g = 0") case True with d show ?thesis by auto next case False define ilc where "ilc = inverse (coeff g (degree g))" from False have ilc: "ilc \ 0" by (auto simp: ilc_def) with False have id: "g = 0 \ False" "coeffs g = [] \ False" "last (coeffs g) = coeff g (degree g)" "coeffs (smult ilc g) = [] \ False" by (auto simp: last_coeffs_eq_coeff_degree) have id2: "hd (rev (coeffs (smult ilc g))) = 1" by (subst hd_rev, insert id ilc, auto simp: coeffs_smult, subst last_map, auto simp: id ilc_def) have id3: "length (coeffs (smult ilc g)) = length (coeffs g)" "rev (coeffs (smult ilc g)) = rev (map ((*) ilc) (coeffs g))" unfolding coeffs_smult using ilc by auto obtain q r where pair: "pseudo_divmod_main_list 1 [] (rev (coeffs f)) (rev (map ((*) ilc) (coeffs g))) (1 + length (coeffs f) - length (coeffs g)) = (q, r)" by force show ?thesis unfolding d Let_def id if_False ilc_def[symmetric] map_prod_def[symmetric] id2 unfolding id3 pair map_prod_def split by (auto simp: Poly_map) qed qed lemma pseudo_divmod_main_list_1: "pseudo_divmod_main_list 1 = divmod_poly_one_main_list" proof (intro ext, goal_cases) case (1 q r d n) have *: "map ((*) 1) xs = xs" for xs :: "'a list" by (induct xs) auto show ?case by (induct n arbitrary: q r d) (auto simp: * Let_def) qed fun divide_poly_main_list :: "'a::idom_divide \ 'a list \ 'a list \ 'a list \ nat \ 'a list" where "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r in if cr = 0 then divide_poly_main_list lc (cCons cr q) (tl r) d n else let a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" | "divide_poly_main_list lc q r d 0 = q" lemma divide_poly_main_list_simp [simp]: "divide_poly_main_list lc q r d (Suc n) = (let cr = hd r; a = cr div lc; qq = cCons a q; rr = minus_poly_rev_list r (map ((*) a) d) in if hd rr = 0 then divide_poly_main_list lc qq (tl rr) d n else [])" by (simp add: Let_def minus_zero_does_nothing) declare divide_poly_main_list.simps(1)[simp del] definition divide_poly_list :: "'a::idom_divide poly \ 'a poly \ 'a poly" where "divide_poly_list f g = (let cg = coeffs g in if cg = [] then g else let cf = coeffs f; cgr = rev cg in poly_of_list (divide_poly_main_list (hd cgr) [] (rev cf) cgr (1 + length cf - length cg)))" lemmas pdivmod_via_divmod_list = pdivmod_via_pseudo_divmod_list[unfolded pseudo_divmod_main_list_1] lemma mod_poly_one_main_list: "snd (divmod_poly_one_main_list q r d n) = mod_poly_one_main_list r d n" by (induct n arbitrary: q r d) (auto simp: Let_def) lemma mod_poly_code [code]: "f mod g = (let cg = coeffs g in if cg = [] then f else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; r = mod_poly_one_main_list (rev cf) (rev ch) (1 + length cf - length cg) in poly_of_list (rev r))" (is "_ = ?rhs") proof - have "snd (f div g, f mod g) = ?rhs" unfolding pdivmod_via_divmod_list Let_def mod_poly_one_main_list [symmetric, of _ _ _ Nil] by (auto split: prod.splits) then show ?thesis by simp qed definition div_field_poly_impl :: "'a :: field poly \ 'a poly \ 'a poly" where "div_field_poly_impl f g = (let cg = coeffs g in if cg = [] then 0 else let cf = coeffs f; ilc = inverse (last cg); ch = map ((*) ilc) cg; q = fst (divmod_poly_one_main_list [] (rev cf) (rev ch) (1 + length cf - length cg)) in poly_of_list ((map ((*) ilc) q)))" text \We do not declare the following lemma as code equation, since then polynomial division on non-fields will no longer be executable. However, a code-unfold is possible, since \div_field_poly_impl\ is a bit more efficient than the generic polynomial division.\ lemma div_field_poly_impl[code_unfold]: "(div) = div_field_poly_impl" proof (intro ext) fix f g :: "'a poly" have "fst (f div g, f mod g) = div_field_poly_impl f g" unfolding div_field_poly_impl_def pdivmod_via_divmod_list Let_def by (auto split: prod.splits) then show "f div g = div_field_poly_impl f g" by simp qed lemma divide_poly_main_list: assumes lc0: "lc \ 0" and lc: "last d = lc" and d: "d \ []" and "n = (1 + length r - length d)" shows "Poly (divide_poly_main_list lc q (rev r) (rev d) n) = divide_poly_main lc (monom 1 n * Poly q) (Poly r) (Poly d) (length r - 1) n" using assms(4-) proof (induct "n" arbitrary: r q) case (Suc n) from Suc.prems have ifCond: "\ Suc (length r) \ length d" by simp with d have r: "r \ []" using Suc_leI length_greater_0_conv list.size(3) by fastforce then obtain rr lcr where r: "r = rr @ [lcr]" by (cases r rule: rev_cases) auto from d lc obtain dd where d: "d = dd @ [lc]" by (cases d rule: rev_cases) auto from Suc(2) ifCond have n: "n = 1 + length rr - length d" by (auto simp: r) from ifCond have len: "length dd \ length rr" by (simp add: r d) show ?case proof (cases "lcr div lc * lc = lcr") case False with r d show ?thesis unfolding Suc(2)[symmetric] by (auto simp add: Let_def nth_default_append) next case True with r d have id: "?thesis \ Poly (divide_poly_main_list lc (cCons (lcr div lc) q) (rev (rev (minus_poly_rev_list (rev rr) (rev (map ((*) (lcr div lc)) dd))))) (rev d) n) = divide_poly_main lc (monom 1 (Suc n) * Poly q + monom (lcr div lc) n) (Poly r - monom (lcr div lc) n * Poly d) (Poly d) (length rr - 1) n" by (cases r rule: rev_cases; cases "d" rule: rev_cases) (auto simp add: Let_def rev_map nth_default_append) have cong: "\x1 x2 x3 x4 y1 y2 y3 y4. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ divide_poly_main lc x1 x2 x3 x4 n = divide_poly_main lc y1 y2 y3 y4 n" by simp show ?thesis unfolding id proof (subst Suc(1), simp add: n, subst minus_poly_rev_list, force simp: len, rule cong[OF _ _ refl], goal_cases) case 2 have "monom lcr (length rr) = monom (lcr div lc) (length rr - length dd) * monom lc (length dd)" by (simp add: mult_monom len True) then show ?case unfolding r d Poly_append n ring_distribs by (auto simp: Poly_map smult_monom smult_monom_mult) qed (auto simp: len monom_Suc smult_monom) qed qed simp lemma divide_poly_list[code]: "f div g = divide_poly_list f g" proof - note d = divide_poly_def divide_poly_list_def show ?thesis proof (cases "g = 0") case True show ?thesis by (auto simp: d True) next case False then obtain cg lcg where cg: "coeffs g = cg @ [lcg]" by (cases "coeffs g" rule: rev_cases) auto with False have id: "(g = 0) = False" "(cg @ [lcg] = []) = False" by auto from cg False have lcg: "coeff g (degree g) = lcg" using last_coeffs_eq_coeff_degree last_snoc by force with False have "lcg \ 0" by auto from cg Poly_coeffs [of g] have ltp: "Poly (cg @ [lcg]) = g" by auto show ?thesis unfolding d cg Let_def id if_False poly_of_list_def by (subst divide_poly_main_list, insert False cg \lcg \ 0\) (auto simp: lcg ltp, simp add: degree_eq_length_coeffs) qed qed subsection \Primality and irreducibility in polynomial rings\ lemma prod_mset_const_poly: "(\x\#A. [:f x:]) = [:prod_mset (image_mset f A):]" by (induct A) (simp_all add: ac_simps) lemma irreducible_const_poly_iff: fixes c :: "'a :: {comm_semiring_1,semiring_no_zero_divisors}" shows "irreducible [:c:] \ irreducible c" proof assume A: "irreducible c" show "irreducible [:c:]" proof (rule irreducibleI) fix a b assume ab: "[:c:] = a * b" hence "degree [:c:] = degree (a * b)" by (simp only: ) also from A ab have "a \ 0" "b \ 0" by auto hence "degree (a * b) = degree a + degree b" by (simp add: degree_mult_eq) finally have "degree a = 0" "degree b = 0" by auto then obtain a' b' where ab': "a = [:a':]" "b = [:b':]" by (auto elim!: degree_eq_zeroE) from ab have "coeff [:c:] 0 = coeff (a * b) 0" by (simp only: ) hence "c = a' * b'" by (simp add: ab' mult_ac) from A and this have "a' dvd 1 \ b' dvd 1" by (rule irreducibleD) with ab' show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed (insert A, auto simp: irreducible_def is_unit_poly_iff) next assume A: "irreducible [:c:]" then have "c \ 0" and "\ c dvd 1" by (auto simp add: irreducible_def is_unit_const_poly_iff) then show "irreducible c" proof (rule irreducibleI) fix a b assume ab: "c = a * b" hence "[:c:] = [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:a:] dvd 1 \ [:b:] dvd 1" by (rule irreducibleD) then show "a dvd 1 \ b dvd 1" by (auto simp add: is_unit_const_poly_iff) qed qed lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i unfolding m_def by (metis (mono_tags, lifting) B Greatest_le_nat that) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) subsection \Content and primitive part of a polynomial\ definition content :: "'a::semiring_gcd poly \ 'a" where "content p = gcd_list (coeffs p)" lemma content_eq_fold_coeffs [code]: "content p = fold_coeffs gcd p 0" by (simp add: content_def Gcd_fin.set_eq_fold fold_coeffs_def foldr_fold fun_eq_iff ac_simps) lemma content_0 [simp]: "content 0 = 0" by (simp add: content_def) lemma content_1 [simp]: "content 1 = 1" by (simp add: content_def) lemma content_const [simp]: "content [:c:] = normalize c" by (simp add: content_def cCons_def) lemma const_poly_dvd_iff_dvd_content: "[:c:] dvd p \ c dvd content p" for c :: "'a::semiring_gcd" proof (cases "p = 0") case True then show ?thesis by simp next case False have "[:c:] dvd p \ (\n. c dvd coeff p n)" by (rule const_poly_dvd_iff) also have "\ \ (\a\set (coeffs p). c dvd a)" proof safe fix n :: nat assume "\a\set (coeffs p). c dvd a" then show "c dvd coeff p n" by (cases "n \ degree p") (auto simp: coeff_eq_0 coeffs_def split: if_splits) qed (auto simp: coeffs_def simp del: upt_Suc split: if_splits) also have "\ \ c dvd content p" by (simp add: content_def dvd_Gcd_fin_iff dvd_mult_unit_iff) finally show ?thesis . qed lemma content_dvd [simp]: "[:content p:] dvd p" by (subst const_poly_dvd_iff_dvd_content) simp_all lemma content_dvd_coeff [simp]: "content p dvd coeff p n" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis by (cases "n \ degree p") (auto simp add: content_def not_le coeff_eq_0 coeff_in_coeffs intro: Gcd_fin_dvd) qed lemma content_dvd_coeffs: "c \ set (coeffs p) \ content p dvd c" by (simp add: content_def Gcd_fin_dvd) lemma normalize_content [simp]: "normalize (content p) = content p" by (simp add: content_def) lemma is_unit_content_iff [simp]: "is_unit (content p) \ content p = 1" proof assume "is_unit (content p)" then have "normalize (content p) = 1" by (simp add: is_unit_normalize del: normalize_content) then show "content p = 1" by simp qed auto lemma content_smult [simp]: fixes c :: "'a :: {normalization_semidom_multiplicative, semiring_gcd}" shows "content (smult c p) = normalize c * content p" by (simp add: content_def coeffs_smult Gcd_fin_mult normalize_mult) lemma content_eq_zero_iff [simp]: "content p = 0 \ p = 0" by (auto simp: content_def simp: poly_eq_iff coeffs_def) definition primitive_part :: "'a :: semiring_gcd poly \ 'a poly" where "primitive_part p = map_poly (\x. x div content p) p" lemma primitive_part_0 [simp]: "primitive_part 0 = 0" by (simp add: primitive_part_def) lemma content_times_primitive_part [simp]: "smult (content p) (primitive_part p) = p" for p :: "'a :: semiring_gcd poly" proof (cases "p = 0") case True then show ?thesis by simp next case False then show ?thesis unfolding primitive_part_def by (auto simp: smult_conv_map_poly map_poly_map_poly o_def content_dvd_coeffs intro: map_poly_idI) qed lemma primitive_part_eq_0_iff [simp]: "primitive_part p = 0 \ p = 0" proof (cases "p = 0") case True then show ?thesis by simp next case False then have "primitive_part p = map_poly (\x. x div content p) p" by (simp add: primitive_part_def) also from False have "\ = 0 \ p = 0" by (intro map_poly_eq_0_iff) (auto simp: dvd_div_eq_0_iff content_dvd_coeffs) finally show ?thesis using False by simp qed lemma content_primitive_part [simp]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" assumes "p \ 0" shows "content (primitive_part p) = 1" proof - have "p = smult (content p) (primitive_part p)" by simp also have "content \ = content (primitive_part p) * content p" by (simp del: content_times_primitive_part add: ac_simps) finally have "1 * content p = content (primitive_part p) * content p" by simp then have "1 * content p div content p = content (primitive_part p) * content p div content p" by simp with assms show ?thesis by simp qed lemma content_decompose: obtains p' :: "'a :: {normalization_semidom_multiplicative, semiring_gcd} poly" where "p = smult (content p) p'" "content p' = 1" proof (cases "p = 0") case True then have "p = smult (content p) 1" "content 1 = 1" by simp_all then show ?thesis .. next case False then have "p = smult (content p) (primitive_part p)" "content (primitive_part p) = 1" by simp_all then show ?thesis .. qed lemma content_dvd_contentI [intro]: "p dvd q \ content p dvd content q" using const_poly_dvd_iff_dvd_content content_dvd dvd_trans by blast lemma primitive_part_const_poly [simp]: "primitive_part [:x:] = [:unit_factor x:]" by (simp add: primitive_part_def map_poly_pCons) lemma primitive_part_prim: "content p = 1 \ primitive_part p = p" by (auto simp: primitive_part_def) lemma degree_primitive_part [simp]: "degree (primitive_part p) = degree p" proof (cases "p = 0") case True then show ?thesis by simp next case False have "p = smult (content p) (primitive_part p)" by simp also from False have "degree \ = degree (primitive_part p)" by (subst degree_smult_eq) simp_all finally show ?thesis .. qed lemma smult_content_normalize_primitive_part [simp]: fixes p :: "'a :: {normalization_semidom_multiplicative, semiring_gcd, idom_divide} poly" shows "smult (content p) (normalize (primitive_part p)) = normalize p" proof - have "smult (content p) (normalize (primitive_part p)) = normalize ([:content p:] * primitive_part p)" by (subst normalize_mult) (simp_all add: normalize_const_poly) also have "[:content p:] * primitive_part p = p" by simp finally show ?thesis . qed context begin private lemma content_1_mult: fixes f g :: "'a :: {semiring_gcd, factorial_semiring} poly" assumes "content f = 1" "content g = 1" shows "content (f * g) = 1" proof (cases "f * g = 0") case False from assms have "f \ 0" "g \ 0" by auto hence "f * g \ 0" by auto { assume "\is_unit (content (f * g))" with False have "\p. p dvd content (f * g) \ prime p" by (intro prime_divisor_exists) simp_all then obtain p where "p dvd content (f * g)" "prime p" by blast from \p dvd content (f * g)\ have "[:p:] dvd f * g" by (simp add: const_poly_dvd_iff_dvd_content) moreover from \prime p\ have "prime_elem [:p:]" by (simp add: lift_prime_elem_poly) ultimately have "[:p:] dvd f \ [:p:] dvd g" by (simp add: prime_elem_dvd_mult_iff) with assms have "is_unit p" by (simp add: const_poly_dvd_iff_dvd_content) with \prime p\ have False by simp } hence "is_unit (content (f * g))" by blast hence "normalize (content (f * g)) = 1" by (simp add: is_unit_normalize del: normalize_content) thus ?thesis by simp qed (insert assms, auto) lemma content_mult: fixes p q :: "'a :: {factorial_semiring, semiring_gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = content p * content q" proof (cases "p * q = 0") case False then have "p \ 0" and "q \ 0" by simp_all then have *: "content (primitive_part p * primitive_part q) = 1" by (auto intro: content_1_mult) have "p * q = smult (content p) (primitive_part p) * smult (content q) (primitive_part q)" by simp also have "\ = smult (content p * content q) (primitive_part p * primitive_part q)" by (metis mult.commute mult_smult_right smult_smult) with * show ?thesis by (simp add: normalize_mult) next case True then show ?thesis by auto qed end lemma primitive_part_mult: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "primitive_part (p * q) = primitive_part p * primitive_part q" proof - have "primitive_part (p * q) = p * q div [:content (p * q):]" by (simp add: primitive_part_def div_const_poly_conv_map_poly) also have "\ = (p div [:content p:]) * (q div [:content q:])" by (subst div_mult_div_if_dvd) (simp_all add: content_mult mult_ac) also have "\ = primitive_part p * primitive_part q" by (simp add: primitive_part_def div_const_poly_conv_map_poly) finally show ?thesis . qed lemma primitive_part_smult: fixes p :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "primitive_part (smult a p) = smult (unit_factor a) (primitive_part p)" proof - have "smult a p = [:a:] * p" by simp also have "primitive_part \ = smult (unit_factor a) (primitive_part p)" by (subst primitive_part_mult) simp_all finally show ?thesis . qed lemma primitive_part_dvd_primitive_partI [intro]: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, ring_gcd, idom_divide, normalization_semidom_multiplicative} poly" shows "p dvd q \ primitive_part p dvd primitive_part q" by (auto elim!: dvdE simp: primitive_part_mult) lemma content_prod_mset: fixes A :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly multiset" shows "content (prod_mset A) = prod_mset (image_mset content A)" by (induction A) (simp_all add: content_mult mult_ac) lemma content_prod_eq_1_iff: fixes p q :: "'a :: {factorial_semiring, semiring_Gcd, normalization_semidom_multiplicative} poly" shows "content (p * q) = 1 \ content p = 1 \ content q = 1" proof safe assume A: "content (p * q) = 1" { fix p q :: "'a poly" assume "content p * content q = 1" hence "1 = content p * content q" by simp hence "content p dvd 1" by (rule dvdI) hence "content p = 1" by simp } note B = this from A B[of p q] B [of q p] show "content p = 1" "content q = 1" by (simp_all add: content_mult mult_ac) qed (auto simp: content_mult) no_notation cCons (infixr "##" 65) end diff --git a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy --- a/src/HOL/Computational_Algebra/Polynomial_Factorial.thy +++ b/src/HOL/Computational_Algebra/Polynomial_Factorial.thy @@ -1,869 +1,885 @@ (* Title: HOL/Computational_Algebra/Polynomial_Factorial.thy Author: Manuel Eberl *) section \Polynomials, fractions and rings\ theory Polynomial_Factorial imports Complex_Main Polynomial Normalized_Fraction begin subsection \Lifting elements into the field of fractions\ definition to_fract :: "'a :: idom \ 'a fract" where "to_fract x = Fract x 1" \ \FIXME: more idiomatic name, abbreviation\ lemma to_fract_0 [simp]: "to_fract 0 = 0" by (simp add: to_fract_def eq_fract Zero_fract_def) lemma to_fract_1 [simp]: "to_fract 1 = 1" by (simp add: to_fract_def eq_fract One_fract_def) lemma to_fract_add [simp]: "to_fract (x + y) = to_fract x + to_fract y" by (simp add: to_fract_def) lemma to_fract_diff [simp]: "to_fract (x - y) = to_fract x - to_fract y" by (simp add: to_fract_def) lemma to_fract_uminus [simp]: "to_fract (-x) = -to_fract x" by (simp add: to_fract_def) lemma to_fract_mult [simp]: "to_fract (x * y) = to_fract x * to_fract y" by (simp add: to_fract_def) lemma to_fract_eq_iff [simp]: "to_fract x = to_fract y \ x = y" by (simp add: to_fract_def eq_fract) lemma to_fract_eq_0_iff [simp]: "to_fract x = 0 \ x = 0" by (simp add: to_fract_def Zero_fract_def eq_fract) lemma snd_quot_of_fract_nonzero [simp]: "snd (quot_of_fract x) \ 0" by transfer simp lemma Fract_quot_of_fract [simp]: "Fract (fst (quot_of_fract x)) (snd (quot_of_fract x)) = x" by transfer (simp del: fractrel_iff, subst fractrel_normalize_quot_left, simp) lemma to_fract_quot_of_fract: assumes "snd (quot_of_fract x) = 1" shows "to_fract (fst (quot_of_fract x)) = x" proof - have "x = Fract (fst (quot_of_fract x)) (snd (quot_of_fract x))" by simp also note assms finally show ?thesis by (simp add: to_fract_def) qed lemma snd_quot_of_fract_Fract_whole: assumes "y dvd x" shows "snd (quot_of_fract (Fract x y)) = 1" using assms by transfer (auto simp: normalize_quot_def Let_def gcd_proj2_if_dvd) lemma Fract_conv_to_fract: "Fract a b = to_fract a / to_fract b" by (simp add: to_fract_def) lemma quot_of_fract_to_fract [simp]: "quot_of_fract (to_fract x) = (x, 1)" unfolding to_fract_def by transfer (simp add: normalize_quot_def) lemma fst_quot_of_fract_eq_0_iff [simp]: "fst (quot_of_fract x) = 0 \ x = 0" by transfer simp lemma snd_quot_of_fract_to_fract [simp]: "snd (quot_of_fract (to_fract x)) = 1" unfolding to_fract_def by (rule snd_quot_of_fract_Fract_whole) simp_all lemma coprime_quot_of_fract: "coprime (fst (quot_of_fract x)) (snd (quot_of_fract x))" by transfer (simp add: coprime_normalize_quot) lemma unit_factor_snd_quot_of_fract: "unit_factor (snd (quot_of_fract x)) = 1" using quot_of_fract_in_normalized_fracts[of x] by (simp add: normalized_fracts_def case_prod_unfold) lemma unit_factor_1_imp_normalized: "unit_factor x = 1 \ normalize x = x" by (subst (2) normalize_mult_unit_factor [symmetric, of x]) (simp del: normalize_mult_unit_factor) lemma normalize_snd_quot_of_fract: "normalize (snd (quot_of_fract x)) = snd (quot_of_fract x)" by (intro unit_factor_1_imp_normalized unit_factor_snd_quot_of_fract) subsection \Lifting polynomial coefficients to the field of fractions\ abbreviation (input) fract_poly where "fract_poly \ map_poly to_fract" abbreviation (input) unfract_poly where "unfract_poly \ map_poly (fst \ quot_of_fract)" lemma fract_poly_smult [simp]: "fract_poly (smult c p) = smult (to_fract c) (fract_poly p)" by (simp add: smult_conv_map_poly map_poly_map_poly o_def) lemma fract_poly_0 [simp]: "fract_poly 0 = 0" by (simp add: poly_eqI coeff_map_poly) lemma fract_poly_1 [simp]: "fract_poly 1 = 1" by (simp add: map_poly_pCons) lemma fract_poly_add [simp]: "fract_poly (p + q) = fract_poly p + fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma fract_poly_diff [simp]: "fract_poly (p - q) = fract_poly p - fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly) lemma to_fract_sum [simp]: "to_fract (sum f A) = sum (\x. to_fract (f x)) A" by (cases "finite A", induction A rule: finite_induct) simp_all lemma fract_poly_mult [simp]: "fract_poly (p * q) = fract_poly p * fract_poly q" by (intro poly_eqI) (simp_all add: coeff_map_poly coeff_mult) lemma fract_poly_eq_iff [simp]: "fract_poly p = fract_poly q \ p = q" by (auto simp: poly_eq_iff coeff_map_poly) lemma fract_poly_eq_0_iff [simp]: "fract_poly p = 0 \ p = 0" using fract_poly_eq_iff[of p 0] by (simp del: fract_poly_eq_iff) lemma fract_poly_dvd: "p dvd q \ fract_poly p dvd fract_poly q" by (auto elim!: dvdE) lemma prod_mset_fract_poly: "(\x\#A. map_poly to_fract (f x)) = fract_poly (prod_mset (image_mset f A))" by (induct A) (simp_all add: ac_simps) lemma is_unit_fract_poly_iff: "p dvd 1 \ fract_poly p dvd 1 \ content p = 1" proof safe assume A: "p dvd 1" with fract_poly_dvd [of p 1] show "is_unit (fract_poly p)" by simp from A show "content p = 1" by (auto simp: is_unit_poly_iff normalize_1_iff) next assume A: "fract_poly p dvd 1" and B: "content p = 1" from A obtain c where c: "fract_poly p = [:c:]" by (auto simp: is_unit_poly_iff) { fix n :: nat assume "n > 0" have "to_fract (coeff p n) = coeff (fract_poly p) n" by (simp add: coeff_map_poly) also note c also from \n > 0\ have "coeff [:c:] n = 0" by (simp add: coeff_pCons split: nat.splits) finally have "coeff p n = 0" by simp } hence "degree p \ 0" by (intro degree_le) simp_all with B show "p dvd 1" by (auto simp: is_unit_poly_iff normalize_1_iff elim!: degree_eq_zeroE) qed lemma fract_poly_is_unit: "p dvd 1 \ fract_poly p dvd 1" using fract_poly_dvd[of p 1] by simp lemma fract_poly_smult_eqE: fixes c :: "'a :: {idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract" assumes "fract_poly p = smult c (fract_poly q)" obtains a b where "c = to_fract b / to_fract a" "smult a p = smult b q" "coprime a b" "normalize a = a" proof - define a b where "a = fst (quot_of_fract c)" and "b = snd (quot_of_fract c)" have "smult (to_fract a) (fract_poly q) = smult (to_fract b) (fract_poly p)" by (subst smult_eq_iff) (simp_all add: a_def b_def Fract_conv_to_fract [symmetric] assms) hence "fract_poly (smult a q) = fract_poly (smult b p)" by (simp del: fract_poly_eq_iff) hence "smult b p = smult a q" by (simp only: fract_poly_eq_iff) moreover have "c = to_fract a / to_fract b" "coprime b a" "normalize b = b" by (simp_all add: a_def b_def coprime_quot_of_fract [of c] ac_simps normalize_snd_quot_of_fract Fract_conv_to_fract [symmetric]) ultimately show ?thesis by (intro that[of a b]) qed subsection \Fractional content\ abbreviation (input) Lcm_coeff_denoms :: "'a :: {semiring_Gcd,idom_divide,ring_gcd,semiring_gcd_mult_normalize} fract poly \ 'a" where "Lcm_coeff_denoms p \ Lcm (snd ` quot_of_fract ` set (coeffs p))" definition fract_content :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a fract" where "fract_content p = (let d = Lcm_coeff_denoms p in Fract (content (unfract_poly (smult (to_fract d) p))) d)" definition primitive_part_fract :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} fract poly \ 'a poly" where "primitive_part_fract p = primitive_part (unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p))" lemma primitive_part_fract_0 [simp]: "primitive_part_fract 0 = 0" by (simp add: primitive_part_fract_def) lemma fract_content_eq_0_iff [simp]: "fract_content p = 0 \ p = 0" unfolding fract_content_def Let_def Zero_fract_def by (subst eq_fract) (auto simp: Lcm_0_iff map_poly_eq_0_iff) lemma content_primitive_part_fract [simp]: fixes p :: "'a :: {semiring_gcd_mult_normalize, factorial_semiring, ring_gcd, semiring_Gcd,idom_divide} fract poly" shows "p \ 0 \ content (primitive_part_fract p) = 1" unfolding primitive_part_fract_def by (rule content_primitive_part) (auto simp: primitive_part_fract_def map_poly_eq_0_iff Lcm_0_iff) lemma content_times_primitive_part_fract: "smult (fract_content p) (fract_poly (primitive_part_fract p)) = p" proof - define p' where "p' = unfract_poly (smult (to_fract (Lcm_coeff_denoms p)) p)" have "fract_poly p' = map_poly (to_fract \ fst \ quot_of_fract) (smult (to_fract (Lcm_coeff_denoms p)) p)" unfolding primitive_part_fract_def p'_def by (subst map_poly_map_poly) (simp_all add: o_assoc) also have "\ = smult (to_fract (Lcm_coeff_denoms p)) p" proof (intro map_poly_idI, unfold o_apply) fix c assume "c \ set (coeffs (smult (to_fract (Lcm_coeff_denoms p)) p))" then obtain c' where c: "c' \ set (coeffs p)" "c = to_fract (Lcm_coeff_denoms p) * c'" by (auto simp add: Lcm_0_iff coeffs_smult split: if_splits) note c(2) also have "c' = Fract (fst (quot_of_fract c')) (snd (quot_of_fract c'))" by simp also have "to_fract (Lcm_coeff_denoms p) * \ = Fract (Lcm_coeff_denoms p * fst (quot_of_fract c')) (snd (quot_of_fract c'))" unfolding to_fract_def by (subst mult_fract) simp_all also have "snd (quot_of_fract \) = 1" by (intro snd_quot_of_fract_Fract_whole dvd_mult2 dvd_Lcm) (insert c(1), auto) finally show "to_fract (fst (quot_of_fract c)) = c" by (rule to_fract_quot_of_fract) qed also have "p' = smult (content p') (primitive_part p')" by (rule content_times_primitive_part [symmetric]) also have "primitive_part p' = primitive_part_fract p" by (simp add: primitive_part_fract_def p'_def) also have "fract_poly (smult (content p') (primitive_part_fract p)) = smult (to_fract (content p')) (fract_poly (primitive_part_fract p))" by simp finally have "smult (to_fract (content p')) (fract_poly (primitive_part_fract p)) = smult (to_fract (Lcm_coeff_denoms p)) p" . thus ?thesis by (subst (asm) smult_eq_iff) (auto simp add: Let_def p'_def Fract_conv_to_fract field_simps Lcm_0_iff fract_content_def) qed lemma fract_content_fract_poly [simp]: "fract_content (fract_poly p) = to_fract (content p)" proof - have "Lcm_coeff_denoms (fract_poly p) = 1" by (auto simp: set_coeffs_map_poly) hence "fract_content (fract_poly p) = to_fract (content (map_poly (fst \ quot_of_fract \ to_fract) p))" by (simp add: fract_content_def to_fract_def fract_collapse map_poly_map_poly del: Lcm_1_iff) also have "map_poly (fst \ quot_of_fract \ to_fract) p = p" by (intro map_poly_idI) simp_all finally show ?thesis . qed lemma content_decompose_fract: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, semiring_gcd_mult_normalize} fract poly" obtains c p' where "p = smult c (map_poly to_fract p')" "content p' = 1" proof (cases "p = 0") case True hence "p = smult 0 (map_poly to_fract 1)" "content 1 = 1" by simp_all thus ?thesis .. next case False thus ?thesis by (rule that[OF content_times_primitive_part_fract [symmetric] content_primitive_part_fract]) qed subsection \More properties of content and primitive part\ lemma lift_prime_elem_poly: assumes "prime_elem (c :: 'a :: semidom)" shows "prime_elem [:c:]" proof (rule prime_elemI) fix a b assume *: "[:c:] dvd a * b" from * have dvd: "c dvd coeff (a * b) n" for n by (subst (asm) const_poly_dvd_iff) blast { define m where "m = (GREATEST m. \c dvd coeff b m)" assume "\[:c:] dvd b" hence A: "\i. \c dvd coeff b i" by (subst (asm) const_poly_dvd_iff) blast have B: "\i. \c dvd coeff b i \ i \ degree b" by (auto intro: le_degree) have coeff_m: "\c dvd coeff b m" unfolding m_def by (rule GreatestI_ex_nat[OF A B]) have "i \ m" if "\c dvd coeff b i" for i unfolding m_def by (blast intro!: Greatest_le_nat that B) hence dvd_b: "c dvd coeff b i" if "i > m" for i using that by force have "c dvd coeff a i" for i proof (induction i rule: nat_descend_induct[of "degree a"]) case (base i) thus ?case by (simp add: coeff_eq_0) next case (descend i) let ?A = "{..i+m} - {i}" have "c dvd coeff (a * b) (i + m)" by (rule dvd) also have "coeff (a * b) (i + m) = (\k\i + m. coeff a k * coeff b (i + m - k))" by (simp add: coeff_mult) also have "{..i+m} = insert i ?A" by auto also have "(\k\\. coeff a k * coeff b (i + m - k)) = coeff a i * coeff b m + (\k\?A. coeff a k * coeff b (i + m - k))" (is "_ = _ + ?S") by (subst sum.insert) simp_all finally have eq: "c dvd coeff a i * coeff b m + ?S" . moreover have "c dvd ?S" proof (rule dvd_sum) fix k assume k: "k \ {..i+m} - {i}" show "c dvd coeff a k * coeff b (i + m - k)" proof (cases "k < i") case False with k have "c dvd coeff a k" by (intro descend.IH) simp thus ?thesis by simp next case True hence "c dvd coeff b (i + m - k)" by (intro dvd_b) simp thus ?thesis by simp qed qed ultimately have "c dvd coeff a i * coeff b m" by (simp add: dvd_add_left_iff) with assms coeff_m show "c dvd coeff a i" by (simp add: prime_elem_dvd_mult_iff) qed hence "[:c:] dvd a" by (subst const_poly_dvd_iff) blast } then show "[:c:] dvd a \ [:c:] dvd b" by blast next from assms show "[:c:] \ 0" and "\ [:c:] dvd 1" by (simp_all add: prime_elem_def is_unit_const_poly_iff) qed lemma prime_elem_const_poly_iff: fixes c :: "'a :: semidom" shows "prime_elem [:c:] \ prime_elem c" proof assume A: "prime_elem [:c:]" show "prime_elem c" proof (rule prime_elemI) fix a b assume "c dvd a * b" hence "[:c:] dvd [:a:] * [:b:]" by (simp add: mult_ac) from A and this have "[:c:] dvd [:a:] \ [:c:] dvd [:b:]" by (rule prime_elem_dvd_multD) thus "c dvd a \ c dvd b" by simp qed (insert A, auto simp: prime_elem_def is_unit_poly_iff) qed (auto intro: lift_prime_elem_poly) lemma fract_poly_dvdD: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide, semiring_gcd_mult_normalize} poly" assumes "fract_poly p dvd fract_poly q" "content p = 1" shows "p dvd q" proof - from assms(1) obtain r where r: "fract_poly q = fract_poly p * r" by (erule dvdE) - from content_decompose_fract[of r] guess c r' . note r' = this + from content_decompose_fract[of r] + obtain c r' where r': "r = smult c (map_poly to_fract r')" "content r' = 1" . from r r' have eq: "fract_poly q = smult c (fract_poly (p * r'))" by simp - from fract_poly_smult_eqE[OF this] guess a b . note ab = this + from fract_poly_smult_eqE[OF this] obtain a b + where ab: + "c = to_fract b / to_fract a" + "smult a q = smult b (p * r')" + "coprime a b" + "normalize a = a" . have "content (smult a q) = content (smult b (p * r'))" by (simp only: ab(2)) hence eq': "normalize b = a * content q" by (simp add: assms content_mult r' ab(4)) have "1 = gcd a (normalize b)" by (simp add: ab) also note eq' also have "gcd a (a * content q) = a" by (simp add: gcd_proj1_if_dvd ab(4)) finally have [simp]: "a = 1" by simp from eq ab have "q = p * ([:b:] * r')" by simp thus ?thesis by (rule dvdI) qed subsection \Polynomials over a field are a Euclidean ring\ context begin interpretation field_poly: normalization_euclidean_semiring_multiplicative where zero = "0 :: 'a :: field poly" and one = 1 and plus = plus and minus = minus and times = times and normalize = "\p. smult (inverse (lead_coeff p)) p" and unit_factor = "\p. [:lead_coeff p:]" and euclidean_size = "\p. if p = 0 then 0 else 2 ^ degree p" and divide = divide and modulo = modulo rewrites "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" and "comm_monoid_mult.prod_mset times 1 = prod_mset" and "comm_semiring_1.irreducible times 1 0 = irreducible" and "comm_semiring_1.prime_elem times 1 0 = prime_elem" proof - show "dvd.dvd (times :: 'a poly \ _) = Rings.dvd" by (simp add: dvd_dict) show "comm_monoid_mult.prod_mset times 1 = prod_mset" by (simp add: prod_mset_dict) show "comm_semiring_1.irreducible times 1 0 = irreducible" by (simp add: irreducible_dict) show "comm_semiring_1.prime_elem times 1 0 = prime_elem" by (simp add: prime_elem_dict) show "class.normalization_euclidean_semiring_multiplicative divide plus minus (0 :: 'a poly) times 1 modulo (\p. if p = 0 then 0 else 2 ^ degree p) (\p. [:lead_coeff p:]) (\p. smult (inverse (lead_coeff p)) p)" proof (standard, fold dvd_dict) fix p :: "'a poly" show "[:lead_coeff p:] * smult (inverse (lead_coeff p)) p = p" by (cases "p = 0") simp_all next fix p :: "'a poly" assume "is_unit p" then show "[:lead_coeff p:] = p" by (elim is_unit_polyE) (auto simp: monom_0 one_poly_def field_simps) next fix p :: "'a poly" assume "p \ 0" then show "is_unit [:lead_coeff p:]" by (simp add: is_unit_pCons_iff) next fix a b :: "'a poly" assume "is_unit a" thus "[:lead_coeff (a * b):] = a * [:lead_coeff b:]" by (auto elim!: is_unit_polyE) qed (auto simp: lead_coeff_mult Rings.div_mult_mod_eq intro!: degree_mod_less' degree_mult_right_le) qed lemma field_poly_irreducible_imp_prime: "prime_elem p" if "irreducible p" for p :: "'a :: field poly" using that by (fact field_poly.irreducible_imp_prime_elem) lemma field_poly_prod_mset_prime_factorization: "prod_mset (field_poly.prime_factorization p) = smult (inverse (lead_coeff p)) p" if "p \ 0" for p :: "'a :: field poly" using that by (fact field_poly.prod_mset_prime_factorization) lemma field_poly_in_prime_factorization_imp_prime: "prime_elem p" if "p \# field_poly.prime_factorization x" for p :: "'a :: field poly" by (rule field_poly.prime_imp_prime_elem, rule field_poly.in_prime_factors_imp_prime) (fact that) subsection \Primality and irreducibility in polynomial rings\ lemma nonconst_poly_irreducible_iff: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "degree p \ 0" shows "irreducible p \ irreducible (fract_poly p) \ content p = 1" proof safe assume p: "irreducible p" - from content_decompose[of p] guess p' . note p' = this + from content_decompose[of p] obtain p' where p': "p = smult (content p) p'" "content p' = 1" . hence "p = [:content p:] * p'" by simp from p this have "[:content p:] dvd 1 \ p' dvd 1" by (rule irreducibleD) moreover have "\p' dvd 1" proof assume "p' dvd 1" hence "degree p = 0" by (subst p') (auto simp: is_unit_poly_iff) with assms show False by contradiction qed ultimately show [simp]: "content p = 1" by (simp add: is_unit_const_poly_iff) show "irreducible (map_poly to_fract p)" proof (rule irreducibleI) have "fract_poly p = 0 \ p = 0" by (intro map_poly_eq_0_iff) auto with assms show "map_poly to_fract p \ 0" by auto next show "\is_unit (fract_poly p)" proof assume "is_unit (map_poly to_fract p)" hence "degree (map_poly to_fract p) = 0" by (auto simp: is_unit_poly_iff) hence "degree p = 0" by (simp add: degree_map_poly) with assms show False by contradiction qed next fix q r assume qr: "fract_poly p = q * r" - from content_decompose_fract[of q] guess cg q' . note q = this - from content_decompose_fract[of r] guess cr r' . note r = this + from content_decompose_fract[of q] + obtain cg q' where q: "q = smult cg (map_poly to_fract q')" "content q' = 1" . + from content_decompose_fract[of r] + obtain cr r' where r: "r = smult cr (map_poly to_fract r')" "content r' = 1" . from qr q r p have nz: "cg \ 0" "cr \ 0" by auto from qr have eq: "fract_poly p = smult (cr * cg) (fract_poly (q' * r'))" by (simp add: q r) - from fract_poly_smult_eqE[OF this] guess a b . note ab = this + from fract_poly_smult_eqE[OF this] obtain a b + where ab: "cr * cg = to_fract b / to_fract a" + "smult a p = smult b (q' * r')" "coprime a b" "normalize a = a" . hence "content (smult a p) = content (smult b (q' * r'))" by (simp only:) with ab(4) have a: "a = normalize b" by (simp add: content_mult q r) then have "normalize b = gcd a b" by simp with \coprime a b\ have "normalize b = 1" by simp then have "a = 1" "is_unit b" by (simp_all add: a normalize_1_iff) note eq also from ab(1) \a = 1\ have "cr * cg = to_fract b" by simp also have "smult \ (fract_poly (q' * r')) = fract_poly (smult b (q' * r'))" by simp finally have "p = ([:b:] * q') * r'" by (simp del: fract_poly_smult) from p and this have "([:b:] * q') dvd 1 \ r' dvd 1" by (rule irreducibleD) hence "q' dvd 1 \ r' dvd 1" by (auto dest: dvd_mult_right simp del: mult_pCons_left) hence "fract_poly q' dvd 1 \ fract_poly r' dvd 1" by (auto simp: fract_poly_is_unit) with q r show "is_unit q \ is_unit r" by (auto simp add: is_unit_smult_iff dvd_field_iff nz) qed next assume irred: "irreducible (fract_poly p)" and primitive: "content p = 1" show "irreducible p" proof (rule irreducibleI) from irred show "p \ 0" by auto next from irred show "\p dvd 1" by (auto simp: irreducible_def dest: fract_poly_is_unit) next fix q r assume qr: "p = q * r" hence "fract_poly p = fract_poly q * fract_poly r" by simp from irred and this have "fract_poly q dvd 1 \ fract_poly r dvd 1" by (rule irreducibleD) with primitive qr show "q dvd 1 \ r dvd 1" by (auto simp: content_prod_eq_1_iff is_unit_fract_poly_iff) qed qed private lemma irreducible_imp_prime_poly: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "irreducible p" shows "prime_elem p" proof (cases "degree p = 0") case True with assms show ?thesis by (auto simp: prime_elem_const_poly_iff irreducible_const_poly_iff intro!: irreducible_imp_prime_elem elim!: degree_eq_zeroE) next case False from assms False have irred: "irreducible (fract_poly p)" and primitive: "content p = 1" by (simp_all add: nonconst_poly_irreducible_iff) from irred have prime: "prime_elem (fract_poly p)" by (rule field_poly_irreducible_imp_prime) show ?thesis proof (rule prime_elemI) fix q r assume "p dvd q * r" hence "fract_poly p dvd fract_poly (q * r)" by (rule fract_poly_dvd) hence "fract_poly p dvd fract_poly q * fract_poly r" by simp from prime and this have "fract_poly p dvd fract_poly q \ fract_poly p dvd fract_poly r" by (rule prime_elem_dvd_multD) with primitive show "p dvd q \ p dvd r" by (auto dest: fract_poly_dvdD) qed (insert assms, auto simp: irreducible_def) qed lemma degree_primitive_part_fract [simp]: "degree (primitive_part_fract p) = degree p" proof - have "p = smult (fract_content p) (fract_poly (primitive_part_fract p))" by (simp add: content_times_primitive_part_fract) also have "degree \ = degree (primitive_part_fract p)" by (auto simp: degree_map_poly) finally show ?thesis .. qed lemma irreducible_primitive_part_fract: fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" assumes "irreducible p" shows "irreducible (primitive_part_fract p)" proof - from assms have deg: "degree (primitive_part_fract p) \ 0" by (intro notI) (auto elim!: degree_eq_zeroE simp: irreducible_def is_unit_poly_iff dvd_field_iff) hence [simp]: "p \ 0" by auto note \irreducible p\ also have "p = [:fract_content p:] * fract_poly (primitive_part_fract p)" by (simp add: content_times_primitive_part_fract) also have "irreducible \ \ irreducible (fract_poly (primitive_part_fract p))" by (intro irreducible_mult_unit_left) (simp_all add: is_unit_poly_iff dvd_field_iff) finally show ?thesis using deg by (simp add: nonconst_poly_irreducible_iff) qed lemma prime_elem_primitive_part_fract: fixes p :: "'a :: {idom_divide, ring_gcd, factorial_semiring, semiring_Gcd,semiring_gcd_mult_normalize} fract poly" shows "irreducible p \ prime_elem (primitive_part_fract p)" by (intro irreducible_imp_prime_poly irreducible_primitive_part_fract) lemma irreducible_linear_field_poly: fixes a b :: "'a::field" assumes "b \ 0" shows "irreducible [:a,b:]" proof (rule irreducibleI) fix p q assume pq: "[:a,b:] = p * q" also from pq assms have "degree \ = degree p + degree q" by (intro degree_mult_eq) auto finally have "degree p = 0 \ degree q = 0" using assms by auto with assms pq show "is_unit p \ is_unit q" by (auto simp: is_unit_const_poly_iff dvd_field_iff elim!: degree_eq_zeroE) qed (insert assms, auto simp: is_unit_poly_iff) lemma prime_elem_linear_field_poly: "(b :: 'a :: field) \ 0 \ prime_elem [:a,b:]" by (rule field_poly_irreducible_imp_prime, rule irreducible_linear_field_poly) lemma irreducible_linear_poly: fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ irreducible [:a,b:]" by (auto intro!: irreducible_linear_field_poly simp: nonconst_poly_irreducible_iff content_def map_poly_pCons) lemma prime_elem_linear_poly: fixes a b :: "'a::{idom_divide,ring_gcd,factorial_semiring,semiring_Gcd,semiring_gcd_mult_normalize}" shows "b \ 0 \ coprime a b \ prime_elem [:a,b:]" by (rule irreducible_imp_prime_poly, rule irreducible_linear_poly) subsection \Prime factorisation of polynomials\ private lemma poly_prime_factorization_exists_content_1: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" "content p = 1" shows "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize p" proof - let ?P = "field_poly.prime_factorization (fract_poly p)" define c where "c = prod_mset (image_mset fract_content ?P)" define c' where "c' = c * to_fract (lead_coeff p)" define e where "e = prod_mset (image_mset primitive_part_fract ?P)" define A where "A = image_mset (normalize \ primitive_part_fract) ?P" have "content e = (\x\#field_poly.prime_factorization (map_poly to_fract p). content (primitive_part_fract x))" by (simp add: e_def content_prod_mset multiset.map_comp o_def) also have "image_mset (\x. content (primitive_part_fract x)) ?P = image_mset (\_. 1) ?P" by (intro image_mset_cong content_primitive_part_fract) auto finally have content_e: "content e = 1" by simp from \p \ 0\ have "fract_poly p = [:lead_coeff (fract_poly p):] * smult (inverse (lead_coeff (fract_poly p))) (fract_poly p)" by simp also have "[:lead_coeff (fract_poly p):] = [:to_fract (lead_coeff p):]" by (simp add: monom_0 degree_map_poly coeff_map_poly) also from assms have "smult (inverse (lead_coeff (fract_poly p))) (fract_poly p) = prod_mset ?P" by (subst field_poly_prod_mset_prime_factorization) simp_all also have "\ = prod_mset (image_mset id ?P)" by simp also have "image_mset id ?P = image_mset (\x. [:fract_content x:] * fract_poly (primitive_part_fract x)) ?P" by (intro image_mset_cong) (auto simp: content_times_primitive_part_fract) also have "prod_mset \ = smult c (fract_poly e)" by (subst prod_mset.distrib) (simp_all add: prod_mset_fract_poly prod_mset_const_poly c_def e_def) also have "[:to_fract (lead_coeff p):] * \ = smult c' (fract_poly e)" by (simp add: c'_def) finally have eq: "fract_poly p = smult c' (fract_poly e)" . also obtain b where b: "c' = to_fract b" "is_unit b" proof - - from fract_poly_smult_eqE[OF eq] guess a b . note ab = this + from fract_poly_smult_eqE[OF eq] + obtain a b where ab: + "c' = to_fract b / to_fract a" + "smult a p = smult b e" + "coprime a b" + "normalize a = a" . from ab(2) have "content (smult a p) = content (smult b e)" by (simp only: ) with assms content_e have "a = normalize b" by (simp add: ab(4)) with ab have ab': "a = 1" "is_unit b" by (simp_all add: normalize_1_iff) with ab ab' have "c' = to_fract b" by auto from this and \is_unit b\ show ?thesis by (rule that) qed hence "smult c' (fract_poly e) = fract_poly (smult b e)" by simp finally have "p = smult b e" by (simp only: fract_poly_eq_iff) hence "p = [:b:] * e" by simp with b have "normalize p = normalize e" by (simp only: normalize_mult) (simp add: is_unit_normalize is_unit_poly_iff) also have "normalize e = prod_mset A" by (simp add: multiset.map_comp e_def A_def normalize_prod_mset) finally have "prod_mset A = normalize p" .. have "prime_elem p" if "p \# A" for p using that by (auto simp: A_def prime_elem_primitive_part_fract prime_elem_imp_irreducible dest!: field_poly_in_prime_factorization_imp_prime ) from this and \prod_mset A = normalize p\ show ?thesis by (intro exI[of _ A]) blast qed lemma poly_prime_factorization_exists: fixes p :: "'a :: {factorial_semiring,semiring_Gcd,ring_gcd,idom_divide,semiring_gcd_mult_normalize} poly" assumes "p \ 0" shows "\A. (\p. p \# A \ prime_elem p) \ normalize (prod_mset A) = normalize p" proof - define B where "B = image_mset (\x. [:x:]) (prime_factorization (content p))" have "\A. (\p. p \# A \ prime_elem p) \ prod_mset A = normalize (primitive_part p)" by (rule poly_prime_factorization_exists_content_1) (insert assms, simp_all) - then guess A by (elim exE conjE) note A = this + then obtain A where A: "\p. p \# A \ prime_elem p" "\\<^sub># A = normalize (primitive_part p)" + by blast have "normalize (prod_mset (A + B)) = normalize (prod_mset A * normalize (prod_mset B))" by simp also from assms have "normalize (prod_mset B) = normalize [:content p:]" by (simp add: prod_mset_const_poly normalize_const_poly prod_mset_prime_factorization_weak B_def) also have "prod_mset A = normalize (primitive_part p)" using A by simp finally have "normalize (prod_mset (A + B)) = normalize (primitive_part p * [:content p:])" by simp moreover have "\p. p \# B \ prime_elem p" by (auto simp: B_def intro!: lift_prime_elem_poly dest: in_prime_factors_imp_prime) ultimately show ?thesis using A by (intro exI[of _ "A + B"]) (auto) qed end subsection \Typeclass instances\ instance poly :: ("{factorial_ring_gcd,semiring_gcd_mult_normalize}") factorial_semiring by standard (rule poly_prime_factorization_exists) instantiation poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") factorial_ring_gcd begin definition gcd_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "gcd_poly = gcd_factorial" definition lcm_poly :: "'a poly \ 'a poly \ 'a poly" where [code del]: "lcm_poly = lcm_factorial" definition Gcd_poly :: "'a poly set \ 'a poly" where [code del]: "Gcd_poly = Gcd_factorial" definition Lcm_poly :: "'a poly set \ 'a poly" where [code del]: "Lcm_poly = Lcm_factorial" instance by standard (simp_all add: gcd_poly_def lcm_poly_def Gcd_poly_def Lcm_poly_def) end instance poly :: ("{factorial_ring_gcd, semiring_gcd_mult_normalize}") semiring_gcd_mult_normalize .. instantiation poly :: ("{field,factorial_ring_gcd,semiring_gcd_mult_normalize}") "{unique_euclidean_ring, normalization_euclidean_semiring}" begin definition euclidean_size_poly :: "'a poly \ nat" where "euclidean_size_poly p = (if p = 0 then 0 else 2 ^ degree p)" definition division_segment_poly :: "'a poly \ 'a poly" where [simp]: "division_segment_poly p = 1" instance proof show "(q * p + r) div p = q" if "p \ 0" and "euclidean_size r < euclidean_size p" for q p r :: "'a poly" proof - from \p \ 0\ eucl_rel_poly [of r p] have "eucl_rel_poly (r + q * p) p (q + r div p, r mod p)" by (simp add: eucl_rel_poly_iff distrib_right) then have "(r + q * p) div p = q + r div p" by (rule div_poly_eq) with that show ?thesis by (cases "r = 0") (simp_all add: euclidean_size_poly_def div_poly_less ac_simps) qed qed (auto simp: euclidean_size_poly_def Rings.div_mult_mod_eq div_poly_less degree_mult_eq power_add intro!: degree_mod_less' split: if_splits) end instance poly :: ("{field, normalization_euclidean_semiring, factorial_ring_gcd, semiring_gcd_mult_normalize}") euclidean_ring_gcd by (rule euclidean_ring_gcd_class.intro, rule factorial_euclidean_semiring_gcdI) standard subsection \Polynomial GCD\ lemma gcd_poly_decompose: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "gcd p q = smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" proof (rule sym, rule gcdI) have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content p:] * primitive_part p" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd p" by simp next have "[:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q) dvd [:content q:] * primitive_part q" by (intro mult_dvd_mono) simp_all thus "smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q)) dvd q" by simp next fix d assume "d dvd p" "d dvd q" hence "[:content d:] * primitive_part d dvd [:gcd (content p) (content q):] * gcd (primitive_part p) (primitive_part q)" by (intro mult_dvd_mono) auto thus "d dvd smult (gcd (content p) (content q)) (gcd (primitive_part p) (primitive_part q))" by simp qed (auto simp: normalize_smult) lemma gcd_poly_pseudo_mod: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" assumes nz: "q \ 0" and prim: "content p = 1" "content q = 1" shows "gcd p q = gcd q (primitive_part (pseudo_mod p q))" proof - define r s where "r = fst (pseudo_divmod p q)" and "s = snd (pseudo_divmod p q)" define a where "a = [:coeff q (degree q) ^ (Suc (degree p) - degree q):]" have [simp]: "primitive_part a = unit_factor a" by (simp add: a_def unit_factor_poly_def unit_factor_power monom_0) from nz have [simp]: "a \ 0" by (auto simp: a_def) have rs: "pseudo_divmod p q = (r, s)" by (simp add: r_def s_def) have "gcd (q * r + s) q = gcd q s" using gcd_add_mult[of q r s] by (simp add: gcd.commute add_ac mult_ac) with pseudo_divmod(1)[OF nz rs] have "gcd (p * a) q = gcd q s" by (simp add: a_def) also from prim have "gcd (p * a) q = gcd p q" by (subst gcd_poly_decompose) (auto simp: primitive_part_mult gcd_mult_unit1 primitive_part_prim simp del: mult_pCons_right ) also from prim have "gcd q s = gcd q (primitive_part s)" by (subst gcd_poly_decompose) (simp_all add: primitive_part_prim) also have "s = pseudo_mod p q" by (simp add: s_def pseudo_mod_def) finally show ?thesis . qed lemma degree_pseudo_mod_less: assumes "q \ 0" "pseudo_mod p q \ 0" shows "degree (pseudo_mod p q) < degree q" using pseudo_mod(2)[of q p] assms by auto function gcd_poly_code_aux :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code_aux p q = (if q = 0 then normalize p else gcd_poly_code_aux q (primitive_part (pseudo_mod p q)))" by auto termination by (relation "measure ((\p. if p = 0 then 0 else Suc (degree p)) \ snd)") (auto simp: degree_pseudo_mod_less) declare gcd_poly_code_aux.simps [simp del] lemma gcd_poly_code_aux_correct: assumes "content p = 1" "q = 0 \ content q = 1" shows "gcd_poly_code_aux p q = gcd p q" using assms proof (induction p q rule: gcd_poly_code_aux.induct) case (1 p q) show ?case proof (cases "q = 0") case True thus ?thesis by (subst gcd_poly_code_aux.simps) auto next case False hence "gcd_poly_code_aux p q = gcd_poly_code_aux q (primitive_part (pseudo_mod p q))" by (subst gcd_poly_code_aux.simps) simp_all also from "1.prems" False have "primitive_part (pseudo_mod p q) = 0 \ content (primitive_part (pseudo_mod p q)) = 1" by (cases "pseudo_mod p q = 0") auto with "1.prems" False have "gcd_poly_code_aux q (primitive_part (pseudo_mod p q)) = gcd q (primitive_part (pseudo_mod p q))" by (intro 1) simp_all also from "1.prems" False have "\ = gcd p q" by (intro gcd_poly_pseudo_mod [symmetric]) auto finally show ?thesis . qed qed definition gcd_poly_code :: "'a :: factorial_ring_gcd poly \ 'a poly \ 'a poly" where "gcd_poly_code p q = (if p = 0 then normalize q else if q = 0 then normalize p else smult (gcd (content p) (content q)) (gcd_poly_code_aux (primitive_part p) (primitive_part q)))" lemma gcd_poly_code [code]: "gcd p q = gcd_poly_code p q" by (simp add: gcd_poly_code_def gcd_poly_code_aux_correct gcd_poly_decompose [symmetric]) lemma lcm_poly_code [code]: fixes p q :: "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly" shows "lcm p q = normalize (p * q div gcd p q)" by (fact lcm_gcd) lemmas Gcd_poly_set_eq_fold [code] = Gcd_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] lemmas Lcm_poly_set_eq_fold [code] = Lcm_set_eq_fold [where ?'a = "'a :: {factorial_ring_gcd,semiring_gcd_mult_normalize} poly"] text \Example: @{lemma "Lcm {[:1, 2, 3:], [:2, 3, 4:]} = [:[:2:], [:7:], [:16:], [:17:], [:12 :: int:]:]" by eval} \ end diff --git a/src/HOL/Number_Theory/Residue_Primitive_Roots.thy b/src/HOL/Number_Theory/Residue_Primitive_Roots.thy --- a/src/HOL/Number_Theory/Residue_Primitive_Roots.thy +++ b/src/HOL/Number_Theory/Residue_Primitive_Roots.thy @@ -1,1152 +1,1160 @@ (* File: HOL/Number_Theory/Residue_Primitive_Roots.thy Author: Manuel Eberl Primitive roots in residue rings: Definition and existence criteria *) section \Primitive roots in residue rings and Carmichael's function\ theory Residue_Primitive_Roots imports Pocklington begin text \ This theory develops the notions of primitive roots (generators) in residue rings. It also provides a definition and all the basic properties of Carmichael's function $\lambda(n)$, which is strongly related to this. The proofs mostly follow Apostol's presentation \ subsection \Primitive roots in residue rings\ text \ A primitive root of a residue ring modulo \n\ is an element \g\ that \<^emph>\generates\ the ring, i.\,e.\ such that for each \x\ coprime to \n\ there exists an \i\ such that $x = g^i$. A simpler definition is that \g\ must have the same order as the cardinality of the multiplicative group, which is $\varphi(n)$. Note that for convenience, this definition does \<^emph>\not\ demand \g < n\. \ inductive residue_primroot :: "nat \ nat \ bool" where "n > 0 \ coprime n g \ ord n g = totient n \ residue_primroot n g" lemma residue_primroot_def [code]: "residue_primroot n x \ n > 0 \ coprime n x \ ord n x = totient n" by (simp add: residue_primroot.simps) lemma not_residue_primroot_0 [simp]: "~residue_primroot 0 x" by (auto simp: residue_primroot_def) lemma residue_primroot_mod [simp]: "residue_primroot n (x mod n) = residue_primroot n x" by (cases "n = 0") (simp_all add: residue_primroot_def) lemma residue_primroot_cong: assumes "[x = x'] (mod n)" shows "residue_primroot n x = residue_primroot n x'" proof - have "residue_primroot n x = residue_primroot n (x mod n)" by simp also have "x mod n = x' mod n" using assms by (simp add: cong_def) also have "residue_primroot n (x' mod n) = residue_primroot n x'" by simp finally show ?thesis . qed lemma not_residue_primroot_0_right [simp]: "residue_primroot n 0 \ n = 1" by (auto simp: residue_primroot_def) lemma residue_primroot_1_iff: "residue_primroot n (Suc 0) \ n \ {1, 2}" proof assume *: "residue_primroot n (Suc 0)" with totient_gt_1[of n] have "n \ 2" by (cases "n \ 2") (auto simp: residue_primroot_def) hence "n \ {0, 1, 2}" by auto thus "n \ {1, 2}" using * by (auto simp: residue_primroot_def) qed (auto simp: residue_primroot_def) subsection \Primitive roots modulo a prime\ text \ For prime \p\, we now analyse the number of elements in the ring $\mathbb{Z}/p\mathbb{Z}$ whose order is precisely \d\ for each \d\. \ context fixes n :: nat and \ assumes n: "n > 1" defines "\ \ (\d. card {x\totatives n. ord n x = d})" begin lemma elements_with_ord_restrict_totatives: "d > 0 \ {x\{..totatives n. ord n x = d}" using n by (auto simp: totatives_def coprime_commute intro!: Nat.gr0I le_neq_trans) lemma prime_elements_with_ord: assumes "\ d \ 0" and "prime n" and a: "a \ totatives n" "ord n a = d" "a \ 1" shows "inj_on (\k. a ^ k mod n) {..{..k. a ^ k mod n) ` {..k. a ^ k mod n) (totatives d) {x\{..k. a ^ k mod n) {.. 0" by (auto simp: totatives_def coprime_commute) moreover have "d \ 1" using a n by (auto simp: ord_eq_Suc_0_iff totatives_less cong_def) ultimately have d: "d > 1" by simp have *: "(\k. a ^ k mod n) ` {..{..{.. d" using assms a by (intro roots_mod_prime_bound) (auto simp: totatives_def coprime_commute) also have "\ = card ((\k. a ^ k mod n) ` {.. {.. \" . next show "(\k. a ^ k mod n) ` {.. {x \ {..auto simp: ord_divides'\) thus "[(a ^ k mod n) ^ d = 1] (mod n)" by (simp add: power_mult [symmetric] cong_def power_mod mult.commute) qed (use \prime n\ in \auto dest: prime_gt_1_nat\) qed auto thus "{x\{..k. a ^ k mod n) ` {..k. a ^ k mod n) (totatives d) {x\{.. (\k. a ^ k mod n) ` totatives d" then obtain k where "b = a ^ k mod n" "k \ totatives d" by auto thus "b \ {b \ {.. {x \ {.. {x\{..coprime n b\ by (auto simp: ord_power) also have "ord n b = d" by (simp add: b) finally have "coprime k d" unfolding coprime_iff_gcd_eq_1 using d a by (subst (asm) div_eq_dividend_iff) auto with k b d have "k \ totatives d" by (auto simp: totatives_def intro!: Nat.gr0I) with k show "b \ (\k. a ^ k mod n) ` totatives d" by blast qed (use d n in \auto simp: totatives_less\) qed lemma prime_card_elements_with_ord: assumes "\ d \ 0" and "prime n" shows "\ d = totient d" proof (cases "d = 1") case True have "\ 1 = 1" using elements_with_ord_1[of n] n by (simp add: \_def) thus ?thesis using True by simp next case False from assms obtain a where a: "a \ totatives n" "ord n a = d" by (auto simp: \_def) from a have "d > 0" by (auto intro!: Nat.gr0I simp: ord_eq_0 totatives_def coprime_commute) from a and False have "a \ 1" by auto from bij_betw_same_card[OF prime_elements_with_ord(3)[OF assms a this]] show "\ d = totient d" using elements_with_ord_restrict_totatives[of d] False a \d > 0\ by (simp add: \_def totient_def) qed lemma prime_sum_card_elements_with_ord_eq_totient: "(\d | d dvd totient n. \ d) = totient n" proof - have "totient n = card (totatives n)" by (simp add: totient_def) also have "totatives n = (\d\{d. d dvd totient n}. {x\totatives n. ord n x = d})" by (force simp: order_divides_totient totatives_def coprime_commute) also have "card \ = (\d | d dvd totient n. \ d)" unfolding \_def using n by (subst card_UN_disjoint) (auto intro!: finite_divisors_nat) finally show ?thesis .. qed text \ We can now show that the number of elements of order \d\ is $\varphi(d)$ if $d\mid p - 1$ and 0 otherwise. \ theorem prime_card_elements_with_ord_eq_totient: assumes "prime n" shows "\ d = (if d dvd n - 1 then totient d else 0)" proof (cases "d dvd totient n") case False thus ?thesis using order_divides_totient[of n] assms by (auto simp: \_def totient_prime totatives_def coprime_commute[of n]) next case True have "\ d = totient d" proof (rule ccontr) assume neq: "\ d \ totient d" have le: "\ d \ totient d" if "d dvd totient n" for d using prime_card_elements_with_ord[of d] assms by (cases "\ d = 0") auto from neq and le[of d] and True have less: "\ d < totient d" by auto have "totient n = (\d | d dvd totient n. \ d)" using prime_sum_card_elements_with_ord_eq_totient .. also have "\ < (\d | d dvd totient n. totient d)" by (rule sum_strict_mono_ex1) (use n le less assms True in \auto intro!: finite_divisors_nat\) also have "\ = totient n" using totient_divisor_sum . finally show False by simp qed with True show ?thesis using assms by (simp add: totient_prime) qed text \ As a corollary, we get that the number of primitive roots modulo a prime \p\ is $\varphi(p - 1)$. Since this number is positive, we also get that there \<^emph>\is\ at least one primitive root modulo \p\. \ lemma assumes "prime n" shows prime_card_primitive_roots: "card {x\totatives n. ord n x = n - 1} = totient (n - 1)" "card {x\{..x. residue_primroot n x" proof - show *: "card {x\totatives n. ord n x = n - 1} = totient (n - 1)" using prime_card_elements_with_ord_eq_totient[of "n - 1"] assms by (auto simp: totient_prime \_def) thus "card {x\{.. 0" using n by auto finally show "\x. residue_primroot n x" using assms prime_gt_1_nat[of n] by (subst (asm) card_gt_0_iff) (auto simp: residue_primroot_def totient_prime totatives_def coprime_commute) qed end subsection \Primitive roots modulo powers of an odd prime\ text \ Any primitive root \g\ modulo an odd prime \p\ is also a primitive root modulo $p^k$ for all $k > 0$ if $[g^{p - 1} \neq 1]\ (\text{mod}\ p^2)$. To show this, we first need the following lemma. \ lemma residue_primroot_power_prime_power_neq_1: assumes "k \ 2" assumes p: "prime p" "odd p" and "residue_primroot p g" and "[g ^ (p - 1) \ 1] (mod p\<^sup>2)" shows "[g ^ totient (p ^ (k - 1)) \ 1] (mod (p ^ k))" using assms(1) proof (induction k rule: dec_induct) case base thus ?case using assms by (simp add: totient_prime) next case (step k) from p have "p > 2" using prime_gt_1_nat[of p] by (cases "p = 2") auto from assms have g: "g > 0" by (auto intro!: Nat.gr0I) have "[g ^ totient (p ^ (k - 1)) = 1] (mod p ^ (k - 1))" using assms by (intro euler_theorem) (auto simp: residue_primroot_def totatives_def coprime_commute) from cong_to_1_nat[OF this] obtain t where *: "g ^ totient (p ^ (k - 1)) - 1 = p ^ (k - 1) * t" by auto have t: "g ^ totient (p ^ (k - 1)) = p ^ (k - 1) * t + 1" using g by (subst * [symmetric]) auto have "\p dvd t" proof assume "p dvd t" then obtain q where [simp]: "t = p * q" by auto from t have "g ^ totient (p ^ (k - 1)) = p ^ k * q + 1" using \k \ 2\ by (cases k) auto hence "[g ^ totient (p ^ (k - 1)) = p ^ k * q + 1] (mod p ^ k)" by simp also have "[p ^ k * q + 1 = 0 * q + 1] (mod p ^ k)" by (intro cong_add cong_mult) (auto simp: cong_0_iff) finally have "[g ^ totient (p ^ (k - 1)) = 1] (mod p ^ k)" by simp with step.IH show False by contradiction qed from t have "(g ^ totient (p ^ (k - 1))) ^ p = (p ^ (k - 1) * t + 1) ^ p" by (rule arg_cong) also have "(g ^ totient (p ^ (k - 1))) ^ p = g ^ (p * totient (p ^ (k - 1)))" by (simp add: power_mult [symmetric] mult.commute) also have "p * totient (p ^ (k - 1)) = totient (p ^ k)" using p \k \ 2\ by (simp add: totient_prime_power Suc_diff_Suc flip: power_Suc) also have "(p ^ (k - 1) * t + 1) ^ p = (\i\p. (p choose i) * t ^ i * p ^ (i * (k - 1)))" by (subst binomial) (simp_all add: mult_ac power_mult_distrib power_mult [symmetric]) finally have "[g ^ totient (p ^ k) = (\i\p. (p choose i) * t ^ i * p ^ (i * (k - 1)))] (mod (p ^ Suc k))" (is "[_ = ?rhs] (mod _)") by simp also have "[?rhs = (\i\p. if i \ 1 then (p choose i) * t ^ i * p ^ (i * (k - 1)) else 0)] (mod (p ^ Suc k))" (is "[sum ?f _ = sum ?g _] (mod _)") proof (intro cong_sum) fix i assume i: "i \ {..p}" consider "i \ 1" | "i = 2" | "i > 2" by force thus "[?f i = ?g i] (mod (p ^ Suc k))" proof cases assume i: "i > 2" have "Suc k \ 3 * (k - 1)" using \k \ 2\ by (simp add: algebra_simps) also have "3 * (k - 1) \ i * (k - 1)" using i by (intro mult_right_mono) auto finally have "p ^ Suc k dvd ?f i" by (intro dvd_mult le_imp_power_dvd) thus "[?f i = ?g i] (mod (p ^ Suc k))" by (simp add: cong_0_iff) next assume [simp]: "i = 2" have "?f i = p * (p - 1) div 2 * t\<^sup>2 * p ^ (2 * (k - 1))" using choose_two[of p] by simp also have "p * (p - 1) div 2 = (p - 1) div 2 * p" using \odd p\ by (auto elim!: oddE) also have "\ * t\<^sup>2 * p ^ (2 * (k - 1)) = (p - 1) div 2 * t\<^sup>2 * (p * p ^ (2 * (k - 1)))" by (simp add: algebra_simps) also have "p * p ^ (2 * (k - 1)) = p ^ (2 * k - 1)" using \k \ 2\ by (cases k) auto also have "p ^ Suc k dvd (p - 1) div 2 * t\<^sup>2 * p ^ (2 * k - 1)" using \k \ 2\ by (intro dvd_mult le_imp_power_dvd) auto finally show "[?f i = ?g i] (mod (p ^ Suc k))" by (simp add: cong_0_iff) qed auto qed also have "(\i\p. ?g i) = (\i\1. ?f i)" using p prime_gt_1_nat[of p] by (intro sum.mono_neutral_cong_right) auto also have "\ = 1 + t * p ^ k" using choose_two[of p] \k \ 2\ by (cases k) simp_all finally have eq: "[g ^ totient (p ^ k) = 1 + t * p ^ k] (mod p ^ Suc k)" . have "[g ^ totient (p ^ k) \ 1] (mod p ^ Suc k)" proof assume "[g ^ totient (p ^ k) = 1] (mod p ^ Suc k)" hence "[g ^ totient (p ^ k) - g ^ totient (p ^ k) = 1 + t * p ^ k - 1] (mod p ^ Suc k)" by (intro cong_diff_nat eq) auto hence "[t * p ^ k = 0] (mod p ^ Suc k)" by (simp add: cong_sym_eq) hence "p * p ^ k dvd t * p ^ k" by (simp add: cong_0_iff) hence "p dvd t" using \p > 2\ by simp with \\p dvd t\ show False by contradiction qed thus ?case by simp qed text \ We can now show that primitive roots modulo \p\ with the above condition are indeed also primitive roots modulo $p^k$. \ proposition residue_primroot_prime_lift_iff: assumes p: "prime p" "odd p" and "residue_primroot p g" shows "(\k>0. residue_primroot (p ^ k) g) \ [g ^ (p - 1) \ 1] (mod p\<^sup>2)" proof - from assms have g: "coprime p g" "ord p g = p - 1" by (auto simp: residue_primroot_def totient_prime) show ?thesis proof assume "\k>0. residue_primroot (p ^ k) g" hence "residue_primroot (p\<^sup>2) g" by auto hence "ord (p\<^sup>2) g = totient (p\<^sup>2)" by (simp_all add: residue_primroot_def) thus "[g ^ (p - 1) \ 1] (mod p\<^sup>2)" using g assms prime_gt_1_nat[of p] by (auto simp: ord_divides' totient_prime_power) next assume g': "[g ^ (p - 1) \ 1] (mod p\<^sup>2)" have "residue_primroot (p ^ k) g" if "k > 0" for k proof (cases "k = 1") case False with that have k: "k > 1" by simp from g have coprime: "coprime (p ^ k) g" by (auto simp: totatives_def coprime_commute) define t where "t = ord (p ^ k) g" have "[g ^ t = 1] (mod (p ^ k))" by (simp add: t_def ord_divides') also have "p ^ k = p * p ^ (k - 1)" using k by (cases k) auto finally have "[g ^ t = 1] (mod p)" by (rule cong_modulus_mult_nat) hence "totient p dvd t" using g p by (simp add: ord_divides' totient_prime) then obtain q where t: "t = totient p * q" by auto have "t dvd totient (p ^ k)" using coprime by (simp add: t_def order_divides_totient) with t p k have "q dvd p ^ (k - 1)" using prime_gt_1_nat[of p] by (auto simp: totient_prime totient_prime_power) then obtain b where b: "b \ k - 1" "q = p ^ b" using divides_primepow_nat[of p q "k - 1"] p by auto have "b = k - 1" proof (rule ccontr) assume "b \ k - 1" with b have "b < k - 1" by simp have "t = p ^ b * (p - 1)" using b p by (simp add: t totient_prime) also have "\ dvd p ^ (k - 2) * (p - 1)" using \b < k - 1\ by (intro mult_dvd_mono le_imp_power_dvd) auto also have "\ = totient (p ^ (k - 1))" using k p by (simp add: totient_prime_power numeral_2_eq_2) finally have "[g ^ totient (p ^ (k - 1)) = 1] (mod (p ^ k))" by (simp add: ord_divides' t_def) with residue_primroot_power_prime_power_neq_1[of k p g] p k assms g' show False by auto qed hence "t = totient (p ^ k)" using p k by (simp add: t b totient_prime totient_prime_power) thus "residue_primroot (p ^ k) g" using g one_less_power[of p k] prime_gt_1_nat[of p] p k by (simp add: residue_primroot_def t_def) qed (use assms in auto) thus "\k>0. residue_primroot (p ^ k) g" by blast qed qed text \ If \p\ is an odd prime, there is always a primitive root \g\ modulo \p\, and if \g\ does not fulfil the above assumption required for it to be liftable to $p^k$, we can use $g + p$, which is also a primitive root modulo \p\ and \<^emph>\does\ fulfil the assumption. This shows that any modulus that is a power of an odd prime has a primitive root. \ theorem residue_primroot_odd_prime_power_exists: assumes p: "prime p" "odd p" obtains g where "\k>0. residue_primroot (p ^ k) g" proof - obtain g where g: "residue_primroot p g" using prime_primitive_root_exists[of p] assms prime_gt_1_nat[of p] by auto have "\g. residue_primroot p g \ [g ^ (p - 1) \ 1] (mod p\<^sup>2)" proof (cases "[g ^ (p - 1) = 1] (mod p\<^sup>2)") case True define g' where "g' = p + g" note g also have "residue_primroot p g \ residue_primroot p g'" unfolding g'_def by (rule residue_primroot_cong) (auto simp: cong_def) finally have g': "residue_primroot p g'" . have "[g' ^ (p - 1) = (\k\p-1. ((p-1) choose k) * g ^ (p - Suc k) * p ^ k)] (mod p\<^sup>2)" (is "[_ = ?rhs] (mod _)") by (simp add: g'_def binomial mult_ac) also have "[?rhs = (\k\p-1. if k \ 1 then ((p-1) choose k) * g ^ (p - Suc k) * p ^ k else 0)] (mod p\<^sup>2)" (is "[sum ?f _ = sum ?g _] (mod _)") proof (intro cong_sum) fix k assume "k \ {..p-1}" show "[?f k = ?g k] (mod p\<^sup>2)" proof (cases "k \ 1") case False have "p\<^sup>2 dvd ?f k" using False by (intro dvd_mult le_imp_power_dvd) auto thus ?thesis using False by (simp add: cong_0_iff) qed auto qed also have "sum ?g {..p-1} = sum ?f {0, 1}" using prime_gt_1_nat[of p] p by (intro sum.mono_neutral_cong_right) auto also have "\ = g ^ (p - 1) + p * (p - 1) * g ^ (p - 2)" using p by (simp add: numeral_2_eq_2) also have "[g ^ (p - 1) + p * (p - 1) * g ^ (p - 2) = 1 + p * (p - 1) * g ^ (p - 2)] (mod p\<^sup>2)" by (intro cong_add True) auto finally have "[g' ^ (p - 1) = 1 + p * (p - 1) * g ^ (p - 2)] (mod p\<^sup>2)" . moreover have "[1 + p * (p - 1) * g ^ (p - 2) \ 1] (mod p\<^sup>2)" proof assume "[1 + p * (p - 1) * g ^ (p - 2) = 1] (mod p\<^sup>2)" hence "[1 + p * (p - 1) * g ^ (p - 2) - 1 = 1 - 1] (mod p\<^sup>2)" by (intro cong_diff_nat) auto hence "p * p dvd p * ((p - 1) * g ^ (p - 2))" by (auto simp: cong_0_iff power2_eq_square) hence "p dvd (p - 1) * g ^ (p - 2)" using p by simp hence "p dvd g ^ (p - 2)" using p dvd_imp_le[of p "p - Suc 0"] prime_gt_1_nat[of p] by (auto simp: prime_dvd_mult_iff) also have "\ dvd g ^ (p - 1)" by (intro le_imp_power_dvd) auto finally have "[g ^ (p - 1) = 0] (mod p)" by (simp add: cong_0_iff) hence "[0 = g ^ (p - 1)] (mod p)" by (simp add: cong_sym_eq) also from \residue_primroot p g\ have "[g ^ (p - 1) = 1] (mod p)" using p by (auto simp: residue_primroot_def ord_divides' totient_prime) finally have "[0 = 1] (mod p)" . thus False using prime_gt_1_nat[of p] p by (simp add: cong_def) qed ultimately have "[g' ^ (p - 1) \ 1] (mod p\<^sup>2)" by (simp add: cong_def) thus "\g. residue_primroot p g \ [g ^ (p - 1) \ 1] (mod p\<^sup>2)" using g' by blast qed (use g in auto) thus ?thesis using residue_primroot_prime_lift_iff[OF assms] that by blast qed subsection \Carmichael's function\ text \ Carmichael's function $\lambda(n)$ gives the LCM of the orders of all elements in the residue ring modulo $n$ -- or, equivalently, the maximum order, as we will show later. Algebraically speaking, it is the exponent of the multiplicative group $(\mathbb{Z}/n\mathbb{Z})^*$. It is not to be confused with Liouville's function, which is also denoted by $\lambda(n)$. \ definition Carmichael where "Carmichael n = (LCM a \ totatives n. ord n a)" lemma Carmichael_0 [simp]: "Carmichael 0 = 1" by (simp add: Carmichael_def) lemma Carmichael_1 [simp]: "Carmichael 1 = 1" by (simp add: Carmichael_def) lemma Carmichael_Suc_0 [simp]: "Carmichael (Suc 0) = 1" by (simp add: Carmichael_def) lemma ord_dvd_Carmichael: assumes "n > 1" "coprime n k" shows "ord n k dvd Carmichael n" proof - have "k mod n \ totatives n" using assms by (auto simp: totatives_def coprime_commute intro!: Nat.gr0I) hence "ord n (k mod n) dvd Carmichael n" by (simp add: Carmichael_def del: ord_mod) thus ?thesis by simp qed lemma Carmichael_divides: assumes "Carmichael n dvd k" "coprime n a" shows "[a ^ k = 1] (mod n)" proof (cases "n < 2 \ a = 1") case False hence "ord n a dvd Carmichael n" using False assms by (intro ord_dvd_Carmichael) auto also have "\ dvd k" by fact finally have "ord n a dvd k" . thus ?thesis using ord_divides by auto next case True then consider "a = 1" | "n = 0" | "n = 1" by force thus ?thesis using assms by cases auto qed lemma Carmichael_dvd_totient: "Carmichael n dvd totient n" unfolding Carmichael_def proof (intro Lcm_least, safe) fix a assume "a \ totatives n" hence "[a ^ totient n = 1] (mod n)" by (intro euler_theorem) (auto simp: totatives_def) thus "ord n a dvd totient n" using ord_divides by blast qed lemma Carmichael_dvd_mono_coprime: assumes "coprime m n" "m > 1" "n > 1" shows "Carmichael m dvd Carmichael (m * n)" unfolding Carmichael_def[of m] proof (intro Lcm_least, safe) fix x assume x: "x \ totatives m" from assms have "totatives n \ {}" by simp then obtain y where y: "y \ totatives n" by blast from binary_chinese_remainder_nat[OF assms(1), of x y] obtain z where z: "[z = x] (mod m)" "[z = y] (mod n)" by blast have z': "coprime z n" "coprime z m" by (rule cong_imp_coprime; use x y z in \force simp: totatives_def cong_sym_eq\)+ from z have "ord m x = ord m z" by (intro ord_cong) (auto simp: cong_sym_eq) also have "ord m z dvd ord (m * n) z" using assms by (auto simp: ord_modulus_mult_coprime) also from z' assms have "\ dvd Carmichael (m * n)" by (intro ord_dvd_Carmichael) (auto simp: coprime_commute intro!:one_less_mult) finally show "ord m x dvd Carmichael (m * n)" . qed text \ $\lambda$ distributes over the product of coprime numbers similarly to $\varphi$, but with LCM instead of multiplication: \ lemma Carmichael_mult_coprime: assumes "coprime m n" shows "Carmichael (m * n) = lcm (Carmichael m) (Carmichael n)" proof (cases "m \ 1 \ n \ 1") case True hence "m = 0 \ n = 0 \ m = 1 \ n = 1" by force thus ?thesis using assms by auto next case False show ?thesis proof (rule dvd_antisym) show "Carmichael (m * n) dvd lcm (Carmichael m) (Carmichael n)" unfolding Carmichael_def[of "m * n"] proof (intro Lcm_least, safe) fix x assume x: "x \ totatives (m * n)" have "ord (m * n) x = lcm (ord m x) (ord n x)" using assms x by (subst ord_modulus_mult_coprime) (auto simp: coprime_commute totatives_def) also have "\ dvd lcm (Carmichael m) (Carmichael n)" using False x by (intro lcm_mono ord_dvd_Carmichael) (auto simp: totatives_def coprime_commute) finally show "ord (m * n) x dvd \" . qed next show "lcm (Carmichael m) (Carmichael n) dvd Carmichael (m * n)" using Carmichael_dvd_mono_coprime[of m n] Carmichael_dvd_mono_coprime[of n m] assms False by (auto intro!: lcm_least simp: coprime_commute mult.commute) qed qed lemma Carmichael_pos [simp, intro]: "Carmichael n > 0" by (auto simp: Carmichael_def ord_eq_0 totatives_def coprime_commute intro!: Nat.gr0I) lemma Carmichael_nonzero [simp]: "Carmichael n \ 0" by simp lemma power_Carmichael_eq_1: assumes "n > 1" "coprime n x" shows "[x ^ Carmichael n = 1] (mod n)" using ord_dvd_Carmichael[of n x] assms by (auto simp: ord_divides') lemma Carmichael_2 [simp]: "Carmichael 2 = 1" using Carmichael_dvd_totient[of 2] by simp lemma Carmichael_4 [simp]: "Carmichael 4 = 2" proof - have "Carmichael 4 dvd 2" using Carmichael_dvd_totient[of 4] by simp hence "Carmichael 4 \ 2" by (rule dvd_imp_le) auto moreover have "Carmichael 4 \ 1" using power_Carmichael_eq_1[of "4::nat" 3] unfolding coprime_iff_gcd_eq_1 by (auto simp: gcd_non_0_nat cong_def) ultimately show ?thesis using Carmichael_pos[of 4] by linarith qed lemma residue_primroot_Carmichael: assumes "residue_primroot n g" shows "Carmichael n = totient n" proof (cases "n = 1") case False show ?thesis proof (intro dvd_antisym Carmichael_dvd_totient) have "ord n g dvd Carmichael n" using assms False by (intro ord_dvd_Carmichael) (auto simp: residue_primroot_def) thus "totient n dvd Carmichael n" using assms by (auto simp: residue_primroot_def) qed qed auto lemma Carmichael_odd_prime_power: assumes "prime p" "odd p" "k > 0" shows "Carmichael (p ^ k) = p ^ (k - 1) * (p - 1)" proof - from assms obtain g where "residue_primroot (p ^ k) g" using residue_primroot_odd_prime_power_exists[of p] assms by metis hence "Carmichael (p ^ k) = totient (p ^ k)" by (intro residue_primroot_Carmichael[of "p ^ k" g]) auto with assms show ?thesis by (simp add: totient_prime_power) qed lemma Carmichael_prime: assumes "prime p" shows "Carmichael p = p - 1" proof (cases "even p") case True with assms have "p = 2" using primes_dvd_imp_eq two_is_prime_nat by blast thus ?thesis by simp next case False with Carmichael_odd_prime_power[of p 1] assms show ?thesis by simp qed lemma Carmichael_twopow_ge_8: assumes "k \ 3" shows "Carmichael (2 ^ k) = 2 ^ (k - 2)" proof (intro dvd_antisym) have "2 ^ (k - 2) = ord (2 ^ k) (3 :: nat)" using ord_twopow_3_5[of k 3] assms by simp also have "\ dvd Carmichael (2 ^ k)" using assms one_less_power[of "2::nat" k] by (intro ord_dvd_Carmichael) auto finally show "2 ^ (k - 2) dvd \" . next show "Carmichael (2 ^ k) dvd 2 ^ (k - 2)" unfolding Carmichael_def proof (intro Lcm_least, safe) fix x assume "x \ totatives (2 ^ k)" hence "odd x" by (auto simp: totatives_def) hence "[x ^ 2 ^ (k - 2) = 1] (mod 2 ^ k)" using assms ord_twopow_aux[of k x] by auto thus "ord (2 ^ k) x dvd 2 ^ (k - 2)" by (simp add: ord_divides') qed qed lemma Carmichael_twopow: "Carmichael (2 ^ k) = (if k \ 2 then 2 ^ (k - 1) else 2 ^ (k - 2))" proof - have "k = 0 \ k = 1 \ k = 2 \ k \ 3" by linarith thus ?thesis by (auto simp: Carmichael_twopow_ge_8) qed lemma Carmichael_prime_power: assumes "prime p" "k > 0" shows "Carmichael (p ^ k) = (if p = 2 \ k > 2 then 2 ^ (k - 2) else p ^ (k - 1) * (p - 1))" proof (cases "p = 2") case True thus ?thesis by (simp add: Carmichael_twopow) next case False with assms have "odd p" "p > 2" using prime_odd_nat[of p] prime_gt_1_nat[of p] by auto thus ?thesis using assms Carmichael_odd_prime_power[of p k] by simp qed lemma Carmichael_prod_coprime: assumes "finite A" "\i j. i \ A \ j \ A \ i \ j \ coprime (f i) (f j)" shows "Carmichael (\i\A. f i) = (LCM i\A. Carmichael (f i))" using assms by (induction A rule: finite_induct) (simp, simp, subst Carmichael_mult_coprime[OF prod_coprime_right], auto) text \ Since $\lambda$ distributes over coprime factors and we know the value of $\lambda(p^k)$ for prime $p$, we can now give a closed formula for $\lambda(n)$ in terms of the prime factorisation of $n$: \ theorem Carmichael_closed_formula: "Carmichael n = (LCM p\prime_factors n. let k = multiplicity p n in if p = 2 \ k > 2 then 2 ^ (k - 2) else p ^ (k - 1) * (p - 1))" (is "_ = Lcm ?A") proof (cases "n = 0") case False hence "n = (\p\prime_factors n. p ^ multiplicity p n)" using prime_factorization_nat by blast also have "Carmichael \ = (LCM p\prime_factors n. Carmichael (p ^ multiplicity p n))" by (subst Carmichael_prod_coprime) (auto simp: in_prime_factors_iff primes_coprime) also have "(\p. Carmichael (p ^ multiplicity p n)) ` prime_factors n = ?A" by (intro image_cong) (auto simp: Let_def Carmichael_prime_power prime_factors_multiplicity) finally show ?thesis . qed auto corollary even_Carmichael: assumes "n > 2" shows "even (Carmichael n)" proof (cases "\k. n = 2 ^ k") case True then obtain k where [simp]: "n = 2 ^ k" by auto from assms have "k \ 0" "k \ 1" by (auto intro!: Nat.gr0I) hence "k \ 2" by auto thus ?thesis by (auto simp: Carmichael_twopow) next case False from assms have "n \ 0" by auto from False have "\p\prime_factors n. p \ 2" using assms Ex_other_prime_factor[of n 2] by auto - from divide_out_primepow_ex[OF \n \ 0\ this] guess p k n' . note p = this + from divide_out_primepow_ex[OF \n \ 0\ this] + obtain p k n' + where p: + "p \ 2" + "prime p" + "p dvd n" + "\ p dvd n'" + "0 < k" + "n = p ^ k * n'" . from p have coprime: "coprime (p ^ k) n'" using p prime_imp_coprime by auto have "odd p" using p primes_dvd_imp_eq[of 2 p] by auto have "even (Carmichael (p ^ k))" using p \odd p\ by (auto simp: Carmichael_prime_power) with p coprime show ?thesis by (auto simp: Carmichael_mult_coprime intro!: dvd_lcmI1) qed lemma eval_Carmichael: assumes "prime_factorization n = A" shows "Carmichael n = (LCM p \ set_mset A. let k = count A p in if p = 2 \ k > 2 then 2 ^ (k - 2) else p ^ (k - 1) * (p - 1))" unfolding assms [symmetric] Carmichael_closed_formula by (intro arg_cong[where f = Lcm] image_cong) (auto simp: Let_def count_prime_factorization) text \ Any residue ring always contains a $\lambda$-root, i.\,e.\ an element whose order is $\lambda(n)$. \ theorem Carmichael_root_exists: assumes "n > (0::nat)" obtains g where "g \ totatives n" and "ord n g = Carmichael n" proof (cases "n = 1") case True thus ?thesis by (intro that[of 1]) (auto simp: totatives_def) next case False have primepow: "\g. coprime (p ^ k) g \ ord (p ^ k) g = Carmichael (p ^ k)" if pk: "prime p" "k > 0" for p k proof (cases "p = 2") case [simp]: True from \k > 0\ consider "k = 1" | "k = 2" | "k \ 3" by force thus ?thesis proof cases assume "k = 1" thus ?thesis by (intro exI[of _ 1]) auto next assume [simp]: "k = 2" have "coprime 4 (3::nat)" by (auto simp: coprime_iff_gcd_eq_1 gcd_non_0_nat) thus ?thesis by (intro exI[of _ 3]) auto next assume k: "k \ 3" have "coprime (2 ^ k :: nat) 3" by auto thus ?thesis using k by (intro exI[of _ 3]) (auto simp: ord_twopow_3_5 Carmichael_twopow) qed next case False hence "odd p" using \prime p\ using primes_dvd_imp_eq two_is_prime_nat by blast then obtain g where "residue_primroot (p ^ k) g" using residue_primroot_odd_prime_power_exists[of p] pk by metis thus ?thesis using False pk by (intro exI[of _ g]) (auto simp: Carmichael_prime_power residue_primroot_def totient_prime_power) qed define P where "P = prime_factors n" define k where "k = (\p. multiplicity p n)" have "\p\P. \g. coprime (p ^ k p) g \ ord (p ^ k p) g = Carmichael (p ^ k p)" using primepow by (auto simp: P_def k_def prime_factors_multiplicity) hence "\g. \p\P. coprime (p ^ k p) (g p) \ ord (p ^ k p) (g p) = Carmichael (p ^ k p)" by (subst (asm) bchoice_iff) then obtain g where g: "\p. p \ P \ coprime (p ^ k p) (g p)" "\p. p \ P \ ord (p ^ k p) (g p) = Carmichael (p ^ k p)" by metis have "\x. \i\P. [x = g i] (mod i ^ k i)" by (intro chinese_remainder_nat) (auto simp: P_def k_def in_prime_factors_iff primes_coprime) then obtain x where x: "\p. p \ P \ [x = g p] (mod p ^ k p)" by metis have "n = (\p\P. p ^ k p)" using assms unfolding P_def k_def by (rule prime_factorization_nat) also have "ord \ x = (LCM p\P. ord (p ^ k p) x)" by (intro ord_modulus_prod_coprime) (auto simp: P_def in_prime_factors_iff primes_coprime) also have "(\p. ord (p ^ k p) x) ` P = (\p. ord (p ^ k p) (g p)) ` P" by (intro image_cong ord_cong x) auto also have "\ = (\p. Carmichael (p ^ k p)) ` P" by (intro image_cong g) auto also have "Lcm \ = Carmichael (\p\P. p ^ k p)" by (intro Carmichael_prod_coprime [symmetric]) (auto simp: P_def in_prime_factors_iff primes_coprime) also have "(\p\P. p ^ k p) = n" using assms unfolding P_def k_def by (rule prime_factorization_nat [symmetric]) finally have "ord n x = Carmichael n" . moreover from this have "coprime n x" by (cases "coprime n x") (auto split: if_splits) ultimately show ?thesis using assms False by (intro that[of "x mod n"]) (auto simp: totatives_def coprime_commute coprime_absorb_left intro!: Nat.gr0I) qed text \ This also means that the Carmichael number is not only the LCM of the orders of the elements of the residue ring, but indeed the maximum of the orders. \ lemma Carmichael_altdef: "Carmichael n = (if n = 0 then 1 else Max (ord n ` totatives n))" proof (cases "n = 0") case False have "Carmichael n = Max (ord n ` totatives n)" proof (intro antisym Max.boundedI Max.coboundedI) fix k assume k: "k \ ord n ` totatives n" thus "k \ Carmichael n" proof (cases "n = 1") case False with \n \ 0\ have "n > 1" by linarith thus ?thesis using k \n \ 0\ by (intro dvd_imp_le) (auto intro!: ord_dvd_Carmichael simp: totatives_def coprime_commute) qed auto next obtain g where "g \ totatives n" and "ord n g = Carmichael n" using Carmichael_root_exists[of n] \n \ 0\ by auto thus "Carmichael n \ ord n ` totatives n" by force qed (use \n \ 0\ in auto) thus ?thesis using False by simp qed auto subsection \Existence of primitive roots for general moduli\ text \ We now related Carmichael's function to the existence of primitive roots and, in the end, use this to show precisely which moduli have primitive roots and which do not. The first criterion for the existence of a primitive root is this: A primitive root modulo $n$ exists iff $\lambda(n) = \varphi(n)$. \ lemma Carmichael_eq_totient_imp_primroot: assumes "n > 0" and "Carmichael n = totient n" shows "\g. residue_primroot n g" proof - from \n > 0\ obtain g where "g \ totatives n" and "ord n g = Carmichael n" using Carmichael_root_exists[of n] by metis with assms show ?thesis by (auto simp: residue_primroot_def totatives_def coprime_commute) qed theorem residue_primroot_iff_Carmichael: "(\g. residue_primroot n g) \ Carmichael n = totient n \ n > 0" proof safe fix g assume g: "residue_primroot n g" thus "n > 0" by (auto simp: residue_primroot_def) have "coprime n g" by (rule ccontr) (use g in \auto simp: residue_primroot_def\) show "Carmichael n = totient n" proof (cases "n = 1") case False with \n > 0\ have "n > 1" by auto with \coprime n g\ have "ord n g dvd Carmichael n" by (intro ord_dvd_Carmichael) auto thus ?thesis using g by (intro dvd_antisym Carmichael_dvd_totient) (auto simp: residue_primroot_def) qed auto qed (use Carmichael_eq_totient_imp_primroot[of n] in auto) text \ Any primitive root modulo $mn$ for coprime $m$, $n$ is also a primitive root modulo $m$ and $n$. The converse does not hold in general. \ lemma residue_primroot_modulus_mult_coprimeD: assumes "coprime m n" and "residue_primroot (m * n) g" shows "residue_primroot m g" "residue_primroot n g" proof - have *: "m > 0" "n > 0" "coprime m g" "coprime n g" "lcm (ord m g) (ord n g) = totient m * totient n" using assms by (auto simp: residue_primroot_def ord_modulus_mult_coprime totient_mult_coprime) define a b where "a = totient m div ord m g" and "b = totient n div ord n g" have ab: "totient m = ord m g * a" "totient n = ord n g * b" using * by (auto simp: a_def b_def order_divides_totient) have "a = 1" "b = 1" "coprime (ord m g) (ord n g)" unfolding coprime_iff_gcd_eq_1 using * by (auto simp: ab lcm_nat_def dvd_div_eq_mult ord_eq_0) with ab and * show "residue_primroot m g" "residue_primroot n g" by (auto simp: residue_primroot_def) qed text \ If a primitive root modulo $mn$ exists for coprime $m, n$, then $\lambda(m)$ and $\lambda(n)$ must also be coprime. This is helpful in establishing that there are no primitive roots modulo $mn$ by showing e.\,g.\ that $\lambda(m)$ and $\lambda(n)$ are both even. \ lemma residue_primroot_modulus_mult_coprime_imp_Carmichael_coprime: assumes "coprime m n" and "residue_primroot (m * n) g" shows "coprime (Carmichael m) (Carmichael n)" proof - from residue_primroot_modulus_mult_coprimeD[OF assms] have *: "residue_primroot m g" "residue_primroot n g" by auto hence [simp]: "Carmichael m = totient m" "Carmichael n = totient n" by (simp_all add: residue_primroot_Carmichael) from * have mn: "m > 0" "n > 0" by (auto simp: residue_primroot_def) from assms have "Carmichael (m * n) = totient (m * n) \ n > 0" using residue_primroot_iff_Carmichael[of "m * n"] by auto with assms have "lcm (totient m) (totient n) = totient m * totient n" by (simp add: Carmichael_mult_coprime totient_mult_coprime) thus ?thesis unfolding coprime_iff_gcd_eq_1 using mn by (simp add: lcm_nat_def dvd_div_eq_mult) qed text \ The following moduli are precisely those that have primitive roots. \ definition cyclic_moduli :: "nat set" where "cyclic_moduli = {1, 2, 4} \ {p ^ k |p k. prime p \ odd p \ k > 0} \ {2 * p ^ k |p k. prime p \ odd p \ k > 0}" theorem residue_primroot_iff_in_cyclic_moduli: "(\g. residue_primroot m g) \ m \ cyclic_moduli" proof - have "(\g. residue_primroot m g)" if "m \ cyclic_moduli" using that unfolding cyclic_moduli_def by (intro Carmichael_eq_totient_imp_primroot) (auto dest: prime_gt_0_nat simp: Carmichael_prime_power totient_prime_power Carmichael_mult_coprime totient_mult_coprime) moreover have "\(\g. residue_primroot m g)" if "m \ cyclic_moduli" proof (cases "m = 0") case False with that have [simp]: "m > 0" "m \ 1" by (auto simp: cyclic_moduli_def) show ?thesis proof (cases "\k. m = 2 ^ k") case True then obtain k where [simp]: "m = 2 ^ k" by auto with that have "k \ 0 \ k \ 1 \ k \ 2" by (auto simp: cyclic_moduli_def) hence "k \ 3" by auto thus ?thesis by (subst residue_primroot_iff_Carmichael) (auto simp: Carmichael_twopow totient_prime_power) next case False hence "\p\prime_factors m. p \ 2" using Ex_other_prime_factor[of m 2] by auto from divide_out_primepow_ex[OF \m \ 0\ this] obtain p k m' where p: "p \ 2" "prime p" "p dvd m" "\p dvd m'" "0 < k" "m = p ^ k * m'" by metis have "odd p" using p primes_dvd_imp_eq[of 2 p] by auto from p have coprime: "coprime (p ^ k) m'" using p prime_imp_coprime by auto have "m \ cyclic_moduli" if "m' = 1" using that p \odd p\ by (auto simp: cyclic_moduli_def) moreover have "m \ cyclic_moduli" if "m' = 2" proof - have "m = 2 * p ^ k" using p that by (simp add: mult.commute) thus "m \ cyclic_moduli" using p \odd p\ unfolding cyclic_moduli_def by blast qed moreover have "m' \ 0" using p by (intro notI) auto ultimately have "m' \ 0 \m' \ 1 \ m' \ 2" using that by auto hence "m' > 2" by linarith show ?thesis proof assume "\g. residue_primroot m g" with coprime p have coprime': "coprime (p - 1) (Carmichael m')" using residue_primroot_modulus_mult_coprime_imp_Carmichael_coprime[OF coprime] by (auto simp: Carmichael_prime_power) moreover have "even (p - 1)" "even (Carmichael m')" using \m' > 2\ \odd p\ by (auto intro!: even_Carmichael) ultimately show False by force qed qed qed auto ultimately show ?thesis by metis qed lemma residue_primroot_is_generator: assumes "m > 1" and "residue_primroot m g" shows "bij_betw (\i. g ^ i mod m) {..i. g ^ i mod m) {..i. g ^ i mod m) {..i. g ^ i mod m) ` {..i. g ^ i mod m) ` {..i. g ^ i mod m) ` {.. totatives m" using \m > 1\ \coprime m g\ power_in_totatives[of m g] by blast qed auto qed text \ Given one primitive root \g\, all the primitive roots are powers $g^i$ for $1\leq i \leq \varphi(n)$ with $\text{gcd}(i, \varphi(n)) = 1$. \ theorem residue_primroot_bij_betw_primroots: assumes "m > 1" and "residue_primroot m g" shows "bij_betw (\i. g ^ i mod m) (totatives (totient m)) {g\totatives m. residue_primroot m g}" proof (cases "m = 2") case [simp]: True have [simp]: "totatives 2 = {1}" by (auto simp: totatives_def elim!: oddE) from assms have "odd g" by (auto simp: residue_primroot_def) hence pow_eq: "(\i. g ^ i mod m) = (\_. 1)" by (auto simp: fun_eq_iff mod_2_eq_odd) have "{g \ totatives m. residue_primroot m g} = {1}" by (auto simp: residue_primroot_def) thus ?thesis using pow_eq by (auto simp: bij_betw_def) next case False thus ?thesis unfolding bij_betw_def proof safe from assms False have "m > 2" by auto from assms \m > 2\ have "totient m > 1" by (intro totient_gt_1) auto from assms have [simp]: "ord m g = totient m" by (simp add: residue_primroot_def) from assms have "coprime m g" by (simp add: residue_primroot_def) hence "inj_on (\i. g ^ i mod m) {..i. g ^ i mod m) (totatives (totient m))" by (rule inj_on_subset) (use assms \totient m > 1\ in \auto simp: totatives_less residue_primroot_def\) { fix i assume i: "i \ totatives (totient m)" from \coprime m g\ and \m > 2\ show "g ^ i mod m \ totatives m" by (intro power_in_totatives) auto show "residue_primroot m (g ^ i mod m)" using i \m > 2\ \coprime m g\ by (auto simp: residue_primroot_def coprime_commute ord_power totatives_def) } { fix x assume x: "x \ totatives m" "residue_primroot m x" then obtain i where i: "i < totient m" "x = (g ^ i mod m)" using assms residue_primroot_is_generator[of m g] by (auto simp: bij_betw_def) from i x \m > 2\ have "i > 0" by (intro Nat.gr0I) (auto simp: residue_primroot_1_iff) have "totient m div gcd i (totient m) = totient m" using x i \coprime m g\ by (auto simp add: residue_primroot_def ord_power) hence "coprime i (totient m)" unfolding coprime_iff_gcd_eq_1 using assms by (subst (asm) dvd_div_eq_mult) auto with i \i > 0\ have "i \ totatives (totient m)" by (auto simp: totatives_def) thus "x \ (\i. g ^ i mod m) ` totatives (totient m)" using i by auto } qed qed text \ It follows from the above statement that any residue ring modulo \n\ that \<^emph>\has\ primitive roots has exactly $\varphi(\varphi(n))$ of them. \ corollary card_residue_primroots: assumes "\g. residue_primroot m g" shows "card {g\totatives m. residue_primroot m g} = totient (totient m)" proof (cases "m = 1") case [simp]: True have "{g \ totatives m. residue_primroot m g} = {1}" by (auto simp: residue_primroot_def) thus ?thesis by simp next case False from assms obtain g where g: "residue_primroot m g" by auto hence "m \ 0" by (intro notI) auto with \m \ 1\ have "m > 1" by linarith from this g have "bij_betw (\i. g ^ i mod m) (totatives (totient m)) {g\totatives m. residue_primroot m g}" by (rule residue_primroot_bij_betw_primroots) hence "card (totatives (totient m)) = card {g\totatives m. residue_primroot m g}" by (rule bij_betw_same_card) thus ?thesis by (simp add: totient_def) qed corollary card_residue_primroots': "card {g\totatives m. residue_primroot m g} = (if m \ cyclic_moduli then totient (totient m) else 0)" by (simp add: residue_primroot_iff_in_cyclic_moduli [symmetric] card_residue_primroots) text \ As an example, we evaluate $\lambda(122200)$ using the prime factorisation. \ lemma "Carmichael 122200 = 1380" proof - have "prime_factorization (2^3 * 5^2 * 13 * 47) = {#2, 2, 2, 5, 5, 13, 47::nat#}" by (intro prime_factorization_eqI) auto from eval_Carmichael[OF this] show "Carmichael 122200 = 1380" by (simp add: lcm_nat_def gcd_non_0_nat) qed (* TODO: definition of index ("discrete logarithm") *) end \ No newline at end of file diff --git a/src/HOL/Probability/Discrete_Topology.thy b/src/HOL/Probability/Discrete_Topology.thy --- a/src/HOL/Probability/Discrete_Topology.thy +++ b/src/HOL/Probability/Discrete_Topology.thy @@ -1,64 +1,64 @@ (* Title: HOL/Probability/Discrete_Topology.thy Author: Fabian Immler, TU München *) theory Discrete_Topology imports "HOL-Analysis.Analysis" begin text \Copy of discrete types with discrete topology. This space is polish.\ typedef 'a discrete = "UNIV::'a set" morphisms of_discrete discrete .. instantiation discrete :: (type) metric_space begin definition dist_discrete :: "'a discrete \ 'a discrete \ real" where "dist_discrete n m = (if n = m then 0 else 1)" definition uniformity_discrete :: "('a discrete \ 'a discrete) filter" where "(uniformity::('a discrete \ 'a discrete) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition "open_discrete" :: "'a discrete set \ bool" where "(open::'a discrete set \ bool) U \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof qed (auto simp: uniformity_discrete_def open_discrete_def dist_discrete_def intro: exI[where x=1]) end lemma open_discrete: "open (S :: 'a discrete set)" unfolding open_dist dist_discrete_def by (auto intro!: exI[of _ "1 / 2"]) instance discrete :: (type) complete_space proof - fix X::"nat\'a discrete" assume "Cauchy X" - hence "\n. \m\n. X n = X m" + fix X::"nat\'a discrete" + assume "Cauchy X" + then obtain n where "\m\n. X n = X m" by (force simp: dist_discrete_def Cauchy_def split: if_split_asm dest:spec[where x=1]) - then guess n .. thus "convergent X" by (intro convergentI[where L="X n"] tendstoI eventually_sequentiallyI[of n]) (simp add: dist_discrete_def) qed instance discrete :: (countable) countable proof have "inj (\c::'a discrete. to_nat (of_discrete c))" by (simp add: inj_on_def of_discrete_inject) thus "\f::'a discrete \ nat. inj f" by blast qed instance discrete :: (countable) second_countable_topology proof let ?B = "range (\n::'a discrete. {n})" have "\S. generate_topology ?B (\x\S. {x})" by (intro generate_topology_Union) (auto intro: generate_topology.intros) then have "open = generate_topology ?B" by (auto intro!: ext simp: open_discrete) moreover have "countable ?B" by simp ultimately show "\B::'a discrete set set. countable B \ open = generate_topology B" by blast qed instance discrete :: (countable) polish_space .. end diff --git a/src/HOL/Probability/Fin_Map.thy b/src/HOL/Probability/Fin_Map.thy --- a/src/HOL/Probability/Fin_Map.thy +++ b/src/HOL/Probability/Fin_Map.thy @@ -1,1335 +1,1351 @@ (* Title: HOL/Probability/Fin_Map.thy Author: Fabian Immler, TU München *) section \Finite Maps\ theory Fin_Map imports "HOL-Analysis.Finite_Product_Measure" "HOL-Library.Finite_Map" begin text \The \<^type>\fmap\ type can be instantiated to \<^class>\polish_space\, needed for the proof of projective limit. \<^const>\extensional\ functions are used for the representation in order to stay close to the developments of (finite) products \<^const>\Pi\<^sub>E\ and their sigma-algebra \<^const>\Pi\<^sub>M\.\ type_notation fmap ("(_ \\<^sub>F /_)" [22, 21] 21) unbundle fmap.lifting subsection \Domain and Application\ lift_definition domain::"('i \\<^sub>F 'a) \ 'i set" is dom . lemma finite_domain[simp, intro]: "finite (domain P)" by transfer simp lift_definition proj :: "('i \\<^sub>F 'a) \ 'i \ 'a" ("'((_)')\<^sub>F" [0] 1000) is "\f x. if x \ dom f then the (f x) else undefined" . declare [[coercion proj]] lemma extensional_proj[simp, intro]: "(P)\<^sub>F \ extensional (domain P)" by transfer (auto simp: extensional_def) lemma proj_undefined[simp, intro]: "i \ domain P \ P i = undefined" using extensional_proj[of P] unfolding extensional_def by auto lemma finmap_eq_iff: "P = Q \ (domain P = domain Q \ (\i\domain P. P i = Q i))" apply transfer apply (safe intro!: ext) subgoal for P Q x by (cases "x \ dom P"; cases "P x") (auto dest!: bspec[where x=x]) done subsection \Constructor of Finite Maps\ lift_definition finmap_of::"'i set \ ('i \ 'a) \ ('i \\<^sub>F 'a)" is "\I f x. if x \ I \ finite I then Some (f x) else None" by (simp add: dom_def) lemma proj_finmap_of[simp]: assumes "finite inds" shows "(finmap_of inds f)\<^sub>F = restrict f inds" using assms by transfer force lemma domain_finmap_of[simp]: assumes "finite inds" shows "domain (finmap_of inds f) = inds" using assms by transfer (auto split: if_splits) lemma finmap_of_eq_iff[simp]: assumes "finite i" "finite j" shows "finmap_of i m = finmap_of j n \ i = j \ (\k\i. m k= n k)" using assms by (auto simp: finmap_eq_iff) lemma finmap_of_inj_on_extensional_finite: assumes "finite K" assumes "S \ extensional K" shows "inj_on (finmap_of K) S" proof (rule inj_onI) fix x y::"'a \ 'b" assume "finmap_of K x = finmap_of K y" hence "(finmap_of K x)\<^sub>F = (finmap_of K y)\<^sub>F" by simp moreover assume "x \ S" "y \ S" hence "x \ extensional K" "y \ extensional K" using assms by auto ultimately show "x = y" using assms by (simp add: extensional_restrict) qed subsection \Product set of Finite Maps\ text \This is \<^term>\Pi\ for Finite Maps, most of this is copied\ definition Pi' :: "'i set \ ('i \ 'a set) \ ('i \\<^sub>F 'a) set" where "Pi' I A = { P. domain P = I \ (\i. i \ I \ (P)\<^sub>F i \ A i) } " syntax "_Pi'" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\'' _\_./ _)" 10) translations "\' x\A. B" == "CONST Pi' A (\x. B)" subsubsection\Basic Properties of \<^term>\Pi'\\ lemma Pi'_I[intro!]: "domain f = A \ (\x. x \ A \ f x \ B x) \ f \ Pi' A B" by (simp add: Pi'_def) lemma Pi'_I'[simp]: "domain f = A \ (\x. x \ A \ f x \ B x) \ f \ Pi' A B" by (simp add:Pi'_def) lemma Pi'_mem: "f\ Pi' A B \ x \ A \ f x \ B x" by (simp add: Pi'_def) lemma Pi'_iff: "f \ Pi' I X \ domain f = I \ (\i\I. f i \ X i)" unfolding Pi'_def by auto lemma Pi'E [elim]: "f \ Pi' A B \ (f x \ B x \ domain f = A \ Q) \ (x \ A \ Q) \ Q" by(auto simp: Pi'_def) lemma in_Pi'_cong: "domain f = domain g \ (\ w. w \ A \ f w = g w) \ f \ Pi' A B \ g \ Pi' A B" by (auto simp: Pi'_def) lemma Pi'_eq_empty[simp]: assumes "finite A" shows "(Pi' A B) = {} \ (\x\A. B x = {})" using assms apply (simp add: Pi'_def, auto) apply (drule_tac x = "finmap_of A (\u. SOME y. y \ B u)" in spec, auto) apply (cut_tac P= "%y. y \ B i" in some_eq_ex, auto) done lemma Pi'_mono: "(\x. x \ A \ B x \ C x) \ Pi' A B \ Pi' A C" by (auto simp: Pi'_def) lemma Pi_Pi': "finite A \ (Pi\<^sub>E A B) = proj ` Pi' A B" apply (auto simp: Pi'_def Pi_def extensional_def) apply (rule_tac x = "finmap_of A (restrict x A)" in image_eqI) apply auto done subsection \Topological Space of Finite Maps\ instantiation fmap :: (type, topological_space) topological_space begin definition open_fmap :: "('a \\<^sub>F 'b) set \ bool" where [code del]: "open_fmap = generate_topology {Pi' a b|a b. \i\a. open (b i)}" lemma open_Pi'I: "(\i. i \ I \ open (A i)) \ open (Pi' I A)" by (auto intro: generate_topology.Basis simp: open_fmap_def) instance using topological_space_generate_topology by intro_classes (auto simp: open_fmap_def class.topological_space_def) end lemma open_restricted_space: shows "open {m. P (domain m)}" proof - have "{m. P (domain m)} = (\i \ Collect P. {m. domain m = i})" by auto also have "open \" proof (rule, safe, cases) fix i::"'a set" assume "finite i" hence "{m. domain m = i} = Pi' i (\_. UNIV)" by (auto simp: Pi'_def) also have "open \" by (auto intro: open_Pi'I simp: \finite i\) finally show "open {m. domain m = i}" . next fix i::"'a set" assume "\ finite i" hence "{m. domain m = i} = {}" by auto also have "open \" by simp finally show "open {m. domain m = i}" . qed finally show ?thesis . qed lemma closed_restricted_space: shows "closed {m. P (domain m)}" using open_restricted_space[of "\x. \ P x"] unfolding closed_def by (rule back_subst) auto lemma tendsto_proj: "((\x. x) \ a) F \ ((\x. (x)\<^sub>F i) \ (a)\<^sub>F i) F" unfolding tendsto_def proof safe fix S::"'b set" let ?S = "Pi' (domain a) (\x. if x = i then S else UNIV)" assume "open S" hence "open ?S" by (auto intro!: open_Pi'I) moreover assume "\S. open S \ a \ S \ eventually (\x. x \ S) F" "a i \ S" ultimately have "eventually (\x. x \ ?S) F" by auto thus "eventually (\x. (x)\<^sub>F i \ S) F" by eventually_elim (insert \a i \ S\, force simp: Pi'_iff split: if_split_asm) qed lemma continuous_proj: shows "continuous_on s (\x. (x)\<^sub>F i)" unfolding continuous_on_def by (safe intro!: tendsto_proj tendsto_ident_at) instance fmap :: (type, first_countable_topology) first_countable_topology proof fix x::"'a\\<^sub>F'b" have "\i. \A. countable A \ (\a\A. x i \ a) \ (\a\A. open a) \ (\S. open S \ x i \ S \ (\a\A. a \ S)) \ (\a b. a \ A \ b \ A \ a \ b \ A)" (is "\i. ?th i") proof - fix i from first_countable_basis_Int_stableE[of "x i"] guess A . + fix i from first_countable_basis_Int_stableE[of "x i"] + obtain A where + "countable A" + "\C. C \ A \ (x)\<^sub>F i \ C" + "\C. C \ A \ open C" + "\S. open S \ (x)\<^sub>F i \ S \ \A\A. A \ S" + "\C D. C \ A \ D \ A \ C \ D \ A" + by auto thus "?th i" by (intro exI[where x=A]) simp qed - then guess A unfolding choice_iff .. note A = this + then obtain A + where A: "\i. countable (A i) \ Ball (A i) ((\) ((x)\<^sub>F i)) \ Ball (A i) open \ + (\S. open S \ (x)\<^sub>F i \ S \ (\a\A i. a \ S)) \ (\a b. a \ A i \ b \ A i \ a \ b \ A i)" + by (auto simp: choice_iff) hence open_sub: "\i S. i\domain x \ open (S i) \ x i\(S i) \ (\a\A i. a\(S i))" by auto have A_notempty: "\i. i \ domain x \ A i \ {}" using open_sub[of _ "\_. UNIV"] by auto let ?A = "(\f. Pi' (domain x) f) ` (Pi\<^sub>E (domain x) A)" show "\A::nat \ ('a\\<^sub>F'b) set. (\i. x \ (A i) \ open (A i)) \ (\S. open S \ x \ S \ (\i. A i \ S))" proof (rule first_countableI[of "?A"], safe) show "countable ?A" using A by (simp add: countable_PiE) next fix S::"('a \\<^sub>F 'b) set" assume "open S" "x \ S" thus "\a\?A. a \ S" unfolding open_fmap_def proof (induct rule: generate_topology.induct) case UNIV thus ?case by (auto simp add: ex_in_conv PiE_eq_empty_iff A_notempty) next case (Int a b) then obtain f g where "f \ Pi\<^sub>E (domain x) A" "Pi' (domain x) f \ a" "g \ Pi\<^sub>E (domain x) A" "Pi' (domain x) g \ b" by auto thus ?case using A by (auto simp: Pi'_iff PiE_iff extensional_def Int_stable_def intro!: bexI[where x="\i. f i \ g i"]) next case (UN B) then obtain b where "x \ b" "b \ B" by auto hence "\a\?A. a \ b" using UN by simp thus ?case using \b \ B\ by blast next case (Basis s) then obtain a b where xs: "x\ Pi' a b" "s = Pi' a b" "\i. i\a \ open (b i)" by auto have "\i. \a. (i \ domain x \ open (b i) \ (x)\<^sub>F i \ b i) \ (a\A i \ a \ b i)" using open_sub[of _ b] by auto then obtain b' where "\i. i \ domain x \ open (b i) \ (x)\<^sub>F i \ b i \ (b' i \A i \ b' i \ b i)" unfolding choice_iff by auto with xs have "\i. i \ a \ (b' i \A i \ b' i \ b i)" "Pi' a b' \ Pi' a b" by (auto simp: Pi'_iff intro!: Pi'_mono) thus ?case using xs by (intro bexI[where x="Pi' a b'"]) (auto simp: Pi'_iff intro!: image_eqI[where x="restrict b' (domain x)"]) qed qed (insert A,auto simp: PiE_iff intro!: open_Pi'I) qed subsection \Metric Space of Finite Maps\ (* TODO: Product of uniform spaces and compatibility with metric_spaces! *) instantiation fmap :: (type, metric_space) dist begin definition dist_fmap where "dist P Q = Max (range (\i. dist ((P)\<^sub>F i) ((Q)\<^sub>F i))) + (if domain P = domain Q then 0 else 1)" instance .. end instantiation fmap :: (type, metric_space) uniformity_dist begin definition [code del]: "(uniformity :: (('a, 'b) fmap \ ('a \\<^sub>F 'b)) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" instance by standard (rule uniformity_fmap_def) end declare uniformity_Abort[where 'a="('a \\<^sub>F 'b::metric_space)", code] instantiation fmap :: (type, metric_space) metric_space begin lemma finite_proj_image': "x \ domain P \ finite ((P)\<^sub>F ` S)" by (rule finite_subset[of _ "proj P ` (domain P \ S \ {x})"]) auto lemma finite_proj_image: "finite ((P)\<^sub>F ` S)" by (cases "\x. x \ domain P") (auto intro: finite_proj_image' finite_subset[where B="domain P"]) lemma finite_proj_diag: "finite ((\i. d ((P)\<^sub>F i) ((Q)\<^sub>F i)) ` S)" proof - have "(\i. d ((P)\<^sub>F i) ((Q)\<^sub>F i)) ` S = (\(i, j). d i j) ` ((\i. ((P)\<^sub>F i, (Q)\<^sub>F i)) ` S)" by auto moreover have "((\i. ((P)\<^sub>F i, (Q)\<^sub>F i)) ` S) \ (\i. (P)\<^sub>F i) ` S \ (\i. (Q)\<^sub>F i) ` S" by auto moreover have "finite \" using finite_proj_image[of P S] finite_proj_image[of Q S] by (intro finite_cartesian_product) simp_all ultimately show ?thesis by (simp add: finite_subset) qed lemma dist_le_1_imp_domain_eq: shows "dist P Q < 1 \ domain P = domain Q" by (simp add: dist_fmap_def finite_proj_diag split: if_split_asm) lemma dist_proj: shows "dist ((x)\<^sub>F i) ((y)\<^sub>F i) \ dist x y" proof - have "dist (x i) (y i) \ Max (range (\i. dist (x i) (y i)))" by (simp add: Max_ge_iff finite_proj_diag) also have "\ \ dist x y" by (simp add: dist_fmap_def) finally show ?thesis . qed lemma dist_finmap_lessI: assumes "domain P = domain Q" assumes "0 < e" assumes "\i. i \ domain P \ dist (P i) (Q i) < e" shows "dist P Q < e" proof - have "dist P Q = Max (range (\i. dist (P i) (Q i)))" using assms by (simp add: dist_fmap_def finite_proj_diag) also have "\ < e" proof (subst Max_less_iff, safe) fix i show "dist ((P)\<^sub>F i) ((Q)\<^sub>F i) < e" using assms by (cases "i \ domain P") simp_all qed (simp add: finite_proj_diag) finally show ?thesis . qed instance proof fix S::"('a \\<^sub>F 'b) set" have *: "open S = (\x\S. \e>0. \y. dist y x < e \ y \ S)" (is "_ = ?od") proof assume "open S" thus ?od unfolding open_fmap_def proof (induct rule: generate_topology.induct) case UNIV thus ?case by (auto intro: zero_less_one) next case (Int a b) show ?case proof safe fix x assume x: "x \ a" "x \ b" with Int x obtain e1 e2 where "e1>0" "\y. dist y x < e1 \ y \ a" "e2>0" "\y. dist y x < e2 \ y \ b" by force thus "\e>0. \y. dist y x < e \ y \ a \ b" by (auto intro!: exI[where x="min e1 e2"]) qed next case (UN K) show ?case proof safe fix x X assume "x \ X" and X: "X \ K" with UN obtain e where "e>0" "\y. dist y x < e \ y \ X" by force with X show "\e>0. \y. dist y x < e \ y \ \K" by auto qed next case (Basis s) then obtain a b where s: "s = Pi' a b" and b: "\i. i\a \ open (b i)" by auto show ?case proof safe fix x assume "x \ s" hence [simp]: "finite a" and a_dom: "a = domain x" using s by (auto simp: Pi'_iff) obtain es where es: "\i \ a. es i > 0 \ (\y. dist y (proj x i) < es i \ y \ b i)" using b \x \ s\ by atomize_elim (intro bchoice, auto simp: open_dist s) hence in_b: "\i y. i \ a \ dist y (proj x i) < es i \ y \ b i" by auto show "\e>0. \y. dist y x < e \ y \ s" proof (cases, rule, safe) assume "a \ {}" show "0 < min 1 (Min (es ` a))" using es by (auto simp: \a \ {}\) fix y assume d: "dist y x < min 1 (Min (es ` a))" show "y \ s" unfolding s proof show "domain y = a" using d s \a \ {}\ by (auto simp: dist_le_1_imp_domain_eq a_dom) fix i assume i: "i \ a" hence "dist ((y)\<^sub>F i) ((x)\<^sub>F i) < es i" using d by (auto simp: dist_fmap_def \a \ {}\ intro!: le_less_trans[OF dist_proj]) with i show "y i \ b i" by (rule in_b) qed next assume "\a \ {}" thus "\e>0. \y. dist y x < e \ y \ s" using s \x \ s\ by (auto simp: Pi'_def dist_le_1_imp_domain_eq intro!: exI[where x=1]) qed qed qed next assume "\x\S. \e>0. \y. dist y x < e \ y \ S" then obtain e where e_pos: "\x. x \ S \ e x > 0" and e_in: "\x y . x \ S \ dist y x < e x \ y \ S" unfolding bchoice_iff by auto have S_eq: "S = \{Pi' a b| a b. \x\S. domain x = a \ b = (\i. ball (x i) (e x))}" proof safe fix x assume "x \ S" thus "x \ \{Pi' a b| a b. \x\S. domain x = a \ b = (\i. ball (x i) (e x))}" using e_pos by (auto intro!: exI[where x="Pi' (domain x) (\i. ball (x i) (e x))"]) next fix x y assume "y \ S" moreover assume "x \ (\' i\domain y. ball (y i) (e y))" hence "dist x y < e y" using e_pos \y \ S\ by (auto simp: dist_fmap_def Pi'_iff finite_proj_diag dist_commute) ultimately show "x \ S" by (rule e_in) qed also have "open \" unfolding open_fmap_def by (intro generate_topology.UN) (auto intro: generate_topology.Basis) finally show "open S" . qed show "open S = (\x\S. \\<^sub>F (x', y) in uniformity. x' = x \ y \ S)" unfolding * eventually_uniformity_metric by (simp del: split_paired_All add: dist_fmap_def dist_commute eq_commute) next fix P Q::"'a \\<^sub>F 'b" have Max_eq_iff: "\A m. finite A \ A \ {} \ (Max A = m) = (m \ A \ (\a\A. a \ m))" by (auto intro: Max_in Max_eqI) show "dist P Q = 0 \ P = Q" by (auto simp: finmap_eq_iff dist_fmap_def Max_ge_iff finite_proj_diag Max_eq_iff add_nonneg_eq_0_iff intro!: Max_eqI image_eqI[where x=undefined]) next fix P Q R::"'a \\<^sub>F 'b" let ?dists = "\P Q i. dist ((P)\<^sub>F i) ((Q)\<^sub>F i)" let ?dpq = "?dists P Q" and ?dpr = "?dists P R" and ?dqr = "?dists Q R" let ?dom = "\P Q. (if domain P = domain Q then 0 else 1::real)" have "dist P Q = Max (range ?dpq) + ?dom P Q" by (simp add: dist_fmap_def) also obtain t where "t \ range ?dpq" "t = Max (range ?dpq)" by (simp add: finite_proj_diag) then obtain i where "Max (range ?dpq) = ?dpq i" by auto also have "?dpq i \ ?dpr i + ?dqr i" by (rule dist_triangle2) also have "?dpr i \ Max (range ?dpr)" by (simp add: finite_proj_diag) also have "?dqr i \ Max (range ?dqr)" by (simp add: finite_proj_diag) also have "?dom P Q \ ?dom P R + ?dom Q R" by simp finally show "dist P Q \ dist P R + dist Q R" by (simp add: dist_fmap_def ac_simps) qed end subsection \Complete Space of Finite Maps\ lemma tendsto_finmap: fixes f::"nat \ ('i \\<^sub>F ('a::metric_space))" assumes ind_f: "\n. domain (f n) = domain g" assumes proj_g: "\i. i \ domain g \ (\n. (f n) i) \ g i" shows "f \ g" unfolding tendsto_iff proof safe fix e::real assume "0 < e" let ?dists = "\x i. dist ((f x)\<^sub>F i) ((g)\<^sub>F i)" have "eventually (\x. \i\domain g. ?dists x i < e) sequentially" using finite_domain[of g] proj_g proof induct case (insert i G) with \0 < e\ have "eventually (\x. ?dists x i < e) sequentially" by (auto simp add: tendsto_iff) moreover from insert have "eventually (\x. \i\G. dist ((f x)\<^sub>F i) ((g)\<^sub>F i) < e) sequentially" by simp ultimately show ?case by eventually_elim auto qed simp thus "eventually (\x. dist (f x) g < e) sequentially" by eventually_elim (auto simp add: dist_fmap_def finite_proj_diag ind_f \0 < e\) qed instance fmap :: (type, complete_space) complete_space proof fix P::"nat \ 'a \\<^sub>F 'b" assume "Cauchy P" then obtain Nd where Nd: "\n. n \ Nd \ dist (P n) (P Nd) < 1" by (force simp: Cauchy_altdef2) define d where "d = domain (P Nd)" with Nd have dim: "\n. n \ Nd \ domain (P n) = d" using dist_le_1_imp_domain_eq by auto have [simp]: "finite d" unfolding d_def by simp define p where "p i n = P n i" for i n define q where "q i = lim (p i)" for i define Q where "Q = finmap_of d q" have q: "\i. i \ d \ q i = Q i" by (auto simp add: Q_def Abs_fmap_inverse) { fix i assume "i \ d" have "Cauchy (p i)" unfolding Cauchy_altdef2 p_def proof safe fix e::real assume "0 < e" with \Cauchy P\ obtain N where N: "\n. n\N \ dist (P n) (P N) < min e 1" by (force simp: Cauchy_altdef2 min_def) hence "\n. n \ N \ domain (P n) = domain (P N)" using dist_le_1_imp_domain_eq by auto with dim have dim: "\n. n \ N \ domain (P n) = d" by (metis nat_le_linear) show "\N. \n\N. dist ((P n) i) ((P N) i) < e" proof (safe intro!: exI[where x="N"]) fix n assume "N \ n" have "N \ N" by simp have "dist ((P n) i) ((P N) i) \ dist (P n) (P N)" using dim[OF \N \ n\] dim[OF \N \ N\] \i \ d\ by (auto intro!: dist_proj) also have "\ < e" using N[OF \N \ n\] by simp finally show "dist ((P n) i) ((P N) i) < e" . qed qed hence "convergent (p i)" by (metis Cauchy_convergent_iff) hence "p i \ q i" unfolding q_def convergent_def by (metis limI) } note p = this have "P \ Q" proof (rule metric_LIMSEQ_I) fix e::real assume "0 < e" have "\ni. \i\d. \n\ni i. dist (p i n) (q i) < e" proof (safe intro!: bchoice) fix i assume "i \ d" from p[OF \i \ d\, THEN metric_LIMSEQ_D, OF \0 < e\] show "\no. \n\no. dist (p i n) (q i) < e" . - qed then guess ni .. note ni = this + qed + then obtain ni where ni: "\i\d. \n\ni i. dist (p i n) (q i) < e" .. define N where "N = max Nd (Max (ni ` d))" show "\N. \n\N. dist (P n) Q < e" proof (safe intro!: exI[where x="N"]) fix n assume "N \ n" hence dom: "domain (P n) = d" "domain Q = d" "domain (P n) = domain Q" using dim by (simp_all add: N_def Q_def dim_def Abs_fmap_inverse) show "dist (P n) Q < e" proof (rule dist_finmap_lessI[OF dom(3) \0 < e\]) fix i assume "i \ domain (P n)" hence "ni i \ Max (ni ` d)" using dom by simp also have "\ \ N" by (simp add: N_def) finally show "dist ((P n)\<^sub>F i) ((Q)\<^sub>F i) < e" using ni \i \ domain (P n)\ \N \ n\ dom by (auto simp: p_def q N_def less_imp_le) qed qed qed thus "convergent P" by (auto simp: convergent_def) qed subsection \Second Countable Space of Finite Maps\ instantiation fmap :: (countable, second_countable_topology) second_countable_topology begin definition basis_proj::"'b set set" where "basis_proj = (SOME B. countable B \ topological_basis B)" lemma countable_basis_proj: "countable basis_proj" and basis_proj: "topological_basis basis_proj" unfolding basis_proj_def by (intro is_basis countable_basis)+ definition basis_finmap::"('a \\<^sub>F 'b) set set" where "basis_finmap = {Pi' I S|I S. finite I \ (\i \ I. S i \ basis_proj)}" lemma in_basis_finmapI: assumes "finite I" assumes "\i. i \ I \ S i \ basis_proj" shows "Pi' I S \ basis_finmap" using assms unfolding basis_finmap_def by auto lemma basis_finmap_eq: assumes "basis_proj \ {}" shows "basis_finmap = (\f. Pi' (domain f) (\i. from_nat_into basis_proj ((f)\<^sub>F i))) ` (UNIV::('a \\<^sub>F nat) set)" (is "_ = ?f ` _") unfolding basis_finmap_def proof safe fix I::"'a set" and S::"'a \ 'b set" assume "finite I" "\i\I. S i \ basis_proj" hence "Pi' I S = ?f (finmap_of I (\x. to_nat_on basis_proj (S x)))" by (force simp: Pi'_def countable_basis_proj) thus "Pi' I S \ range ?f" by simp next fix x and f::"'a \\<^sub>F nat" show "\I S. (\' i\domain f. from_nat_into basis_proj ((f)\<^sub>F i)) = Pi' I S \ finite I \ (\i\I. S i \ basis_proj)" using assms by (auto intro: from_nat_into) qed lemma basis_finmap_eq_empty: "basis_proj = {} \ basis_finmap = {Pi' {} undefined}" by (auto simp: Pi'_iff basis_finmap_def) lemma countable_basis_finmap: "countable basis_finmap" by (cases "basis_proj = {}") (auto simp: basis_finmap_eq basis_finmap_eq_empty) lemma finmap_topological_basis: "topological_basis basis_finmap" proof (subst topological_basis_iff, safe) fix B' assume "B' \ basis_finmap" thus "open B'" by (auto intro!: open_Pi'I topological_basis_open[OF basis_proj] simp: topological_basis_def basis_finmap_def Let_def) next fix O'::"('a \\<^sub>F 'b) set" and x assume O': "open O'" "x \ O'" then obtain a where a: "x \ Pi' (domain x) a" "Pi' (domain x) a \ O'" "\i. i\domain x \ open (a i)" unfolding open_fmap_def proof (atomize_elim, induct rule: generate_topology.induct) case (Int a b) let ?p="\a f. x \ Pi' (domain x) f \ Pi' (domain x) f \ a \ (\i. i \ domain x \ open (f i))" from Int obtain f g where "?p a f" "?p b g" by auto thus ?case by (force intro!: exI[where x="\i. f i \ g i"] simp: Pi'_def) next case (UN k) then obtain kk a where "x \ kk" "kk \ k" "x \ Pi' (domain x) a" "Pi' (domain x) a \ kk" "\i. i\domain x \ open (a i)" by force thus ?case by blast qed (auto simp: Pi'_def) have "\B. (\i\domain x. x i \ B i \ B i \ a i \ B i \ basis_proj)" proof (rule bchoice, safe) fix i assume "i \ domain x" hence "open (a i)" "x i \ a i" using a by auto - from topological_basisE[OF basis_proj this] guess b' . + from topological_basisE[OF basis_proj this] obtain b' + where "b' \ basis_proj" "(x)\<^sub>F i \ b'" "b' \ a i" + by blast thus "\y. x i \ y \ y \ a i \ y \ basis_proj" by auto qed - then guess B .. note B = this + then obtain B where B: "\i\domain x. (x)\<^sub>F i \ B i \ B i \ a i \ B i \ basis_proj" + by auto define B' where "B' = Pi' (domain x) (\i. (B i)::'b set)" have "B' \ Pi' (domain x) a" using B by (auto intro!: Pi'_mono simp: B'_def) also note \\ \ O'\ finally show "\B'\basis_finmap. x \ B' \ B' \ O'" using B by (auto intro!: bexI[where x=B'] Pi'_mono in_basis_finmapI simp: B'_def) qed lemma range_enum_basis_finmap_imp_open: assumes "x \ basis_finmap" shows "open x" using finmap_topological_basis assms by (auto simp: topological_basis_def) instance proof qed (blast intro: finmap_topological_basis countable_basis_finmap topological_basis_imp_subbasis) end subsection \Polish Space of Finite Maps\ instance fmap :: (countable, polish_space) polish_space proof qed subsection \Product Measurable Space of Finite Maps\ definition "PiF I M \ sigma (\J \ I. (\' j\J. space (M j))) {(\' j\J. X j) |X J. J \ I \ X \ (\ j\J. sets (M j))}" abbreviation "Pi\<^sub>F I M \ PiF I M" syntax "_PiF" :: "pttrn \ 'i set \ 'a measure \ ('i => 'a) measure" ("(3\\<^sub>F _\_./ _)" 10) translations "\\<^sub>F x\I. M" == "CONST PiF I (%x. M)" lemma PiF_gen_subset: "{(\' j\J. X j) |X J. J \ I \ X \ (\ j\J. sets (M j))} \ Pow (\J \ I. (\' j\J. space (M j)))" by (auto simp: Pi'_def) (blast dest: sets.sets_into_space) lemma space_PiF: "space (PiF I M) = (\J \ I. (\' j\J. space (M j)))" unfolding PiF_def using PiF_gen_subset by (rule space_measure_of) lemma sets_PiF: "sets (PiF I M) = sigma_sets (\J \ I. (\' j\J. space (M j))) {(\' j\J. X j) |X J. J \ I \ X \ (\ j\J. sets (M j))}" unfolding PiF_def using PiF_gen_subset by (rule sets_measure_of) lemma sets_PiF_singleton: "sets (PiF {I} M) = sigma_sets (\' j\I. space (M j)) {(\' j\I. X j) |X. X \ (\ j\I. sets (M j))}" unfolding sets_PiF by simp lemma in_sets_PiFI: assumes "X = (Pi' J S)" "J \ I" "\i. i\J \ S i \ sets (M i)" shows "X \ sets (PiF I M)" unfolding sets_PiF using assms by blast lemma product_in_sets_PiFI: assumes "J \ I" "\i. i\J \ S i \ sets (M i)" shows "(Pi' J S) \ sets (PiF I M)" unfolding sets_PiF using assms by blast lemma singleton_space_subset_in_sets: fixes J assumes "J \ I" assumes "finite J" shows "space (PiF {J} M) \ sets (PiF I M)" using assms by (intro in_sets_PiFI[where J=J and S="\i. space (M i)"]) (auto simp: product_def space_PiF) lemma singleton_subspace_set_in_sets: assumes A: "A \ sets (PiF {J} M)" assumes "finite J" assumes "J \ I" shows "A \ sets (PiF I M)" using A[unfolded sets_PiF] apply (induct A) unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric] using assms by (auto intro: in_sets_PiFI intro!: singleton_space_subset_in_sets) lemma finite_measurable_singletonI: assumes "finite I" assumes "\J. J \ I \ finite J" assumes MN: "\J. J \ I \ A \ measurable (PiF {J} M) N" shows "A \ measurable (PiF I M) N" unfolding measurable_def proof safe fix y assume "y \ sets N" have "A -` y \ space (PiF I M) = (\J\I. A -` y \ space (PiF {J} M))" by (auto simp: space_PiF) also have "\ \ sets (PiF I M)" proof (rule sets.finite_UN) show "finite I" by fact fix J assume "J \ I" with assms have "finite J" by simp show "A -` y \ space (PiF {J} M) \ sets (PiF I M)" by (rule singleton_subspace_set_in_sets[OF measurable_sets[OF assms(3)]]) fact+ qed finally show "A -` y \ space (PiF I M) \ sets (PiF I M)" . next fix x assume "x \ space (PiF I M)" thus "A x \ space N" using MN[of "domain x"] by (auto simp: space_PiF measurable_space Pi'_def) qed lemma countable_finite_comprehension: fixes f :: "'a::countable set \ _" assumes "\s. P s \ finite s" assumes "\s. P s \ f s \ sets M" shows "\{f s|s. P s} \ sets M" proof - have "\{f s|s. P s} = (\n::nat. let s = set (from_nat n) in if P s then f s else {})" proof safe fix x X s assume *: "x \ f s" "P s" with assms obtain l where "s = set l" using finite_list by blast with * show "x \ (\n. let s = set (from_nat n) in if P s then f s else {})" using \P s\ by (auto intro!: exI[where x="to_nat l"]) next fix x n assume "x \ (let s = set (from_nat n) in if P s then f s else {})" thus "x \ \{f s|s. P s}" using assms by (auto simp: Let_def split: if_split_asm) qed hence "\{f s|s. P s} = (\n. let s = set (from_nat n) in if P s then f s else {})" by simp also have "\ \ sets M" using assms by (auto simp: Let_def) finally show ?thesis . qed lemma space_subset_in_sets: fixes J::"'a::countable set set" assumes "J \ I" assumes "\j. j \ J \ finite j" shows "space (PiF J M) \ sets (PiF I M)" proof - have "space (PiF J M) = \{space (PiF {j} M)|j. j \ J}" unfolding space_PiF by blast also have "\ \ sets (PiF I M)" using assms by (intro countable_finite_comprehension) (auto simp: singleton_space_subset_in_sets) finally show ?thesis . qed lemma subspace_set_in_sets: fixes J::"'a::countable set set" assumes A: "A \ sets (PiF J M)" assumes "J \ I" assumes "\j. j \ J \ finite j" shows "A \ sets (PiF I M)" using A[unfolded sets_PiF] apply (induct A) unfolding sets_PiF[symmetric] unfolding space_PiF[symmetric] using assms by (auto intro: in_sets_PiFI intro!: space_subset_in_sets) lemma countable_measurable_PiFI: fixes I::"'a::countable set set" assumes MN: "\J. J \ I \ finite J \ A \ measurable (PiF {J} M) N" shows "A \ measurable (PiF I M) N" unfolding measurable_def proof safe fix y assume "y \ sets N" have "A -` y = (\{A -` y \ {x. domain x = J}|J. finite J})" by auto { fix x::"'a \\<^sub>F 'b" from finite_list[of "domain x"] obtain xs where "set xs = domain x" by auto hence "\n. domain x = set (from_nat n)" by (intro exI[where x="to_nat xs"]) auto } hence "A -` y \ space (PiF I M) = (\n. A -` y \ space (PiF ({set (from_nat n)}\I) M))" by (auto simp: space_PiF Pi'_def) also have "\ \ sets (PiF I M)" apply (intro sets.Int sets.countable_nat_UN subsetI, safe) apply (case_tac "set (from_nat i) \ I") apply simp_all apply (rule singleton_subspace_set_in_sets[OF measurable_sets[OF MN]]) using assms \y \ sets N\ apply (auto simp: space_PiF) done finally show "A -` y \ space (PiF I M) \ sets (PiF I M)" . next fix x assume "x \ space (PiF I M)" thus "A x \ space N" using MN[of "domain x"] by (auto simp: space_PiF measurable_space Pi'_def) qed lemma measurable_PiF: assumes f: "\x. x \ space N \ domain (f x) \ I \ (\i\domain (f x). (f x) i \ space (M i))" assumes S: "\J S. J \ I \ (\i. i \ J \ S i \ sets (M i)) \ f -` (Pi' J S) \ space N \ sets N" shows "f \ measurable N (PiF I M)" unfolding PiF_def using PiF_gen_subset apply (rule measurable_measure_of) using f apply force apply (insert S, auto) done lemma restrict_sets_measurable: assumes A: "A \ sets (PiF I M)" and "J \ I" shows "A \ {m. domain m \ J} \ sets (PiF J M)" using A[unfolded sets_PiF] proof (induct A) case (Basic a) then obtain K S where S: "a = Pi' K S" "K \ I" "(\i\K. S i \ sets (M i))" by auto show ?case proof cases assume "K \ J" hence "a \ {m. domain m \ J} \ {Pi' K X |X K. K \ J \ X \ (\ j\K. sets (M j))}" using S by (auto intro!: exI[where x=K] exI[where x=S] simp: Pi'_def) also have "\ \ sets (PiF J M)" unfolding sets_PiF by auto finally show ?thesis . next assume "K \ J" hence "a \ {m. domain m \ J} = {}" using S by (auto simp: Pi'_def) also have "\ \ sets (PiF J M)" by simp finally show ?thesis . qed next case (Union a) have "\(a ` UNIV) \ {m. domain m \ J} = (\i. (a i \ {m. domain m \ J}))" by simp also have "\ \ sets (PiF J M)" using Union by (intro sets.countable_nat_UN) auto finally show ?case . next case (Compl a) have "(space (PiF I M) - a) \ {m. domain m \ J} = (space (PiF J M) - (a \ {m. domain m \ J}))" using \J \ I\ by (auto simp: space_PiF Pi'_def) also have "\ \ sets (PiF J M)" using Compl by auto finally show ?case by (simp add: space_PiF) qed simp lemma measurable_finmap_of: assumes f: "\i. (\x \ space N. i \ J x) \ (\x. f x i) \ measurable N (M i)" assumes J: "\x. x \ space N \ J x \ I" "\x. x \ space N \ finite (J x)" assumes JN: "\S. {x. J x = S} \ space N \ sets N" shows "(\x. finmap_of (J x) (f x)) \ measurable N (PiF I M)" proof (rule measurable_PiF) fix x assume "x \ space N" with J[of x] measurable_space[OF f] show "domain (finmap_of (J x) (f x)) \ I \ (\i\domain (finmap_of (J x) (f x)). (finmap_of (J x) (f x)) i \ space (M i))" by auto next fix K S assume "K \ I" and *: "\i. i \ K \ S i \ sets (M i)" with J have eq: "(\x. finmap_of (J x) (f x)) -` Pi' K S \ space N = (if \x \ space N. K = J x \ finite K then if K = {} then {x \ space N. J x = K} else (\i\K. (\x. f x i) -` S i \ {x \ space N. J x = K}) else {})" by (auto simp: Pi'_def) have r: "{x \ space N. J x = K} = space N \ ({x. J x = K} \ space N)" by auto show "(\x. finmap_of (J x) (f x)) -` Pi' K S \ space N \ sets N" unfolding eq r apply (simp del: INT_simps add: ) apply (intro conjI impI sets.finite_INT JN sets.Int[OF sets.top]) apply simp apply assumption apply (subst Int_assoc[symmetric]) apply (rule sets.Int) apply (intro measurable_sets[OF f] *) apply force apply assumption apply (intro JN) done qed lemma measurable_PiM_finmap_of: assumes "finite J" shows "finmap_of J \ measurable (Pi\<^sub>M J M) (PiF {J} M)" apply (rule measurable_finmap_of) apply (rule measurable_component_singleton) apply simp apply rule apply (rule \finite J\) apply simp done lemma proj_measurable_singleton: assumes "A \ sets (M i)" shows "(\x. (x)\<^sub>F i) -` A \ space (PiF {I} M) \ sets (PiF {I} M)" proof cases assume "i \ I" hence "(\x. (x)\<^sub>F i) -` A \ space (PiF {I} M) = Pi' I (\x. if x = i then A else space (M x))" using sets.sets_into_space[OF ] \A \ sets (M i)\ assms by (auto simp: space_PiF Pi'_def) thus ?thesis using assms \A \ sets (M i)\ by (intro in_sets_PiFI) auto next assume "i \ I" hence "(\x. (x)\<^sub>F i) -` A \ space (PiF {I} M) = (if undefined \ A then space (PiF {I} M) else {})" by (auto simp: space_PiF Pi'_def) thus ?thesis by simp qed lemma measurable_proj_singleton: assumes "i \ I" shows "(\x. (x)\<^sub>F i) \ measurable (PiF {I} M) (M i)" by (unfold measurable_def, intro CollectI conjI ballI proj_measurable_singleton assms) (insert \i \ I\, auto simp: space_PiF) lemma measurable_proj_countable: fixes I::"'a::countable set set" assumes "y \ space (M i)" shows "(\x. if i \ domain x then (x)\<^sub>F i else y) \ measurable (PiF I M) (M i)" proof (rule countable_measurable_PiFI) fix J assume "J \ I" "finite J" show "(\x. if i \ domain x then x i else y) \ measurable (PiF {J} M) (M i)" unfolding measurable_def proof safe fix z assume "z \ sets (M i)" have "(\x. if i \ domain x then x i else y) -` z \ space (PiF {J} M) = (\x. if i \ J then (x)\<^sub>F i else y) -` z \ space (PiF {J} M)" by (auto simp: space_PiF Pi'_def) also have "\ \ sets (PiF {J} M)" using \z \ sets (M i)\ \finite J\ by (cases "i \ J") (auto intro!: measurable_sets[OF measurable_proj_singleton]) finally show "(\x. if i \ domain x then x i else y) -` z \ space (PiF {J} M) \ sets (PiF {J} M)" . qed (insert \y \ space (M i)\, auto simp: space_PiF Pi'_def) qed lemma measurable_restrict_proj: assumes "J \ II" "finite J" shows "finmap_of J \ measurable (PiM J M) (PiF II M)" using assms by (intro measurable_finmap_of measurable_component_singleton) auto lemma measurable_proj_PiM: fixes J K ::"'a::countable set" and I::"'a set set" assumes "finite J" "J \ I" assumes "x \ space (PiM J M)" shows "proj \ measurable (PiF {J} M) (PiM J M)" proof (rule measurable_PiM_single) show "proj \ space (PiF {J} M) \ (\\<^sub>E i \ J. space (M i))" using assms by (auto simp add: space_PiM space_PiF extensional_def sets_PiF Pi'_def) next fix A i assume A: "i \ J" "A \ sets (M i)" show "{\ \ space (PiF {J} M). (\)\<^sub>F i \ A} \ sets (PiF {J} M)" proof have "{\ \ space (PiF {J} M). (\)\<^sub>F i \ A} = (\\. (\)\<^sub>F i) -` A \ space (PiF {J} M)" by auto also have "\ \ sets (PiF {J} M)" using assms A by (auto intro: measurable_sets[OF measurable_proj_singleton] simp: space_PiM) finally show ?thesis . qed simp qed lemma space_PiF_singleton_eq_product: assumes "finite I" shows "space (PiF {I} M) = (\' i\I. space (M i))" by (auto simp: product_def space_PiF assms) text \adapted from @{thm sets_PiM_single}\ lemma sets_PiF_single: assumes "finite I" "I \ {}" shows "sets (PiF {I} M) = sigma_sets (\' i\I. space (M i)) {{f\\' i\I. space (M i). f i \ A} | i A. i \ I \ A \ sets (M i)}" (is "_ = sigma_sets ?\ ?R") unfolding sets_PiF_singleton proof (rule sigma_sets_eqI) interpret R: sigma_algebra ?\ "sigma_sets ?\ ?R" by (rule sigma_algebra_sigma_sets) auto fix A assume "A \ {Pi' I X |X. X \ (\ j\I. sets (M j))}" then obtain X where X: "A = Pi' I X" "X \ (\ j\I. sets (M j))" by auto show "A \ sigma_sets ?\ ?R" proof - from \I \ {}\ X have "A = (\j\I. {f\space (PiF {I} M). f j \ X j})" using sets.sets_into_space by (auto simp: space_PiF product_def) blast also have "\ \ sigma_sets ?\ ?R" using X \I \ {}\ assms by (intro R.finite_INT) (auto simp: space_PiF) finally show "A \ sigma_sets ?\ ?R" . qed next fix A assume "A \ ?R" then obtain i B where A: "A = {f\\' i\I. space (M i). f i \ B}" "i \ I" "B \ sets (M i)" by auto then have "A = (\' j \ I. if j = i then B else space (M j))" using sets.sets_into_space[OF A(3)] apply (auto simp: Pi'_iff split: if_split_asm) apply blast done also have "\ \ sigma_sets ?\ {Pi' I X |X. X \ (\ j\I. sets (M j))}" using A by (intro sigma_sets.Basic ) (auto intro: exI[where x="\j. if j = i then B else space (M j)"]) finally show "A \ sigma_sets ?\ {Pi' I X |X. X \ (\ j\I. sets (M j))}" . qed text \adapted from @{thm PiE_cong}\ lemma Pi'_cong: assumes "finite I" assumes "\i. i \ I \ f i = g i" shows "Pi' I f = Pi' I g" using assms by (auto simp: Pi'_def) text \adapted from @{thm Pi_UN}\ lemma Pi'_UN: fixes A :: "nat \ 'i \ 'a set" assumes "finite I" assumes mono: "\i n m. i \ I \ n \ m \ A n i \ A m i" shows "(\n. Pi' I (A n)) = Pi' I (\i. \n. A n i)" proof (intro set_eqI iffI) fix f assume "f \ Pi' I (\i. \n. A n i)" then have "\i\I. \n. f i \ A n i" "domain f = I" by (auto simp: \finite I\ Pi'_def) from bchoice[OF this(1)] obtain n where n: "\i. i \ I \ f i \ (A (n i) i)" by auto obtain k where k: "\i. i \ I \ n i \ k" using \finite I\ finite_nat_set_iff_bounded_le[of "n`I"] by auto have "f \ Pi' I (\i. A k i)" proof fix i assume "i \ I" from mono[OF this, of "n i" k] k[OF this] n[OF this] \domain f = I\ \i \ I\ show "f i \ A k i " by (auto simp: \finite I\) qed (simp add: \domain f = I\ \finite I\) then show "f \ (\n. Pi' I (A n))" by auto qed (auto simp: Pi'_def \finite I\) text \adapted from @{thm sets_PiM_sigma}\ lemma sigma_fprod_algebra_sigma_eq: fixes E :: "'i \ 'a set set" and S :: "'i \ nat \ 'a set" assumes [simp]: "finite I" "I \ {}" and S_union: "\i. i \ I \ (\j. S i j) = space (M i)" and S_in_E: "\i. i \ I \ range (S i) \ E i" assumes E_closed: "\i. i \ I \ E i \ Pow (space (M i))" and E_generates: "\i. i \ I \ sets (M i) = sigma_sets (space (M i)) (E i)" defines "P == { Pi' I F | F. \i\I. F i \ E i }" shows "sets (PiF {I} M) = sigma_sets (space (PiF {I} M)) P" proof let ?P = "sigma (space (Pi\<^sub>F {I} M)) P" - from \finite I\[THEN ex_bij_betw_finite_nat] guess T .. + from \finite I\[THEN ex_bij_betw_finite_nat] obtain T where "bij_betw T I {0..i. i \ I \ T i < card I" "\i. i\I \ the_inv_into I T (T i) = i" by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f simp del: \finite I\) have P_closed: "P \ Pow (space (Pi\<^sub>F {I} M))" using E_closed by (auto simp: space_PiF P_def Pi'_iff subset_eq) then have space_P: "space ?P = (\' i\I. space (M i))" by (simp add: space_PiF) have "sets (PiF {I} M) = sigma_sets (space ?P) {{f \ \' i\I. space (M i). f i \ A} |i A. i \ I \ A \ sets (M i)}" using sets_PiF_single[of I M] by (simp add: space_P) also have "\ \ sets (sigma (space (PiF {I} M)) P)" proof (safe intro!: sets.sigma_sets_subset) fix i A assume "i \ I" and A: "A \ sets (M i)" have "(\x. (x)\<^sub>F i) \ measurable ?P (sigma (space (M i)) (E i))" proof (subst measurable_iff_measure_of) show "E i \ Pow (space (M i))" using \i \ I\ by fact from space_P \i \ I\ show "(\x. (x)\<^sub>F i) \ space ?P \ space (M i)" by auto show "\A\E i. (\x. (x)\<^sub>F i) -` A \ space ?P \ sets ?P" proof fix A assume A: "A \ E i" from T have *: "\xs. length xs = card I \ (\j\I. (g)\<^sub>F j \ (if i = j then A else S j (xs ! T j)))" if "domain g = I" and "\j\I. (g)\<^sub>F j \ (if i = j then A else S j (f j))" for g f using that by (auto intro!: exI [of _ "map (\n. f (the_inv_into I T n)) [0..x. (x)\<^sub>F i) -` A \ space ?P = (\' j\I. if i = j then A else space (M j))" using E_closed \i \ I\ by (auto simp: space_P Pi_iff subset_eq split: if_split_asm) also have "\ = (\' j\I. \n. if i = j then A else S j n)" by (intro Pi'_cong) (simp_all add: S_union) also have "\ = {g. domain g = I \ (\f. \j\I. (g)\<^sub>F j \ (if i = j then A else S j (f j)))}" by (rule set_eqI) (simp del: if_image_distrib add: Pi'_iff bchoice_iff) also have "\ = (\xs\{xs. length xs = card I}. \' j\I. if i = j then A else S j (xs ! T j))" using * by (auto simp add: Pi'_iff split del: if_split) also have "\ \ sets ?P" proof (safe intro!: sets.countable_UN) fix xs show "(\' j\I. if i = j then A else S j (xs ! T j)) \ sets ?P" using A S_in_E by (simp add: P_closed) (auto simp: P_def subset_eq intro!: exI[of _ "\j. if i = j then A else S j (xs ! T j)"]) qed finally show "(\x. (x)\<^sub>F i) -` A \ space ?P \ sets ?P" using P_closed by simp qed qed from measurable_sets[OF this, of A] A \i \ I\ E_closed have "(\x. (x)\<^sub>F i) -` A \ space ?P \ sets ?P" by (simp add: E_generates) also have "(\x. (x)\<^sub>F i) -` A \ space ?P = {f \ \' i\I. space (M i). f i \ A}" using P_closed by (auto simp: space_PiF) finally show "\ \ sets ?P" . qed finally show "sets (PiF {I} M) \ sigma_sets (space (PiF {I} M)) P" by (simp add: P_closed) show "sigma_sets (space (PiF {I} M)) P \ sets (PiF {I} M)" using \finite I\ \I \ {}\ by (auto intro!: sets.sigma_sets_subset product_in_sets_PiFI simp: E_generates P_def) qed lemma product_open_generates_sets_PiF_single: assumes "I \ {}" assumes [simp]: "finite I" shows "sets (PiF {I} (\_. borel::'b::second_countable_topology measure)) = sigma_sets (space (PiF {I} (\_. borel))) {Pi' I F |F. (\i\I. F i \ Collect open)}" proof - - from open_countable_basisE[OF open_UNIV] guess S::"'b set set" . note S = this + from open_countable_basisE[OF open_UNIV] obtain S::"'b set set" + where S: "S \ (SOME B. countable B \ topological_basis B)" "UNIV = \ S" + by auto show ?thesis proof (rule sigma_fprod_algebra_sigma_eq) show "finite I" by simp show "I \ {}" by fact define S' where "S' = from_nat_into S" show "(\j. S' j) = space borel" using S apply (auto simp add: from_nat_into countable_basis_proj S'_def basis_proj_def) apply (metis (lifting, mono_tags) UNIV_I UnionE basis_proj_def countable_basis_proj countable_subset from_nat_into_surj) done show "range S' \ Collect open" using S apply (auto simp add: from_nat_into countable_basis_proj S'_def) apply (metis UNIV_not_empty Union_empty from_nat_into subsetD topological_basis_open[OF basis_proj] basis_proj_def) done show "Collect open \ Pow (space borel)" by simp show "sets borel = sigma_sets (space borel) (Collect open)" by (simp add: borel_def) qed qed lemma finmap_UNIV[simp]: "(\J\Collect finite. \' j\J. UNIV) = UNIV" by auto lemma borel_eq_PiF_borel: shows "(borel :: ('i::countable \\<^sub>F 'a::polish_space) measure) = PiF (Collect finite) (\_. borel :: 'a measure)" unfolding borel_def PiF_def proof (rule measure_eqI, clarsimp, rule sigma_sets_eqI) fix a::"('i \\<^sub>F 'a) set" assume "a \ Collect open" hence "open a" by simp then obtain B' where B': "B'\basis_finmap" "a = \B'" using finmap_topological_basis by (force simp add: topological_basis_def) have "a \ sigma UNIV {Pi' J X |X J. finite J \ X \ J \ sigma_sets UNIV (Collect open)}" unfolding \a = \B'\ proof (rule sets.countable_Union) from B' countable_basis_finmap show "countable B'" by (metis countable_subset) next show "B' \ sets (sigma UNIV {Pi' J X |X J. finite J \ X \ J \ sigma_sets UNIV (Collect open)})" (is "_ \ sets ?s") proof fix x assume "x \ B'" with B' have "x \ basis_finmap" by auto then obtain J X where "x = Pi' J X" "finite J" "X \ J \ sigma_sets UNIV (Collect open)" by (auto simp: basis_finmap_def topological_basis_open[OF basis_proj]) thus "x \ sets ?s" by auto qed qed thus "a \ sigma_sets UNIV {Pi' J X |X J. finite J \ X \ J \ sigma_sets UNIV (Collect open)}" by simp next fix b::"('i \\<^sub>F 'a) set" assume "b \ {Pi' J X |X J. finite J \ X \ J \ sigma_sets UNIV (Collect open)}" hence b': "b \ sets (Pi\<^sub>F (Collect finite) (\_. borel))" by (auto simp: sets_PiF borel_def) let ?b = "\J. b \ {x. domain x = J}" have "b = \((\J. ?b J) ` Collect finite)" by auto also have "\ \ sets borel" proof (rule sets.countable_Union, safe) fix J::"'i set" assume "finite J" { assume ef: "J = {}" have "?b J \ sets borel" proof cases assume "?b J \ {}" then obtain f where "f \ b" "domain f = {}" using ef by auto hence "?b J = {f}" using \J = {}\ by (auto simp: finmap_eq_iff) also have "{f} \ sets borel" by simp finally show ?thesis . qed simp } moreover { assume "J \ ({}::'i set)" have "(?b J) = b \ {m. domain m \ {J}}" by auto also have "\ \ sets (PiF {J} (\_. borel))" using b' by (rule restrict_sets_measurable) (auto simp: \finite J\) also have "\ = sigma_sets (space (PiF {J} (\_. borel))) {Pi' (J) F |F. (\j\J. F j \ Collect open)}" (is "_ = sigma_sets _ ?P") by (rule product_open_generates_sets_PiF_single[OF \J \ {}\ \finite J\]) also have "\ \ sigma_sets UNIV (Collect open)" by (intro sigma_sets_mono'') (auto intro!: open_Pi'I simp: space_PiF) finally have "(?b J) \ sets borel" by (simp add: borel_def) } ultimately show "(?b J) \ sets borel" by blast qed (simp add: countable_Collect_finite) finally show "b \ sigma_sets UNIV (Collect open)" by (simp add: borel_def) qed (simp add: emeasure_sigma borel_def PiF_def) subsection \Isomorphism between Functions and Finite Maps\ lemma measurable_finmap_compose: shows "(\m. compose J m f) \ measurable (PiM (f ` J) (\_. M)) (PiM J (\_. M))" unfolding compose_def by measurable lemma measurable_compose_inv: assumes inj: "\j. j \ J \ f' (f j) = j" shows "(\m. compose (f ` J) m f') \ measurable (PiM J (\_. M)) (PiM (f ` J) (\_. M))" unfolding compose_def by (rule measurable_restrict) (auto simp: inj) locale function_to_finmap = fixes J::"'a set" and f :: "'a \ 'b::countable" and f' assumes [simp]: "finite J" assumes inv: "i \ J \ f' (f i) = i" begin text \to measure finmaps\ definition "fm = (finmap_of (f ` J)) o (\g. compose (f ` J) g f')" lemma domain_fm[simp]: "domain (fm x) = f ` J" unfolding fm_def by simp lemma fm_restrict[simp]: "fm (restrict y J) = fm y" unfolding fm_def by (auto simp: compose_def inv intro: restrict_ext) lemma fm_product: assumes "\i. space (M i) = UNIV" shows "fm -` Pi' (f ` J) S \ space (Pi\<^sub>M J M) = (\\<^sub>E j \ J. S (f j))" using assms by (auto simp: inv fm_def compose_def space_PiM Pi'_def) lemma fm_measurable: assumes "f ` J \ N" shows "fm \ measurable (Pi\<^sub>M J (\_. M)) (Pi\<^sub>F N (\_. M))" unfolding fm_def proof (rule measurable_comp, rule measurable_compose_inv) show "finmap_of (f ` J) \ measurable (Pi\<^sub>M (f ` J) (\_. M)) (PiF N (\_. M)) " using assms by (intro measurable_finmap_of measurable_component_singleton) auto qed (simp_all add: inv) lemma proj_fm: assumes "x \ J" shows "fm m (f x) = m x" using assms by (auto simp: fm_def compose_def o_def inv) lemma inj_on_compose_f': "inj_on (\g. compose (f ` J) g f') (extensional J)" proof (rule inj_on_inverseI) fix x::"'a \ 'c" assume "x \ extensional J" thus "(\x. compose J x f) (compose (f ` J) x f') = x" by (auto simp: compose_def inv extensional_def) qed lemma inj_on_fm: assumes "\i. space (M i) = UNIV" shows "inj_on fm (space (Pi\<^sub>M J M))" using assms apply (auto simp: fm_def space_PiM PiE_def) apply (rule comp_inj_on) apply (rule inj_on_compose_f') apply (rule finmap_of_inj_on_extensional_finite) apply simp apply (auto) done text \to measure functions\ definition "mf = (\g. compose J g f) o proj" lemma mf_fm: assumes "x \ space (Pi\<^sub>M J (\_. M))" shows "mf (fm x) = x" proof - have "mf (fm x) \ extensional J" by (auto simp: mf_def extensional_def compose_def) moreover have "x \ extensional J" using assms sets.sets_into_space by (force simp: space_PiM PiE_def) moreover { fix i assume "i \ J" hence "mf (fm x) i = x i" by (auto simp: inv mf_def compose_def fm_def) } ultimately show ?thesis by (rule extensionalityI) qed lemma mf_measurable: assumes "space M = UNIV" shows "mf \ measurable (PiF {f ` J} (\_. M)) (PiM J (\_. M))" unfolding mf_def proof (rule measurable_comp, rule measurable_proj_PiM) show "(\g. compose J g f) \ measurable (Pi\<^sub>M (f ` J) (\x. M)) (Pi\<^sub>M J (\_. M))" by (rule measurable_finmap_compose) qed (auto simp add: space_PiM extensional_def assms) lemma fm_image_measurable: assumes "space M = UNIV" assumes "X \ sets (Pi\<^sub>M J (\_. M))" shows "fm ` X \ sets (PiF {f ` J} (\_. M))" proof - have "fm ` X = (mf) -` X \ space (PiF {f ` J} (\_. M))" proof safe fix x assume "x \ X" with mf_fm[of x] sets.sets_into_space[OF assms(2)] show "fm x \ mf -` X" by auto show "fm x \ space (PiF {f ` J} (\_. M))" by (simp add: space_PiF assms) next fix y x assume x: "mf y \ X" assume y: "y \ space (PiF {f ` J} (\_. M))" thus "y \ fm ` X" by (intro image_eqI[OF _ x], unfold finmap_eq_iff) (auto simp: space_PiF fm_def mf_def compose_def inv Pi'_def) qed also have "\ \ sets (PiF {f ` J} (\_. M))" using assms by (intro measurable_sets[OF mf_measurable]) auto finally show ?thesis . qed lemma fm_image_measurable_finite: assumes "space M = UNIV" assumes "X \ sets (Pi\<^sub>M J (\_. M::'c measure))" shows "fm ` X \ sets (PiF (Collect finite) (\_. M::'c measure))" using fm_image_measurable[OF assms] by (rule subspace_set_in_sets) (auto simp: finite_subset) text \measure on finmaps\ definition "mapmeasure M N = distr M (PiF (Collect finite) N) (fm)" lemma sets_mapmeasure[simp]: "sets (mapmeasure M N) = sets (PiF (Collect finite) N)" unfolding mapmeasure_def by simp lemma space_mapmeasure[simp]: "space (mapmeasure M N) = space (PiF (Collect finite) N)" unfolding mapmeasure_def by simp lemma mapmeasure_PiF: assumes s1: "space M = space (Pi\<^sub>M J (\_. N))" assumes s2: "sets M = sets (Pi\<^sub>M J (\_. N))" assumes "space N = UNIV" assumes "X \ sets (PiF (Collect finite) (\_. N))" shows "emeasure (mapmeasure M (\_. N)) X = emeasure M ((fm -` X \ extensional J))" using assms by (auto simp: measurable_cong_sets[OF s2 refl] mapmeasure_def emeasure_distr fm_measurable space_PiM PiE_def) lemma mapmeasure_PiM: fixes N::"'c measure" assumes s1: "space M = space (Pi\<^sub>M J (\_. N))" assumes s2: "sets M = (Pi\<^sub>M J (\_. N))" assumes N: "space N = UNIV" assumes X: "X \ sets M" shows "emeasure M X = emeasure (mapmeasure M (\_. N)) (fm ` X)" unfolding mapmeasure_def proof (subst emeasure_distr, subst measurable_cong_sets[OF s2 refl], rule fm_measurable) have "X \ space (Pi\<^sub>M J (\_. N))" using assms by (simp add: sets.sets_into_space) from assms inj_on_fm[of "\_. N"] subsetD[OF this] have "fm -` fm ` X \ space (Pi\<^sub>M J (\_. N)) = X" by (auto simp: vimage_image_eq inj_on_def) thus "emeasure M X = emeasure M (fm -` fm ` X \ space M)" using s1 by simp show "fm ` X \ sets (PiF (Collect finite) (\_. N))" by (rule fm_image_measurable_finite[OF N X[simplified s2]]) qed simp end end diff --git a/src/HOL/Probability/Independent_Family.thy b/src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy +++ b/src/HOL/Probability/Independent_Family.thy @@ -1,1400 +1,1408 @@ (* Title: HOL/Probability/Independent_Family.thy Author: Johannes Hölzl, TU München Author: Sudeep Kanav, TU München *) section \Independent families of events, event sets, and random variables\ theory Independent_Family imports Infinite_Product_Measure begin definition (in prob_space) "indep_sets F I \ (\i\I. F i \ events) \ (\J\I. J \ {} \ finite J \ (\A\Pi J F. prob (\j\J. A j) = (\j\J. prob (A j))))" definition (in prob_space) "indep_set A B \ indep_sets (case_bool A B) UNIV" definition (in prob_space) indep_events_def_alt: "indep_events A I \ indep_sets (\i. {A i}) I" lemma (in prob_space) indep_events_def: "indep_events A I \ (A`I \ events) \ (\J\I. J \ {} \ finite J \ prob (\j\J. A j) = (\j\J. prob (A j)))" unfolding indep_events_def_alt indep_sets_def apply (simp add: Ball_def Pi_iff image_subset_iff_funcset) apply (intro conj_cong refl arg_cong[where f=All] ext imp_cong) apply auto done lemma (in prob_space) indep_eventsI: "(\i. i \ I \ F i \ sets M) \ (\J. J \ I \ finite J \ J \ {} \ prob (\i\J. F i) = (\i\J. prob (F i))) \ indep_events F I" by (auto simp: indep_events_def) definition (in prob_space) "indep_event A B \ indep_events (case_bool A B) UNIV" lemma (in prob_space) indep_sets_cong: "I = J \ (\i. i \ I \ F i = G i) \ indep_sets F I \ indep_sets G J" by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+ lemma (in prob_space) indep_events_finite_index_events: "indep_events F I \ (\J\I. J \ {} \ finite J \ indep_events F J)" by (auto simp: indep_events_def) lemma (in prob_space) indep_sets_finite_index_sets: "indep_sets F I \ (\J\I. J \ {} \ finite J \ indep_sets F J)" proof (intro iffI allI impI) assume *: "\J\I. J \ {} \ finite J \ indep_sets F J" show "indep_sets F I" unfolding indep_sets_def proof (intro conjI ballI allI impI) fix i assume "i \ I" with *[THEN spec, of "{i}"] show "F i \ events" by (auto simp: indep_sets_def) qed (insert *, auto simp: indep_sets_def) qed (auto simp: indep_sets_def) lemma (in prob_space) indep_sets_mono_index: "J \ I \ indep_sets F I \ indep_sets F J" unfolding indep_sets_def by auto lemma (in prob_space) indep_sets_mono_sets: assumes indep: "indep_sets F I" assumes mono: "\i. i\I \ G i \ F i" shows "indep_sets G I" proof - have "(\i\I. F i \ events) \ (\i\I. G i \ events)" using mono by auto moreover have "\A J. J \ I \ A \ (\ j\J. G j) \ A \ (\ j\J. F j)" using mono by (auto simp: Pi_iff) ultimately show ?thesis using indep by (auto simp: indep_sets_def) qed lemma (in prob_space) indep_sets_mono: assumes indep: "indep_sets F I" assumes mono: "J \ I" "\i. i\J \ G i \ F i" shows "indep_sets G J" apply (rule indep_sets_mono_sets) apply (rule indep_sets_mono_index) apply (fact +) done lemma (in prob_space) indep_setsI: assumes "\i. i \ I \ F i \ events" and "\A J. J \ {} \ J \ I \ finite J \ (\j\J. A j \ F j) \ prob (\j\J. A j) = (\j\J. prob (A j))" shows "indep_sets F I" using assms unfolding indep_sets_def by (auto simp: Pi_iff) lemma (in prob_space) indep_setsD: assumes "indep_sets F I" and "J \ I" "J \ {}" "finite J" "\j\J. A j \ F j" shows "prob (\j\J. A j) = (\j\J. prob (A j))" using assms unfolding indep_sets_def by auto lemma (in prob_space) indep_setI: assumes ev: "A \ events" "B \ events" and indep: "\a b. a \ A \ b \ B \ prob (a \ b) = prob a * prob b" shows "indep_set A B" unfolding indep_set_def proof (rule indep_setsI) fix F J assume "J \ {}" "J \ UNIV" and F: "\j\J. F j \ (case j of True \ A | False \ B)" have "J \ Pow UNIV" by auto with F \J \ {}\ indep[of "F True" "F False"] show "prob (\j\J. F j) = (\j\J. prob (F j))" unfolding UNIV_bool Pow_insert by (auto simp: ac_simps) qed (auto split: bool.split simp: ev) lemma (in prob_space) indep_setD: assumes indep: "indep_set A B" and ev: "a \ A" "b \ B" shows "prob (a \ b) = prob a * prob b" using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "case_bool a b"] ev by (simp add: ac_simps UNIV_bool) lemma (in prob_space) assumes indep: "indep_set A B" shows indep_setD_ev1: "A \ events" and indep_setD_ev2: "B \ events" using indep unfolding indep_set_def indep_sets_def UNIV_bool by auto lemma (in prob_space) indep_sets_Dynkin: assumes indep: "indep_sets F I" shows "indep_sets (\i. Dynkin (space M) (F i)) I" (is "indep_sets ?F I") proof (subst indep_sets_finite_index_sets, intro allI impI ballI) fix J assume "finite J" "J \ I" "J \ {}" with indep have "indep_sets F J" by (subst (asm) indep_sets_finite_index_sets) auto { fix J K assume "indep_sets F K" let ?G = "\S i. if i \ S then ?F i else F i" assume "finite J" "J \ K" then have "indep_sets (?G J) K" proof induct case (insert j J) moreover define G where "G = ?G J" ultimately have G: "indep_sets G K" "\i. i \ K \ G i \ events" and "j \ K" by (auto simp: indep_sets_def) let ?D = "{E\events. indep_sets (G(j := {E})) K }" { fix X assume X: "X \ events" assume indep: "\J A. J \ {} \ J \ K \ finite J \ j \ J \ (\i\J. A i \ G i) \ prob ((\i\J. A i) \ X) = prob X * (\i\J. prob (A i))" have "indep_sets (G(j := {X})) K" proof (rule indep_setsI) fix i assume "i \ K" then show "(G(j:={X})) i \ events" using G X by auto next fix A J assume J: "J \ {}" "J \ K" "finite J" "\i\J. A i \ (G(j := {X})) i" show "prob (\j\J. A j) = (\j\J. prob (A j))" proof cases assume "j \ J" with J have "A j = X" by auto show ?thesis proof cases assume "J = {j}" then show ?thesis by simp next assume "J \ {j}" have "prob (\i\J. A i) = prob ((\i\J-{j}. A i) \ X)" using \j \ J\ \A j = X\ by (auto intro!: arg_cong[where f=prob] split: if_split_asm) also have "\ = prob X * (\i\J-{j}. prob (A i))" proof (rule indep) show "J - {j} \ {}" "J - {j} \ K" "finite (J - {j})" "j \ J - {j}" using J \J \ {j}\ \j \ J\ by auto show "\i\J - {j}. A i \ G i" using J by auto qed also have "\ = prob (A j) * (\i\J-{j}. prob (A i))" using \A j = X\ by simp also have "\ = (\i\J. prob (A i))" unfolding prod.insert_remove[OF \finite J\, symmetric, of "\i. prob (A i)"] using \j \ J\ by (simp add: insert_absorb) finally show ?thesis . qed next assume "j \ J" with J have "\i\J. A i \ G i" by (auto split: if_split_asm) with J show ?thesis by (intro indep_setsD[OF G(1)]) auto qed qed } note indep_sets_insert = this have "Dynkin_system (space M) ?D" proof (rule Dynkin_systemI', simp_all cong del: indep_sets_cong, safe) show "indep_sets (G(j := {{}})) K" by (rule indep_sets_insert) auto next fix X assume X: "X \ events" and G': "indep_sets (G(j := {X})) K" show "indep_sets (G(j := {space M - X})) K" proof (rule indep_sets_insert) fix J A assume J: "J \ {}" "J \ K" "finite J" "j \ J" and A: "\i\J. A i \ G i" then have A_sets: "\i. i\J \ A i \ events" using G by auto have "prob ((\j\J. A j) \ (space M - X)) = prob ((\j\J. A j) - (\i\insert j J. (A(j := X)) i))" using A_sets sets.sets_into_space[of _ M] X \J \ {}\ by (auto intro!: arg_cong[where f=prob] split: if_split_asm) also have "\ = prob (\j\J. A j) - prob (\i\insert j J. (A(j := X)) i)" using J \J \ {}\ \j \ J\ A_sets X sets.sets_into_space by (auto intro!: finite_measure_Diff sets.finite_INT split: if_split_asm) finally have "prob ((\j\J. A j) \ (space M - X)) = prob (\j\J. A j) - prob (\i\insert j J. (A(j := X)) i)" . moreover { have "prob (\j\J. A j) = (\j\J. prob (A j))" using J A \finite J\ by (intro indep_setsD[OF G(1)]) auto then have "prob (\j\J. A j) = prob (space M) * (\i\J. prob (A i))" using prob_space by simp } moreover { have "prob (\i\insert j J. (A(j := X)) i) = (\i\insert j J. prob ((A(j := X)) i))" using J A \j \ K\ by (intro indep_setsD[OF G']) auto then have "prob (\i\insert j J. (A(j := X)) i) = prob X * (\i\J. prob (A i))" using \finite J\ \j \ J\ by (auto intro!: prod.cong) } ultimately have "prob ((\j\J. A j) \ (space M - X)) = (prob (space M) - prob X) * (\i\J. prob (A i))" by (simp add: field_simps) also have "\ = prob (space M - X) * (\i\J. prob (A i))" using X A by (simp add: finite_measure_compl) finally show "prob ((\j\J. A j) \ (space M - X)) = prob (space M - X) * (\i\J. prob (A i))" . qed (insert X, auto) next fix F :: "nat \ 'a set" assume disj: "disjoint_family F" and "range F \ ?D" then have F: "\i. F i \ events" "\i. indep_sets (G(j:={F i})) K" by auto show "indep_sets (G(j := {\k. F k})) K" proof (rule indep_sets_insert) fix J A assume J: "j \ J" "J \ {}" "J \ K" "finite J" and A: "\i\J. A i \ G i" then have A_sets: "\i. i\J \ A i \ events" using G by auto have "prob ((\j\J. A j) \ (\k. F k)) = prob (\k. (\i\insert j J. (A(j := F k)) i))" using \J \ {}\ \j \ J\ \j \ K\ by (auto intro!: arg_cong[where f=prob] split: if_split_asm) moreover have "(\k. prob (\i\insert j J. (A(j := F k)) i)) sums prob (\k. (\i\insert j J. (A(j := F k)) i))" proof (rule finite_measure_UNION) show "disjoint_family (\k. \i\insert j J. (A(j := F k)) i)" using disj by (rule disjoint_family_on_bisimulation) auto show "range (\k. \i\insert j J. (A(j := F k)) i) \ events" using A_sets F \finite J\ \J \ {}\ \j \ J\ by (auto intro!: sets.Int) qed moreover { fix k from J A \j \ K\ have "prob (\i\insert j J. (A(j := F k)) i) = prob (F k) * (\i\J. prob (A i))" by (subst indep_setsD[OF F(2)]) (auto intro!: prod.cong split: if_split_asm) also have "\ = prob (F k) * prob (\i\J. A i)" using J A \j \ K\ by (subst indep_setsD[OF G(1)]) auto finally have "prob (\i\insert j J. (A(j := F k)) i) = prob (F k) * prob (\i\J. A i)" . } ultimately have "(\k. prob (F k) * prob (\i\J. A i)) sums (prob ((\j\J. A j) \ (\k. F k)))" by simp moreover have "(\k. prob (F k) * prob (\i\J. A i)) sums (prob (\k. F k) * prob (\i\J. A i))" using disj F(1) by (intro finite_measure_UNION sums_mult2) auto then have "(\k. prob (F k) * prob (\i\J. A i)) sums (prob (\k. F k) * (\i\J. prob (A i)))" using J A \j \ K\ by (subst indep_setsD[OF G(1), symmetric]) auto ultimately show "prob ((\j\J. A j) \ (\k. F k)) = prob (\k. F k) * (\j\J. prob (A j))" by (auto dest!: sums_unique) qed (insert F, auto) qed (insert sets.sets_into_space, auto) then have mono: "Dynkin (space M) (G j) \ {E \ events. indep_sets (G(j := {E})) K}" proof (rule Dynkin_system.Dynkin_subset, safe) fix X assume "X \ G j" then show "X \ events" using G \j \ K\ by auto from \indep_sets G K\ show "indep_sets (G(j := {X})) K" by (rule indep_sets_mono_sets) (insert \X \ G j\, auto) qed have "indep_sets (G(j:=?D)) K" proof (rule indep_setsI) fix i assume "i \ K" then show "(G(j := ?D)) i \ events" using G(2) by auto next fix A J assume J: "J\{}" "J \ K" "finite J" and A: "\i\J. A i \ (G(j := ?D)) i" show "prob (\j\J. A j) = (\j\J. prob (A j))" proof cases assume "j \ J" with A have indep: "indep_sets (G(j := {A j})) K" by auto from J A show ?thesis by (intro indep_setsD[OF indep]) auto next assume "j \ J" with J A have "\i\J. A i \ G i" by (auto split: if_split_asm) with J show ?thesis by (intro indep_setsD[OF G(1)]) auto qed qed then have "indep_sets (G(j := Dynkin (space M) (G j))) K" by (rule indep_sets_mono_sets) (insert mono, auto) then show ?case by (rule indep_sets_mono_sets) (insert \j \ K\ \j \ J\, auto simp: G_def) qed (insert \indep_sets F K\, simp) } from this[OF \indep_sets F J\ \finite J\ subset_refl] show "indep_sets ?F J" by (rule indep_sets_mono_sets) auto qed lemma (in prob_space) indep_sets_sigma: assumes indep: "indep_sets F I" assumes stable: "\i. i \ I \ Int_stable (F i)" shows "indep_sets (\i. sigma_sets (space M) (F i)) I" proof - from indep_sets_Dynkin[OF indep] show ?thesis proof (rule indep_sets_mono_sets, subst sigma_eq_Dynkin, simp_all add: stable) fix i assume "i \ I" with indep have "F i \ events" by (auto simp: indep_sets_def) with sets.sets_into_space show "F i \ Pow (space M)" by auto qed qed lemma (in prob_space) indep_sets_sigma_sets_iff: assumes "\i. i \ I \ Int_stable (F i)" shows "indep_sets (\i. sigma_sets (space M) (F i)) I \ indep_sets F I" proof assume "indep_sets F I" then show "indep_sets (\i. sigma_sets (space M) (F i)) I" by (rule indep_sets_sigma) fact next assume "indep_sets (\i. sigma_sets (space M) (F i)) I" then show "indep_sets F I" by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic) qed definition (in prob_space) indep_vars_def2: "indep_vars M' X I \ (\i\I. random_variable (M' i) (X i)) \ indep_sets (\i. { X i -` A \ space M | A. A \ sets (M' i)}) I" definition (in prob_space) "indep_var Ma A Mb B \ indep_vars (case_bool Ma Mb) (case_bool A B) UNIV" lemma (in prob_space) indep_vars_def: "indep_vars M' X I \ (\i\I. random_variable (M' i) (X i)) \ indep_sets (\i. sigma_sets (space M) { X i -` A \ space M | A. A \ sets (M' i)}) I" unfolding indep_vars_def2 apply (rule conj_cong[OF refl]) apply (rule indep_sets_sigma_sets_iff[symmetric]) apply (auto simp: Int_stable_def) apply (rule_tac x="A \ Aa" in exI) apply auto done lemma (in prob_space) indep_var_eq: "indep_var S X T Y \ (random_variable S X \ random_variable T Y) \ indep_set (sigma_sets (space M) { X -` A \ space M | A. A \ sets S}) (sigma_sets (space M) { Y -` A \ space M | A. A \ sets T})" unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool by (intro arg_cong2[where f="(\)"] arg_cong2[where f=indep_sets] ext) (auto split: bool.split) lemma (in prob_space) indep_sets2_eq: "indep_set A B \ A \ events \ B \ events \ (\a\A. \b\B. prob (a \ b) = prob a * prob b)" unfolding indep_set_def proof (intro iffI ballI conjI) assume indep: "indep_sets (case_bool A B) UNIV" { fix a b assume "a \ A" "b \ B" with indep_setsD[OF indep, of UNIV "case_bool a b"] show "prob (a \ b) = prob a * prob b" unfolding UNIV_bool by (simp add: ac_simps) } from indep show "A \ events" "B \ events" unfolding indep_sets_def UNIV_bool by auto next assume *: "A \ events \ B \ events \ (\a\A. \b\B. prob (a \ b) = prob a * prob b)" show "indep_sets (case_bool A B) UNIV" proof (rule indep_setsI) fix i show "(case i of True \ A | False \ B) \ events" using * by (auto split: bool.split) next fix J X assume "J \ {}" "J \ UNIV" and X: "\j\J. X j \ (case j of True \ A | False \ B)" then have "J = {True} \ J = {False} \ J = {True,False}" by (auto simp: UNIV_bool) then show "prob (\j\J. X j) = (\j\J. prob (X j))" using X * by auto qed qed lemma (in prob_space) indep_set_sigma_sets: assumes "indep_set A B" assumes A: "Int_stable A" and B: "Int_stable B" shows "indep_set (sigma_sets (space M) A) (sigma_sets (space M) B)" proof - have "indep_sets (\i. sigma_sets (space M) (case i of True \ A | False \ B)) UNIV" proof (rule indep_sets_sigma) show "indep_sets (case_bool A B) UNIV" by (rule \indep_set A B\[unfolded indep_set_def]) fix i show "Int_stable (case i of True \ A | False \ B)" using A B by (cases i) auto qed then show ?thesis unfolding indep_set_def by (rule indep_sets_mono_sets) (auto split: bool.split) qed lemma (in prob_space) indep_eventsI_indep_vars: assumes indep: "indep_vars N X I" assumes P: "\i. i \ I \ {x\space (N i). P i x} \ sets (N i)" shows "indep_events (\i. {x\space M. P i (X i x)}) I" proof - have "indep_sets (\i. {X i -` A \ space M |A. A \ sets (N i)}) I" using indep unfolding indep_vars_def2 by auto then show ?thesis unfolding indep_events_def_alt proof (rule indep_sets_mono_sets) fix i assume "i \ I" then have "{{x \ space M. P i (X i x)}} = {X i -` {x\space (N i). P i x} \ space M}" using indep by (auto simp: indep_vars_def dest: measurable_space) also have "\ \ {X i -` A \ space M |A. A \ sets (N i)}" using P[OF \i \ I\] by blast finally show "{{x \ space M. P i (X i x)}} \ {X i -` A \ space M |A. A \ sets (N i)}" . qed qed lemma (in prob_space) indep_sets_collect_sigma: fixes I :: "'j \ 'i set" and J :: "'j set" and E :: "'i \ 'a set set" assumes indep: "indep_sets E (\j\J. I j)" assumes Int_stable: "\i j. j \ J \ i \ I j \ Int_stable (E i)" assumes disjoint: "disjoint_family_on I J" shows "indep_sets (\j. sigma_sets (space M) (\i\I j. E i)) J" proof - let ?E = "\j. {\k\K. E' k| E' K. finite K \ K \ {} \ K \ I j \ (\k\K. E' k \ E k) }" from indep have E: "\j i. j \ J \ i \ I j \ E i \ events" unfolding indep_sets_def by auto { fix j let ?S = "sigma_sets (space M) (\i\I j. E i)" assume "j \ J" from E[OF this] interpret S: sigma_algebra "space M" ?S using sets.sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto have "sigma_sets (space M) (\i\I j. E i) = sigma_sets (space M) (?E j)" proof (rule sigma_sets_eqI) fix A assume "A \ (\i\I j. E i)" - then guess i .. + then obtain i where "i \ I j" "A \ E i" .. then show "A \ sigma_sets (space M) (?E j)" by (auto intro!: sigma_sets.intros(2-) exI[of _ "{i}"] exI[of _ "\i. A"]) next fix A assume "A \ ?E j" then obtain E' K where "finite K" "K \ {}" "K \ I j" "\k. k \ K \ E' k \ E k" and A: "A = (\k\K. E' k)" by auto then have "A \ ?S" unfolding A by (safe intro!: S.finite_INT) auto then show "A \ sigma_sets (space M) (\i\I j. E i)" by simp qed } moreover have "indep_sets (\j. sigma_sets (space M) (?E j)) J" proof (rule indep_sets_sigma) show "indep_sets ?E J" proof (intro indep_setsI) fix j assume "j \ J" with E show "?E j \ events" by (force intro!: sets.finite_INT) next fix K A assume K: "K \ {}" "K \ J" "finite K" and "\j\K. A j \ ?E j" then have "\j\K. \E' L. A j = (\l\L. E' l) \ finite L \ L \ {} \ L \ I j \ (\l\L. E' l \ E l)" by simp - from bchoice[OF this] guess E' .. + from bchoice[OF this] obtain E' + where "\x\K. \L. A x = \ (E' x ` L) \ finite L \ L \ {} \ L \ I x \ (\l\L. E' x l \ E l)" + .. from bchoice[OF this] obtain L where A: "\j. j\K \ A j = (\l\L j. E' j l)" and L: "\j. j\K \ finite (L j)" "\j. j\K \ L j \ {}" "\j. j\K \ L j \ I j" and E': "\j l. j\K \ l \ L j \ E' j l \ E l" by auto - { fix k l j assume "k \ K" "j \ K" "l \ L j" "l \ L k" have "k = j" proof (rule ccontr) assume "k \ j" with disjoint \K \ J\ \k \ K\ \j \ K\ have "I k \ I j = {}" unfolding disjoint_family_on_def by auto with L(2,3)[OF \j \ K\] L(2,3)[OF \k \ K\] show False using \l \ L k\ \l \ L j\ by auto qed } note L_inj = this define k where "k l = (SOME k. k \ K \ l \ L k)" for l { fix x j l assume *: "j \ K" "l \ L j" have "k l = j" unfolding k_def proof (rule some_equality) fix k assume "k \ K \ l \ L k" with * L_inj show "k = j" by auto qed (insert *, simp) } note k_simp[simp] = this let ?E' = "\l. E' (k l) l" have "prob (\j\K. A j) = prob (\l\(\k\K. L k). ?E' l)" by (auto simp: A intro!: arg_cong[where f=prob]) also have "\ = (\l\(\k\K. L k). prob (?E' l))" using L K E' by (intro indep_setsD[OF indep]) (simp_all add: UN_mono) also have "\ = (\j\K. \l\L j. prob (E' j l))" using K L L_inj by (subst prod.UNION_disjoint) auto also have "\ = (\j\K. prob (A j))" using K L E' by (auto simp add: A intro!: prod.cong indep_setsD[OF indep, symmetric]) blast finally show "prob (\j\K. A j) = (\j\K. prob (A j))" . qed next fix j assume "j \ J" show "Int_stable (?E j)" proof (rule Int_stableI) fix a assume "a \ ?E j" then obtain Ka Ea where a: "a = (\k\Ka. Ea k)" "finite Ka" "Ka \ {}" "Ka \ I j" "\k. k\Ka \ Ea k \ E k" by auto fix b assume "b \ ?E j" then obtain Kb Eb where b: "b = (\k\Kb. Eb k)" "finite Kb" "Kb \ {}" "Kb \ I j" "\k. k\Kb \ Eb k \ E k" by auto let ?f = "\k. (if k \ Ka \ Kb then Ea k \ Eb k else if k \ Kb then Eb k else if k \ Ka then Ea k else {})" have "Ka \ Kb = (Ka \ Kb) \ (Kb - Ka) \ (Ka - Kb)" by blast moreover have "(\x\Ka \ Kb. Ea x \ Eb x) \ (\x\Kb - Ka. Eb x) \ (\x\Ka - Kb. Ea x) = (\k\Ka. Ea k) \ (\k\Kb. Eb k)" by auto ultimately have "(\k\Ka \ Kb. ?f k) = (\k\Ka. Ea k) \ (\k\Kb. Eb k)" (is "?lhs = ?rhs") by (simp only: image_Un Inter_Un_distrib) simp then have "a \ b = (\k\Ka \ Kb. ?f k)" by (simp only: a(1) b(1)) with a b \j \ J\ Int_stableD[OF Int_stable] show "a \ b \ ?E j" by (intro CollectI exI[of _ "Ka \ Kb"] exI[of _ ?f]) auto qed qed ultimately show ?thesis by (simp cong: indep_sets_cong) qed lemma (in prob_space) indep_vars_restrict: assumes ind: "indep_vars M' X I" and K: "\j. j \ L \ K j \ I" and J: "disjoint_family_on K L" shows "indep_vars (\j. PiM (K j) M') (\j \. restrict (\i. X i \) (K j)) L" unfolding indep_vars_def proof safe fix j assume "j \ L" then show "random_variable (Pi\<^sub>M (K j) M') (\\. \i\K j. X i \)" using K ind by (auto simp: indep_vars_def intro!: measurable_restrict) next have X: "\i. i \ I \ X i \ measurable M (M' i)" using ind by (auto simp: indep_vars_def) let ?proj = "\j S. {(\\. \i\K j. X i \) -` A \ space M |A. A \ S}" let ?UN = "\j. sigma_sets (space M) (\i\K j. { X i -` A \ space M| A. A \ sets (M' i) })" show "indep_sets (\i. sigma_sets (space M) (?proj i (sets (Pi\<^sub>M (K i) M')))) L" proof (rule indep_sets_mono_sets) fix j assume j: "j \ L" have "sigma_sets (space M) (?proj j (sets (Pi\<^sub>M (K j) M'))) = sigma_sets (space M) (sigma_sets (space M) (?proj j (prod_algebra (K j) M')))" using j K X[THEN measurable_space] unfolding sets_PiM by (subst sigma_sets_vimage_commute) (auto simp add: Pi_iff) also have "\ = sigma_sets (space M) (?proj j (prod_algebra (K j) M'))" by (rule sigma_sets_sigma_sets_eq) auto also have "\ \ ?UN j" proof (rule sigma_sets_mono, safe del: disjE elim!: prod_algebraE) fix J E assume J: "finite J" "J \ {} \ K j = {}" "J \ K j" and E: "\i. i \ J \ E i \ sets (M' i)" show "(\\. \i\K j. X i \) -` prod_emb (K j) M' J (Pi\<^sub>E J E) \ space M \ ?UN j" proof cases assume "K j = {}" with J show ?thesis by (auto simp add: sigma_sets_empty_eq prod_emb_def) next assume "K j \ {}" with J have "J \ {}" by auto { interpret sigma_algebra "space M" "?UN j" by (rule sigma_algebra_sigma_sets) auto have "\A. (\i. i \ J \ A i \ ?UN j) \ \(A ` J) \ ?UN j" using \finite J\ \J \ {}\ by (rule finite_INT) blast } note INT = this from \J \ {}\ J K E[rule_format, THEN sets.sets_into_space] j have "(\\. \i\K j. X i \) -` prod_emb (K j) M' J (Pi\<^sub>E J E) \ space M = (\i\J. X i -` E i \ space M)" apply (subst prod_emb_PiE[OF _ ]) apply auto [] apply auto [] apply (auto simp add: Pi_iff intro!: X[THEN measurable_space]) apply (erule_tac x=i in ballE) apply auto done also have "\ \ ?UN j" apply (rule INT) apply (rule sigma_sets.Basic) using \J \ K j\ E apply auto done finally show ?thesis . qed qed finally show "sigma_sets (space M) (?proj j (sets (Pi\<^sub>M (K j) M'))) \ ?UN j" . next show "indep_sets ?UN L" proof (rule indep_sets_collect_sigma) show "indep_sets (\i. {X i -` A \ space M |A. A \ sets (M' i)}) (\j\L. K j)" proof (rule indep_sets_mono_index) show "indep_sets (\i. {X i -` A \ space M |A. A \ sets (M' i)}) I" using ind unfolding indep_vars_def2 by auto show "(\l\L. K l) \ I" using K by auto qed next fix l i assume "l \ L" "i \ K l" show "Int_stable {X i -` A \ space M |A. A \ sets (M' i)}" apply (auto simp: Int_stable_def) apply (rule_tac x="A \ Aa" in exI) apply auto done qed fact qed qed lemma (in prob_space) indep_var_restrict: assumes ind: "indep_vars M' X I" and AB: "A \ B = {}" "A \ I" "B \ I" shows "indep_var (PiM A M') (\\. restrict (\i. X i \) A) (PiM B M') (\\. restrict (\i. X i \) B)" proof - have *: "case_bool (Pi\<^sub>M A M') (Pi\<^sub>M B M') = (\b. PiM (case_bool A B b) M')" "case_bool (\\. \i\A. X i \) (\\. \i\B. X i \) = (\b \. \i\case_bool A B b. X i \)" by (simp_all add: fun_eq_iff split: bool.split) show ?thesis unfolding indep_var_def * using AB by (intro indep_vars_restrict[OF ind]) (auto simp: disjoint_family_on_def split: bool.split) qed lemma (in prob_space) indep_vars_subset: assumes "indep_vars M' X I" "J \ I" shows "indep_vars M' X J" using assms unfolding indep_vars_def indep_sets_def by auto lemma (in prob_space) indep_vars_cong: "I = J \ (\i. i \ I \ X i = Y i) \ (\i. i \ I \ M' i = N' i) \ indep_vars M' X I \ indep_vars N' Y J" unfolding indep_vars_def2 by (intro conj_cong indep_sets_cong) auto definition (in prob_space) tail_events where "tail_events A = (\n. sigma_sets (space M) (\ (A ` {n..})))" lemma (in prob_space) tail_events_sets: assumes A: "\i::nat. A i \ events" shows "tail_events A \ events" proof fix X assume X: "X \ tail_events A" let ?A = "(\n. sigma_sets (space M) (\ (A ` {n..})))" from X have "\n::nat. X \ sigma_sets (space M) (\ (A ` {n..}))" by (auto simp: tail_events_def) from this[of 0] have "X \ sigma_sets (space M) (\(A ` UNIV))" by simp then show "X \ events" by induct (insert A, auto) qed lemma (in prob_space) sigma_algebra_tail_events: assumes "\i::nat. sigma_algebra (space M) (A i)" shows "sigma_algebra (space M) (tail_events A)" unfolding tail_events_def proof (simp add: sigma_algebra_iff2, safe) let ?A = "(\n. sigma_sets (space M) (\ (A ` {n..})))" interpret A: sigma_algebra "space M" "A i" for i by fact { fix X x assume "X \ ?A" "x \ X" then have "\n. X \ sigma_sets (space M) (\ (A ` {n..}))" by auto from this[of 0] have "X \ sigma_sets (space M) (\(A ` UNIV))" by simp then have "X \ space M" by induct (insert A.sets_into_space, auto) with \x \ X\ show "x \ space M" by auto } { fix F :: "nat \ 'a set" and n assume "range F \ ?A" then show "(\(F ` UNIV)) \ sigma_sets (space M) (\ (A ` {n..}))" by (intro sigma_sets.Union) auto } qed (auto intro!: sigma_sets.Compl sigma_sets.Empty) lemma (in prob_space) kolmogorov_0_1_law: fixes A :: "nat \ 'a set set" assumes "\i::nat. sigma_algebra (space M) (A i)" assumes indep: "indep_sets A UNIV" and X: "X \ tail_events A" shows "prob X = 0 \ prob X = 1" proof - have A: "\i. A i \ events" using indep unfolding indep_sets_def by simp let ?D = "{D \ events. prob (X \ D) = prob X * prob D}" interpret A: sigma_algebra "space M" "A i" for i by fact interpret T: sigma_algebra "space M" "tail_events A" by (rule sigma_algebra_tail_events) fact have "X \ space M" using T.space_closed X by auto have X_in: "X \ events" using tail_events_sets A X by auto interpret D: Dynkin_system "space M" ?D proof (rule Dynkin_systemI) fix D assume "D \ ?D" then show "D \ space M" using sets.sets_into_space by auto next show "space M \ ?D" using prob_space \X \ space M\ by (simp add: Int_absorb2) next fix A assume A: "A \ ?D" have "prob (X \ (space M - A)) = prob (X - (X \ A))" using \X \ space M\ by (auto intro!: arg_cong[where f=prob]) also have "\ = prob X - prob (X \ A)" using X_in A by (intro finite_measure_Diff) auto also have "\ = prob X * prob (space M) - prob X * prob A" using A prob_space by auto also have "\ = prob X * prob (space M - A)" using X_in A sets.sets_into_space by (subst finite_measure_Diff) (auto simp: field_simps) finally show "space M - A \ ?D" using A \X \ space M\ by auto next fix F :: "nat \ 'a set" assume dis: "disjoint_family F" and "range F \ ?D" then have F: "range F \ events" "\i. prob (X \ F i) = prob X * prob (F i)" by auto have "(\i. prob (X \ F i)) sums prob (\i. X \ F i)" proof (rule finite_measure_UNION) show "range (\i. X \ F i) \ events" using F X_in by auto show "disjoint_family (\i. X \ F i)" using dis by (rule disjoint_family_on_bisimulation) auto qed with F have "(\i. prob X * prob (F i)) sums prob (X \ (\i. F i))" by simp moreover have "(\i. prob X * prob (F i)) sums (prob X * prob (\i. F i))" by (intro sums_mult finite_measure_UNION F dis) ultimately have "prob (X \ (\i. F i)) = prob X * prob (\i. F i)" by (auto dest!: sums_unique) with F show "(\i. F i) \ ?D" by auto qed { fix n have "indep_sets (\b. sigma_sets (space M) (\m\case_bool {..n} {Suc n..} b. A m)) UNIV" proof (rule indep_sets_collect_sigma) have *: "(\b. case b of True \ {..n} | False \ {Suc n..}) = UNIV" (is "?U = _") by (simp split: bool.split add: set_eq_iff) (metis not_less_eq_eq) with indep show "indep_sets A ?U" by simp show "disjoint_family (case_bool {..n} {Suc n..})" unfolding disjoint_family_on_def by (auto split: bool.split) fix m show "Int_stable (A m)" unfolding Int_stable_def using A.Int by auto qed also have "(\b. sigma_sets (space M) (\m\case_bool {..n} {Suc n..} b. A m)) = case_bool (sigma_sets (space M) (\m\{..n}. A m)) (sigma_sets (space M) (\m\{Suc n..}. A m))" by (auto intro!: ext split: bool.split) finally have indep: "indep_set (sigma_sets (space M) (\m\{..n}. A m)) (sigma_sets (space M) (\m\{Suc n..}. A m))" unfolding indep_set_def by simp have "sigma_sets (space M) (\m\{..n}. A m) \ ?D" proof (simp add: subset_eq, rule) fix D assume D: "D \ sigma_sets (space M) (\m\{..n}. A m)" have "X \ sigma_sets (space M) (\m\{Suc n..}. A m)" using X unfolding tail_events_def by simp from indep_setD[OF indep D this] indep_setD_ev1[OF indep] D show "D \ events \ prob (X \ D) = prob X * prob D" by (auto simp add: ac_simps) qed } then have "(\n. sigma_sets (space M) (\m\{..n}. A m)) \ ?D" (is "?A \ _") by auto note \X \ tail_events A\ also { have "\n. sigma_sets (space M) (\i\{n..}. A i) \ sigma_sets (space M) ?A" by (intro sigma_sets_subseteq UN_mono) auto then have "tail_events A \ sigma_sets (space M) ?A" unfolding tail_events_def by auto } also have "sigma_sets (space M) ?A = Dynkin (space M) ?A" proof (rule sigma_eq_Dynkin) { fix B n assume "B \ sigma_sets (space M) (\m\{..n}. A m)" then have "B \ space M" by induct (insert A sets.sets_into_space[of _ M], auto) } then show "?A \ Pow (space M)" by auto show "Int_stable ?A" proof (rule Int_stableI) - fix a assume "a \ ?A" then guess n .. note a = this - fix b assume "b \ ?A" then guess m .. note b = this + fix a b assume "a \ ?A" "b \ ?A" then obtain n m + where a: "n \ UNIV" "a \ sigma_sets (space M) (\ (A ` {..n}))" + and b: "m \ UNIV" "b \ sigma_sets (space M) (\ (A ` {..m}))" by auto interpret Amn: sigma_algebra "space M" "sigma_sets (space M) (\i\{..max m n}. A i)" using A sets.sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto have "sigma_sets (space M) (\i\{..n}. A i) \ sigma_sets (space M) (\i\{..max m n}. A i)" by (intro sigma_sets_subseteq UN_mono) auto with a have "a \ sigma_sets (space M) (\i\{..max m n}. A i)" by auto moreover have "sigma_sets (space M) (\i\{..m}. A i) \ sigma_sets (space M) (\i\{..max m n}. A i)" by (intro sigma_sets_subseteq UN_mono) auto with b have "b \ sigma_sets (space M) (\i\{..max m n}. A i)" by auto ultimately have "a \ b \ sigma_sets (space M) (\i\{..max m n}. A i)" using Amn.Int[of a b] by simp then show "a \ b \ (\n. sigma_sets (space M) (\i\{..n}. A i))" by auto qed qed also have "Dynkin (space M) ?A \ ?D" using \?A \ ?D\ by (auto intro!: D.Dynkin_subset) finally show ?thesis by auto qed lemma (in prob_space) borel_0_1_law: fixes F :: "nat \ 'a set" assumes F2: "indep_events F UNIV" shows "prob (\n. \m\{n..}. F m) = 0 \ prob (\n. \m\{n..}. F m) = 1" proof (rule kolmogorov_0_1_law[of "\i. sigma_sets (space M) { F i }"]) have F1: "range F \ events" using F2 by (simp add: indep_events_def subset_eq) { fix i show "sigma_algebra (space M) (sigma_sets (space M) {F i})" using sigma_algebra_sigma_sets[of "{F i}" "space M"] F1 sets.sets_into_space by auto } show "indep_sets (\i. sigma_sets (space M) {F i}) UNIV" proof (rule indep_sets_sigma) show "indep_sets (\i. {F i}) UNIV" unfolding indep_events_def_alt[symmetric] by fact fix i show "Int_stable {F i}" unfolding Int_stable_def by simp qed let ?Q = "\n. \i\{n..}. F i" show "(\n. \m\{n..}. F m) \ tail_events (\i. sigma_sets (space M) {F i})" unfolding tail_events_def proof fix j interpret S: sigma_algebra "space M" "sigma_sets (space M) (\i\{j..}. sigma_sets (space M) {F i})" using order_trans[OF F1 sets.space_closed] by (intro sigma_algebra_sigma_sets) (simp add: sigma_sets_singleton subset_eq) have "(\n. ?Q n) = (\n\{j..}. ?Q n)" by (intro decseq_SucI INT_decseq_offset UN_mono) auto also have "\ \ sigma_sets (space M) (\i\{j..}. sigma_sets (space M) {F i})" using order_trans[OF F1 sets.space_closed] by (safe intro!: S.countable_INT S.countable_UN) (auto simp: sigma_sets_singleton intro!: sigma_sets.Basic bexI) finally show "(\n. ?Q n) \ sigma_sets (space M) (\i\{j..}. sigma_sets (space M) {F i})" by simp qed qed lemma (in prob_space) borel_0_1_law_AE: fixes P :: "nat \ 'a \ bool" assumes "indep_events (\m. {x\space M. P m x}) UNIV" (is "indep_events ?P _") shows "(AE x in M. infinite {m. P m x}) \ (AE x in M. finite {m. P m x})" proof - have [measurable]: "\m. {x\space M. P m x} \ sets M" using assms by (auto simp: indep_events_def) have *: "(\n. \m\{n..}. {x \ space M. P m x}) \ events" by simp from assms have "prob (\n. \m\{n..}. ?P m) = 0 \ prob (\n. \m\{n..}. ?P m) = 1" by (rule borel_0_1_law) also have "prob (\n. \m\{n..}. ?P m) = 1 \ (AE x in M. infinite {m. P m x})" using * by (simp add: prob_eq_1) (simp add: Bex_def infinite_nat_iff_unbounded_le) also have "prob (\n. \m\{n..}. ?P m) = 0 \ (AE x in M. finite {m. P m x})" using * by (simp add: prob_eq_0) (auto simp add: Ball_def finite_nat_iff_bounded not_less [symmetric]) finally show ?thesis by blast qed lemma (in prob_space) indep_sets_finite: assumes I: "I \ {}" "finite I" and F: "\i. i \ I \ F i \ events" "\i. i \ I \ space M \ F i" shows "indep_sets F I \ (\A\Pi I F. prob (\j\I. A j) = (\j\I. prob (A j)))" proof assume *: "indep_sets F I" from I show "\A\Pi I F. prob (\j\I. A j) = (\j\I. prob (A j))" by (intro indep_setsD[OF *] ballI) auto next assume indep: "\A\Pi I F. prob (\j\I. A j) = (\j\I. prob (A j))" show "indep_sets F I" proof (rule indep_setsI[OF F(1)]) fix A J assume J: "J \ {}" "J \ I" "finite J" assume A: "\j\J. A j \ F j" let ?A = "\j. if j \ J then A j else space M" have "prob (\j\I. ?A j) = prob (\j\J. A j)" using subset_trans[OF F(1) sets.space_closed] J A by (auto intro!: arg_cong[where f=prob] split: if_split_asm) blast also from A F have "(\j. if j \ J then A j else space M) \ Pi I F" (is "?A \ _") by (auto split: if_split_asm) with indep have "prob (\j\I. ?A j) = (\j\I. prob (?A j))" by auto also have "\ = (\j\J. prob (A j))" unfolding if_distrib prod.If_cases[OF \finite I\] using prob_space \J \ I\ by (simp add: Int_absorb1 prod.neutral_const) finally show "prob (\j\J. A j) = (\j\J. prob (A j))" .. qed qed lemma (in prob_space) indep_vars_finite: fixes I :: "'i set" assumes I: "I \ {}" "finite I" and M': "\i. i \ I \ sets (M' i) = sigma_sets (space (M' i)) (E i)" and rv: "\i. i \ I \ random_variable (M' i) (X i)" and Int_stable: "\i. i \ I \ Int_stable (E i)" and space: "\i. i \ I \ space (M' i) \ E i" and closed: "\i. i \ I \ E i \ Pow (space (M' i))" shows "indep_vars M' X I \ (\A\(\ i\I. E i). prob (\j\I. X j -` A j \ space M) = (\j\I. prob (X j -` A j \ space M)))" proof - from rv have X: "\i. i \ I \ X i \ space M \ space (M' i)" unfolding measurable_def by simp { fix i assume "i\I" from closed[OF \i \ I\] have "sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)} = sigma_sets (space M) {X i -` A \ space M |A. A \ E i}" unfolding sigma_sets_vimage_commute[OF X, OF \i \ I\, symmetric] M'[OF \i \ I\] by (subst sigma_sets_sigma_sets_eq) auto } note sigma_sets_X = this { fix i assume "i\I" have "Int_stable {X i -` A \ space M |A. A \ E i}" proof (rule Int_stableI) fix a assume "a \ {X i -` A \ space M |A. A \ E i}" then obtain A where "a = X i -` A \ space M" "A \ E i" by auto moreover fix b assume "b \ {X i -` A \ space M |A. A \ E i}" then obtain B where "b = X i -` B \ space M" "B \ E i" by auto moreover have "(X i -` A \ space M) \ (X i -` B \ space M) = X i -` (A \ B) \ space M" by auto moreover note Int_stable[OF \i \ I\] ultimately show "a \ b \ {X i -` A \ space M |A. A \ E i}" by (auto simp del: vimage_Int intro!: exI[of _ "A \ B"] dest: Int_stableD) qed } note indep_sets_X = indep_sets_sigma_sets_iff[OF this] { fix i assume "i \ I" { fix A assume "A \ E i" with M'[OF \i \ I\] have "A \ sets (M' i)" by auto moreover from rv[OF \i\I\] have "X i \ measurable M (M' i)" by auto ultimately have "X i -` A \ space M \ sets M" by (auto intro: measurable_sets) } with X[OF \i\I\] space[OF \i\I\] have "{X i -` A \ space M |A. A \ E i} \ events" "space M \ {X i -` A \ space M |A. A \ E i}" by (auto intro!: exI[of _ "space (M' i)"]) } note indep_sets_finite_X = indep_sets_finite[OF I this] have "(\A\\ i\I. {X i -` A \ space M |A. A \ E i}. prob (\(A ` I)) = (\j\I. prob (A j))) = (\A\\ i\I. E i. prob ((\j\I. X j -` A j) \ space M) = (\x\I. prob (X x -` A x \ space M)))" (is "?L = ?R") proof safe fix A assume ?L and A: "A \ (\ i\I. E i)" from \?L\[THEN bspec, of "\i. X i -` A i \ space M"] A \I \ {}\ show "prob ((\j\I. X j -` A j) \ space M) = (\x\I. prob (X x -` A x \ space M))" by (auto simp add: Pi_iff) next fix A assume ?R and A: "A \ (\ i\I. {X i -` A \ space M |A. A \ E i})" from A have "\i\I. \B. A i = X i -` B \ space M \ B \ E i" by auto from bchoice[OF this] obtain B where B: "\i\I. A i = X i -` B i \ space M" "B \ (\ i\I. E i)" by auto from \?R\[THEN bspec, OF B(2)] B(1) \I \ {}\ show "prob (\(A ` I)) = (\j\I. prob (A j))" by simp qed then show ?thesis using \I \ {}\ by (simp add: rv indep_vars_def indep_sets_X sigma_sets_X indep_sets_finite_X cong: indep_sets_cong) qed lemma (in prob_space) indep_vars_compose: assumes "indep_vars M' X I" assumes rv: "\i. i \ I \ Y i \ measurable (M' i) (N i)" shows "indep_vars N (\i. Y i \ X i) I" unfolding indep_vars_def proof from rv \indep_vars M' X I\ show "\i\I. random_variable (N i) (Y i \ X i)" by (auto simp: indep_vars_def) have "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" using \indep_vars M' X I\ by (simp add: indep_vars_def) then show "indep_sets (\i. sigma_sets (space M) {(Y i \ X i) -` A \ space M |A. A \ sets (N i)}) I" proof (rule indep_sets_mono_sets) fix i assume "i \ I" with \indep_vars M' X I\ have X: "X i \ space M \ space (M' i)" unfolding indep_vars_def measurable_def by auto { fix A assume "A \ sets (N i)" then have "\B. (Y i \ X i) -` A \ space M = X i -` B \ space M \ B \ sets (M' i)" by (intro exI[of _ "Y i -` A \ space (M' i)"]) (auto simp: vimage_comp intro!: measurable_sets rv \i \ I\ funcset_mem[OF X]) } then show "sigma_sets (space M) {(Y i \ X i) -` A \ space M |A. A \ sets (N i)} \ sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" by (intro sigma_sets_subseteq) (auto simp: vimage_comp) qed qed lemma (in prob_space) indep_vars_compose2: assumes "indep_vars M' X I" assumes rv: "\i. i \ I \ Y i \ measurable (M' i) (N i)" shows "indep_vars N (\i x. Y i (X i x)) I" using indep_vars_compose [OF assms] by (simp add: comp_def) lemma (in prob_space) indep_var_compose: assumes "indep_var M1 X1 M2 X2" "Y1 \ measurable M1 N1" "Y2 \ measurable M2 N2" shows "indep_var N1 (Y1 \ X1) N2 (Y2 \ X2)" proof - have "indep_vars (case_bool N1 N2) (\b. case_bool Y1 Y2 b \ case_bool X1 X2 b) UNIV" using assms by (intro indep_vars_compose[where M'="case_bool M1 M2"]) (auto simp: indep_var_def split: bool.split) also have "(\b. case_bool Y1 Y2 b \ case_bool X1 X2 b) = case_bool (Y1 \ X1) (Y2 \ X2)" by (simp add: fun_eq_iff split: bool.split) finally show ?thesis unfolding indep_var_def . qed lemma (in prob_space) indep_vars_Min: fixes X :: "'i \ 'a \ real" assumes I: "finite I" "i \ I" and indep: "indep_vars (\_. borel) X (insert i I)" shows "indep_var borel (X i) borel (\\. Min ((\i. X i \)`I))" proof - have "indep_var borel ((\f. f i) \ (\\. restrict (\i. X i \) {i})) borel ((\f. Min (f`I)) \ (\\. restrict (\i. X i \) I))" using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] borel_measurable_Min) auto also have "((\f. f i) \ (\\. restrict (\i. X i \) {i})) = X i" by auto also have "((\f. Min (f`I)) \ (\\. restrict (\i. X i \) I)) = (\\. Min ((\i. X i \)`I))" by (auto cong: rev_conj_cong) finally show ?thesis unfolding indep_var_def . qed lemma (in prob_space) indep_vars_sum: fixes X :: "'i \ 'a \ real" assumes I: "finite I" "i \ I" and indep: "indep_vars (\_. borel) X (insert i I)" shows "indep_var borel (X i) borel (\\. \i\I. X i \)" proof - have "indep_var borel ((\f. f i) \ (\\. restrict (\i. X i \) {i})) borel ((\f. \i\I. f i) \ (\\. restrict (\i. X i \) I))" using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] ) auto also have "((\f. f i) \ (\\. restrict (\i. X i \) {i})) = X i" by auto also have "((\f. \i\I. f i) \ (\\. restrict (\i. X i \) I)) = (\\. \i\I. X i \)" by (auto cong: rev_conj_cong) finally show ?thesis . qed lemma (in prob_space) indep_vars_prod: fixes X :: "'i \ 'a \ real" assumes I: "finite I" "i \ I" and indep: "indep_vars (\_. borel) X (insert i I)" shows "indep_var borel (X i) borel (\\. \i\I. X i \)" proof - have "indep_var borel ((\f. f i) \ (\\. restrict (\i. X i \) {i})) borel ((\f. \i\I. f i) \ (\\. restrict (\i. X i \) I))" using I by (intro indep_var_compose[OF indep_var_restrict[OF indep]] ) auto also have "((\f. f i) \ (\\. restrict (\i. X i \) {i})) = X i" by auto also have "((\f. \i\I. f i) \ (\\. restrict (\i. X i \) I)) = (\\. \i\I. X i \)" by (auto cong: rev_conj_cong) finally show ?thesis . qed lemma (in prob_space) indep_varsD_finite: assumes X: "indep_vars M' X I" assumes I: "I \ {}" "finite I" "\i. i \ I \ A i \ sets (M' i)" shows "prob (\i\I. X i -` A i \ space M) = (\i\I. prob (X i -` A i \ space M))" proof (rule indep_setsD) show "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" using X by (auto simp: indep_vars_def) show "I \ I" "I \ {}" "finite I" using I by auto show "\i\I. X i -` A i \ space M \ sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" using I by auto qed lemma (in prob_space) indep_varsD: assumes X: "indep_vars M' X I" assumes I: "J \ {}" "finite J" "J \ I" "\i. i \ J \ A i \ sets (M' i)" shows "prob (\i\J. X i -` A i \ space M) = (\i\J. prob (X i -` A i \ space M))" proof (rule indep_setsD) show "indep_sets (\i. sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}) I" using X by (auto simp: indep_vars_def) show "\i\J. X i -` A i \ space M \ sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)}" using I by auto qed fact+ lemma (in prob_space) indep_vars_iff_distr_eq_PiM: fixes I :: "'i set" and X :: "'i \ 'a \ 'b" assumes "I \ {}" assumes rv: "\i. random_variable (M' i) (X i)" shows "indep_vars M' X I \ distr M (\\<^sub>M i\I. M' i) (\x. \i\I. X i x) = (\\<^sub>M i\I. distr M (M' i) (X i))" proof - let ?P = "\\<^sub>M i\I. M' i" let ?X = "\x. \i\I. X i x" let ?D = "distr M ?P ?X" have X: "random_variable ?P ?X" by (intro measurable_restrict rv) interpret D: prob_space ?D by (intro prob_space_distr X) let ?D' = "\i. distr M (M' i) (X i)" let ?P' = "\\<^sub>M i\I. distr M (M' i) (X i)" interpret D': prob_space "?D' i" for i by (intro prob_space_distr rv) interpret P: product_prob_space ?D' I .. show ?thesis proof assume "indep_vars M' X I" show "?D = ?P'" proof (rule measure_eqI_generator_eq) show "Int_stable (prod_algebra I M')" by (rule Int_stable_prod_algebra) show "prod_algebra I M' \ Pow (space ?P)" using prod_algebra_sets_into_space by (simp add: space_PiM) show "sets ?D = sigma_sets (space ?P) (prod_algebra I M')" by (simp add: sets_PiM space_PiM) show "sets ?P' = sigma_sets (space ?P) (prod_algebra I M')" by (simp add: sets_PiM space_PiM cong: prod_algebra_cong) let ?A = "\i. \\<^sub>E i\I. space (M' i)" show "range ?A \ prod_algebra I M'" "(\i. ?A i) = space (Pi\<^sub>M I M')" by (auto simp: space_PiM intro!: space_in_prod_algebra cong: prod_algebra_cong) { fix i show "emeasure ?D (\\<^sub>E i\I. space (M' i)) \ \" by auto } next fix E assume E: "E \ prod_algebra I M'" - from prod_algebraE[OF E] guess J Y . note J = this - + from prod_algebraE[OF E] obtain J Y + where J: + "E = prod_emb I M' J (Pi\<^sub>E J Y)" + "finite J" + "J \ {} \ I = {}" + "J \ I" + "\i. i \ J \ Y i \ sets (M' i)" + by auto from E have "E \ sets ?P" by (auto simp: sets_PiM) then have "emeasure ?D E = emeasure M (?X -` E \ space M)" by (simp add: emeasure_distr X) also have "?X -` E \ space M = (\i\J. X i -` Y i \ space M)" using J \I \ {}\ measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: if_split_asm) also have "emeasure M (\i\J. X i -` Y i \ space M) = (\ i\J. emeasure M (X i -` Y i \ space M))" using \indep_vars M' X I\ J \I \ {}\ using indep_varsD[of M' X I J] by (auto simp: emeasure_eq_measure prod_ennreal measure_nonneg prod_nonneg) also have "\ = (\ i\J. emeasure (?D' i) (Y i))" using rv J by (simp add: emeasure_distr) also have "\ = emeasure ?P' E" using P.emeasure_PiM_emb[of J Y] J by (simp add: prod_emb_def) finally show "emeasure ?D E = emeasure ?P' E" . qed next assume "?D = ?P'" show "indep_vars M' X I" unfolding indep_vars_def proof (intro conjI indep_setsI ballI rv) fix i show "sigma_sets (space M) {X i -` A \ space M |A. A \ sets (M' i)} \ events" by (auto intro!: sets.sigma_sets_subset measurable_sets rv) next fix J Y' assume J: "J \ {}" "J \ I" "finite J" assume Y': "\j\J. Y' j \ sigma_sets (space M) {X j -` A \ space M |A. A \ sets (M' j)}" have "\j\J. \Y. Y' j = X j -` Y \ space M \ Y \ sets (M' j)" proof fix j assume "j \ J" from Y'[rule_format, OF this] rv[of j] show "\Y. Y' j = X j -` Y \ space M \ Y \ sets (M' j)" by (subst (asm) sigma_sets_vimage_commute[symmetric, of _ _ "space (M' j)"]) (auto dest: measurable_space simp: sets.sigma_sets_eq) qed from bchoice[OF this] obtain Y where Y: "\j. j \ J \ Y' j = X j -` Y j \ space M" "\j. j \ J \ Y j \ sets (M' j)" by auto let ?E = "prod_emb I M' J (Pi\<^sub>E J Y)" from Y have "(\j\J. Y' j) = ?X -` ?E \ space M" using J \I \ {}\ measurable_space[OF rv] by (auto simp: prod_emb_def PiE_iff split: if_split_asm) then have "emeasure M (\j\J. Y' j) = emeasure M (?X -` ?E \ space M)" by simp also have "\ = emeasure ?D ?E" using Y J by (intro emeasure_distr[symmetric] X sets_PiM_I) auto also have "\ = emeasure ?P' ?E" using \?D = ?P'\ by simp also have "\ = (\ i\J. emeasure (?D' i) (Y i))" using P.emeasure_PiM_emb[of J Y] J Y by (simp add: prod_emb_def) also have "\ = (\ i\J. emeasure M (Y' i))" using rv J Y by (simp add: emeasure_distr) finally have "emeasure M (\j\J. Y' j) = (\ i\J. emeasure M (Y' i))" . then show "prob (\j\J. Y' j) = (\ i\J. prob (Y' i))" by (auto simp: emeasure_eq_measure prod_ennreal measure_nonneg prod_nonneg) qed qed qed lemma (in prob_space) indep_vars_iff_distr_eq_PiM': fixes I :: "'i set" and X :: "'i \ 'a \ 'b" assumes "I \ {}" assumes rv: "\i. i \ I \ random_variable (M' i) (X i)" shows "indep_vars M' X I \ distr M (\\<^sub>M i\I. M' i) (\x. \i\I. X i x) = (\\<^sub>M i\I. distr M (M' i) (X i))" proof - from assms obtain j where j: "j \ I" by auto define N' where "N' = (\i. if i \ I then M' i else M' j)" define Y where "Y = (\i. if i \ I then X i else X j)" have rv: "random_variable (N' i) (Y i)" for i using j by (auto simp: N'_def Y_def intro: assms) have "indep_vars M' X I = indep_vars N' Y I" by (intro indep_vars_cong) (auto simp: N'_def Y_def) also have "\ \ distr M (\\<^sub>M i\I. N' i) (\x. \i\I. Y i x) = (\\<^sub>M i\I. distr M (N' i) (Y i))" by (intro indep_vars_iff_distr_eq_PiM rv assms) also have "(\\<^sub>M i\I. N' i) = (\\<^sub>M i\I. M' i)" by (intro PiM_cong) (simp_all add: N'_def) also have "(\x. \i\I. Y i x) = (\x. \i\I. X i x)" by (simp_all add: Y_def fun_eq_iff) also have "(\\<^sub>M i\I. distr M (N' i) (Y i)) = (\\<^sub>M i\I. distr M (M' i) (X i))" by (intro PiM_cong distr_cong) (simp_all add: N'_def Y_def) finally show ?thesis . qed lemma (in prob_space) indep_varD: assumes indep: "indep_var Ma A Mb B" assumes sets: "Xa \ sets Ma" "Xb \ sets Mb" shows "prob ((\x. (A x, B x)) -` (Xa \ Xb) \ space M) = prob (A -` Xa \ space M) * prob (B -` Xb \ space M)" proof - have "prob ((\x. (A x, B x)) -` (Xa \ Xb) \ space M) = prob (\i\UNIV. (case_bool A B i -` case_bool Xa Xb i \ space M))" by (auto intro!: arg_cong[where f=prob] simp: UNIV_bool) also have "\ = (\i\UNIV. prob (case_bool A B i -` case_bool Xa Xb i \ space M))" using indep unfolding indep_var_def by (rule indep_varsD) (auto split: bool.split intro: sets) also have "\ = prob (A -` Xa \ space M) * prob (B -` Xb \ space M)" unfolding UNIV_bool by simp finally show ?thesis . qed lemma (in prob_space) prob_indep_random_variable: assumes ind[simp]: "indep_var N X N Y" assumes [simp]: "A \ sets N" "B \ sets N" shows "\

(x in M. X x \ A \ Y x \ B) = \

(x in M. X x \ A) * \

(x in M. Y x \ B)" proof- have " \

(x in M. (X x)\A \ (Y x)\ B ) = prob ((\x. (X x, Y x)) -` (A \ B) \ space M)" by (auto intro!: arg_cong[where f= prob]) also have "...= prob (X -` A \ space M) * prob (Y -` B \ space M)" by (auto intro!: indep_varD[where Ma=N and Mb=N]) also have "... = \

(x in M. X x \ A) * \

(x in M. Y x \ B)" by (auto intro!: arg_cong2[where f= "(*)"] arg_cong[where f= prob]) finally show ?thesis . qed lemma (in prob_space) assumes "indep_var S X T Y" shows indep_var_rv1: "random_variable S X" and indep_var_rv2: "random_variable T Y" proof - have "\i\UNIV. random_variable (case_bool S T i) (case_bool X Y i)" using assms unfolding indep_var_def indep_vars_def by auto then show "random_variable S X" "random_variable T Y" unfolding UNIV_bool by auto qed lemma (in prob_space) indep_var_distribution_eq: "indep_var S X T Y \ random_variable S X \ random_variable T Y \ distr M S X \\<^sub>M distr M T Y = distr M (S \\<^sub>M T) (\x. (X x, Y x))" (is "_ \ _ \ _ \ ?S \\<^sub>M ?T = ?J") proof safe assume "indep_var S X T Y" then show rvs: "random_variable S X" "random_variable T Y" by (blast dest: indep_var_rv1 indep_var_rv2)+ then have XY: "random_variable (S \\<^sub>M T) (\x. (X x, Y x))" by (rule measurable_Pair) interpret X: prob_space ?S by (rule prob_space_distr) fact interpret Y: prob_space ?T by (rule prob_space_distr) fact interpret XY: pair_prob_space ?S ?T .. show "?S \\<^sub>M ?T = ?J" proof (rule pair_measure_eqI) show "sigma_finite_measure ?S" .. show "sigma_finite_measure ?T" .. fix A B assume A: "A \ sets ?S" and B: "B \ sets ?T" have "emeasure ?J (A \ B) = emeasure M ((\x. (X x, Y x)) -` (A \ B) \ space M)" using A B by (intro emeasure_distr[OF XY]) auto also have "\ = emeasure M (X -` A \ space M) * emeasure M (Y -` B \ space M)" using indep_varD[OF \indep_var S X T Y\, of A B] A B by (simp add: emeasure_eq_measure measure_nonneg ennreal_mult) also have "\ = emeasure ?S A * emeasure ?T B" using rvs A B by (simp add: emeasure_distr) finally show "emeasure ?S A * emeasure ?T B = emeasure ?J (A \ B)" by simp qed simp next assume rvs: "random_variable S X" "random_variable T Y" then have XY: "random_variable (S \\<^sub>M T) (\x. (X x, Y x))" by (rule measurable_Pair) let ?S = "distr M S X" and ?T = "distr M T Y" interpret X: prob_space ?S by (rule prob_space_distr) fact interpret Y: prob_space ?T by (rule prob_space_distr) fact interpret XY: pair_prob_space ?S ?T .. assume "?S \\<^sub>M ?T = ?J" { fix S and X have "Int_stable {X -` A \ space M |A. A \ sets S}" proof (safe intro!: Int_stableI) fix A B assume "A \ sets S" "B \ sets S" then show "\C. (X -` A \ space M) \ (X -` B \ space M) = (X -` C \ space M) \ C \ sets S" by (intro exI[of _ "A \ B"]) auto qed } note Int_stable = this show "indep_var S X T Y" unfolding indep_var_eq proof (intro conjI indep_set_sigma_sets Int_stable rvs) show "indep_set {X -` A \ space M |A. A \ sets S} {Y -` A \ space M |A. A \ sets T}" proof (safe intro!: indep_setI) { fix A assume "A \ sets S" then show "X -` A \ space M \ sets M" using \X \ measurable M S\ by (auto intro: measurable_sets) } { fix A assume "A \ sets T" then show "Y -` A \ space M \ sets M" using \Y \ measurable M T\ by (auto intro: measurable_sets) } next fix A B assume ab: "A \ sets S" "B \ sets T" then have "prob ((X -` A \ space M) \ (Y -` B \ space M)) = emeasure ?J (A \ B)" using XY by (auto simp add: emeasure_distr emeasure_eq_measure measure_nonneg intro!: arg_cong[where f="prob"]) also have "\ = emeasure (?S \\<^sub>M ?T) (A \ B)" unfolding \?S \\<^sub>M ?T = ?J\ .. also have "\ = emeasure ?S A * emeasure ?T B" using ab by (simp add: Y.emeasure_pair_measure_Times) finally show "prob ((X -` A \ space M) \ (Y -` B \ space M)) = prob (X -` A \ space M) * prob (Y -` B \ space M)" using rvs ab by (simp add: emeasure_eq_measure emeasure_distr measure_nonneg ennreal_mult[symmetric]) qed qed qed lemma (in prob_space) distributed_joint_indep: assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes X: "distributed M S X Px" and Y: "distributed M T Y Py" assumes indep: "indep_var S X T Y" shows "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) (\(x, y). Px x * Py y)" using indep_var_distribution_eq[of S X T Y] indep by (intro distributed_joint_indep'[OF S T X Y]) auto lemma (in prob_space) indep_vars_nn_integral: assumes I: "finite I" "indep_vars (\_. borel) X I" "\i \. i \ I \ 0 \ X i \" shows "(\\<^sup>+\. (\i\I. X i \) \M) = (\i\I. \\<^sup>+\. X i \ \M)" proof cases assume "I \ {}" define Y where [abs_def]: "Y i \ = (if i \ I then X i \ else 0)" for i \ { fix i have "i \ I \ random_variable borel (X i)" using I(2) by (cases "i\I") (auto simp: indep_vars_def) } note rv_X = this { fix i have "random_variable borel (Y i)" using I(2) by (cases "i\I") (auto simp: Y_def rv_X) } note rv_Y = this[measurable] interpret Y: prob_space "distr M borel (Y i)" for i using I(2) by (cases "i \ I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return) interpret product_sigma_finite "\i. distr M borel (Y i)" .. have indep_Y: "indep_vars (\i. borel) Y I" by (rule indep_vars_cong[THEN iffD1, OF _ _ _ I(2)]) (auto simp: Y_def) have "(\\<^sup>+\. (\i\I. X i \) \M) = (\\<^sup>+\. (\i\I. Y i \) \M)" using I(3) by (auto intro!: nn_integral_cong prod.cong simp add: Y_def max_def) also have "\ = (\\<^sup>+\. (\i\I. \ i) \distr M (Pi\<^sub>M I (\i. borel)) (\x. \i\I. Y i x))" by (subst nn_integral_distr) auto also have "\ = (\\<^sup>+\. (\i\I. \ i) \Pi\<^sub>M I (\i. distr M borel (Y i)))" unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \I \ {}\ rv_Y indep_Y] .. also have "\ = (\i\I. (\\<^sup>+\. \ \distr M borel (Y i)))" by (rule product_nn_integral_prod) (auto intro: \finite I\) also have "\ = (\i\I. \\<^sup>+\. X i \ \M)" by (intro prod.cong nn_integral_cong) (auto simp: nn_integral_distr Y_def rv_X) finally show ?thesis . qed (simp add: emeasure_space_1) lemma (in prob_space) fixes X :: "'i \ 'a \ 'b::{real_normed_field, banach, second_countable_topology}" assumes I: "finite I" "indep_vars (\_. borel) X I" "\i. i \ I \ integrable M (X i)" shows indep_vars_lebesgue_integral: "(\\. (\i\I. X i \) \M) = (\i\I. \\. X i \ \M)" (is ?eq) and indep_vars_integrable: "integrable M (\\. (\i\I. X i \))" (is ?int) proof (induct rule: case_split) assume "I \ {}" define Y where [abs_def]: "Y i \ = (if i \ I then X i \ else 0)" for i \ { fix i have "i \ I \ random_variable borel (X i)" using I(2) by (cases "i\I") (auto simp: indep_vars_def) } note rv_X = this[measurable] { fix i have "random_variable borel (Y i)" using I(2) by (cases "i\I") (auto simp: Y_def rv_X) } note rv_Y = this[measurable] { fix i have "integrable M (Y i)" using I(3) by (cases "i\I") (auto simp: Y_def) } note int_Y = this interpret Y: prob_space "distr M borel (Y i)" for i using I(2) by (cases "i \ I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return) interpret product_sigma_finite "\i. distr M borel (Y i)" .. have indep_Y: "indep_vars (\i. borel) Y I" by (rule indep_vars_cong[THEN iffD1, OF _ _ _ I(2)]) (auto simp: Y_def) have "(\\. (\i\I. X i \) \M) = (\\. (\i\I. Y i \) \M)" using I(3) by (simp add: Y_def) also have "\ = (\\. (\i\I. \ i) \distr M (Pi\<^sub>M I (\i. borel)) (\x. \i\I. Y i x))" by (subst integral_distr) auto also have "\ = (\\. (\i\I. \ i) \Pi\<^sub>M I (\i. distr M borel (Y i)))" unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \I \ {}\ rv_Y indep_Y] .. also have "\ = (\i\I. (\\. \ \distr M borel (Y i)))" by (rule product_integral_prod) (auto intro: \finite I\ simp: integrable_distr_eq int_Y) also have "\ = (\i\I. \\. X i \ \M)" by (intro prod.cong integral_cong) (auto simp: integral_distr Y_def rv_X) finally show ?eq . have "integrable (distr M (Pi\<^sub>M I (\i. borel)) (\x. \i\I. Y i x)) (\\. (\i\I. \ i))" unfolding indep_vars_iff_distr_eq_PiM[THEN iffD1, OF \I \ {}\ rv_Y indep_Y] by (intro product_integrable_prod[OF \finite I\]) (simp add: integrable_distr_eq int_Y) then show ?int by (simp add: integrable_distr_eq Y_def) qed (simp_all add: prob_space) lemma (in prob_space) fixes X1 X2 :: "'a \ 'b::{real_normed_field, banach, second_countable_topology}" assumes "indep_var borel X1 borel X2" "integrable M X1" "integrable M X2" shows indep_var_lebesgue_integral: "(\\. X1 \ * X2 \ \M) = (\\. X1 \ \M) * (\\. X2 \ \M)" (is ?eq) and indep_var_integrable: "integrable M (\\. X1 \ * X2 \)" (is ?int) unfolding indep_var_def proof - have *: "(\\. X1 \ * X2 \) = (\\. \i\UNIV. (case_bool X1 X2 i \))" by (simp add: UNIV_bool mult.commute) have **: "(\ _. borel) = case_bool borel borel" by (rule ext, metis (full_types) bool.simps(3) bool.simps(4)) show ?eq apply (subst *) apply (subst indep_vars_lebesgue_integral) apply (auto) apply (subst **, subst indep_var_def [symmetric], rule assms) apply (simp split: bool.split add: assms) by (simp add: UNIV_bool mult.commute) show ?int apply (subst *) apply (rule indep_vars_integrable) apply auto apply (subst **, subst indep_var_def [symmetric], rule assms) by (simp split: bool.split add: assms) qed end diff --git a/src/HOL/Probability/Information.thy b/src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy +++ b/src/HOL/Probability/Information.thy @@ -1,1968 +1,1974 @@ (* Title: HOL/Probability/Information.thy Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Information theory\ theory Information imports Independent_Family begin lemma log_le: "1 < a \ 0 < x \ x \ y \ log a x \ log a y" by (subst log_le_cancel_iff) auto lemma log_less: "1 < a \ 0 < x \ x < y \ log a x < log a y" by (subst log_less_cancel_iff) auto lemma sum_cartesian_product': "(\x\A \ B. f x) = (\x\A. sum (\y. f (x, y)) B)" unfolding sum.cartesian_product by simp lemma split_pairs: "((A, B) = X) \ (fst X = A \ snd X = B)" and "(X = (A, B)) \ (fst X = A \ snd X = B)" by auto subsection "Information theory" locale information_space = prob_space + fixes b :: real assumes b_gt_1: "1 < b" context information_space begin text \Introduce some simplification rules for logarithm of base \<^term>\b\.\ lemma log_neg_const: assumes "x \ 0" shows "log b x = log b 0" proof - { fix u :: real have "x \ 0" by fact also have "0 < exp u" using exp_gt_zero . finally have "exp u \ x" by auto } then show "log b x = log b 0" by (simp add: log_def ln_real_def) qed lemma log_mult_eq: "log b (A * B) = (if 0 < A * B then log b \A\ + log b \B\ else log b 0)" using log_mult[of b "\A\" "\B\"] b_gt_1 log_neg_const[of "A * B"] by (auto simp: zero_less_mult_iff mult_le_0_iff) lemma log_inverse_eq: "log b (inverse B) = (if 0 < B then - log b B else log b 0)" using log_inverse[of b B] log_neg_const[of "inverse B"] b_gt_1 by simp lemma log_divide_eq: "log b (A / B) = (if 0 < A * B then log b \A\ - log b \B\ else log b 0)" unfolding divide_inverse log_mult_eq log_inverse_eq abs_inverse by (auto simp: zero_less_mult_iff mult_le_0_iff) lemmas log_simps = log_mult_eq log_inverse_eq log_divide_eq end subsection "Kullback$-$Leibler divergence" text \The Kullback$-$Leibler divergence is also known as relative entropy or Kullback$-$Leibler distance.\ definition "entropy_density b M N = log b \ enn2real \ RN_deriv M N" definition "KL_divergence b M N = integral\<^sup>L N (entropy_density b M N)" lemma measurable_entropy_density[measurable]: "entropy_density b M N \ borel_measurable M" unfolding entropy_density_def by auto lemma (in sigma_finite_measure) KL_density: fixes f :: "'a \ real" assumes "1 < b" assumes f[measurable]: "f \ borel_measurable M" and nn: "AE x in M. 0 \ f x" shows "KL_divergence b M (density M f) = (\x. f x * log b (f x) \M)" unfolding KL_divergence_def proof (subst integral_real_density) show [measurable]: "entropy_density b M (density M (\x. ennreal (f x))) \ borel_measurable M" using f by (auto simp: comp_def entropy_density_def) have "density M (RN_deriv M (density M f)) = density M f" using f nn by (intro density_RN_deriv_density) auto then have eq: "AE x in M. RN_deriv M (density M f) x = f x" using f nn by (intro density_unique) auto show "(\x. f x * entropy_density b M (density M (\x. ennreal (f x))) x \M) = (\x. f x * log b (f x) \M)" apply (intro integral_cong_AE) apply measurable using eq nn apply eventually_elim apply (auto simp: entropy_density_def) done qed fact+ lemma (in sigma_finite_measure) KL_density_density: fixes f g :: "'a \ real" assumes "1 < b" assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" assumes ac: "AE x in M. f x = 0 \ g x = 0" shows "KL_divergence b (density M f) (density M g) = (\x. g x * log b (g x / f x) \M)" proof - interpret Mf: sigma_finite_measure "density M f" using f by (subst sigma_finite_iff_density_finite) auto have "KL_divergence b (density M f) (density M g) = KL_divergence b (density M f) (density (density M f) (\x. g x / f x))" using f g ac by (subst density_density_divide) simp_all also have "\ = (\x. (g x / f x) * log b (g x / f x) \density M f)" using f g \1 < b\ by (intro Mf.KL_density) (auto simp: AE_density) also have "\ = (\x. g x * log b (g x / f x) \M)" using ac f g \1 < b\ by (subst integral_density) (auto intro!: integral_cong_AE) finally show ?thesis . qed lemma (in information_space) KL_gt_0: fixes D :: "'a \ real" assumes "prob_space (density M D)" assumes D: "D \ borel_measurable M" "AE x in M. 0 \ D x" assumes int: "integrable M (\x. D x * log b (D x))" assumes A: "density M D \ M" shows "0 < KL_divergence b M (density M D)" proof - interpret N: prob_space "density M D" by fact obtain A where "A \ sets M" "emeasure (density M D) A \ emeasure M A" using measure_eqI[of "density M D" M] \density M D \ M\ by auto let ?D_set = "{x\space M. D x \ 0}" have [simp, intro]: "?D_set \ sets M" using D by auto have D_neg: "(\\<^sup>+ x. ennreal (- D x) \M) = 0" using D by (subst nn_integral_0_iff_AE) (auto simp: ennreal_neg) have "(\\<^sup>+ x. ennreal (D x) \M) = emeasure (density M D) (space M)" using D by (simp add: emeasure_density cong: nn_integral_cong) then have D_pos: "(\\<^sup>+ x. ennreal (D x) \M) = 1" using N.emeasure_space_1 by simp have "integrable M D" using D D_pos D_neg unfolding real_integrable_def real_lebesgue_integral_def by simp_all then have "integral\<^sup>L M D = 1" using D D_pos D_neg by (simp add: real_lebesgue_integral_def) have "0 \ 1 - measure M ?D_set" using prob_le_1 by (auto simp: field_simps) also have "\ = (\ x. D x - indicator ?D_set x \M)" using \integrable M D\ \integral\<^sup>L M D = 1\ by (simp add: emeasure_eq_measure) also have "\ < (\ x. D x * (ln b * log b (D x)) \M)" proof (rule integral_less_AE) show "integrable M (\x. D x - indicator ?D_set x)" using \integrable M D\ by (auto simp: less_top[symmetric]) next from integrable_mult_left(1)[OF int, of "ln b"] show "integrable M (\x. D x * (ln b * log b (D x)))" by (simp add: ac_simps) next show "emeasure M {x\space M. D x \ 1 \ D x \ 0} \ 0" proof assume eq_0: "emeasure M {x\space M. D x \ 1 \ D x \ 0} = 0" then have disj: "AE x in M. D x = 1 \ D x = 0" using D(1) by (auto intro!: AE_I[OF subset_refl] sets.sets_Collect) have "emeasure M {x\space M. D x = 1} = (\\<^sup>+ x. indicator {x\space M. D x = 1} x \M)" using D(1) by auto also have "\ = (\\<^sup>+ x. ennreal (D x) \M)" using disj by (auto intro!: nn_integral_cong_AE simp: indicator_def one_ennreal_def) finally have "AE x in M. D x = 1" using D D_pos by (intro AE_I_eq_1) auto then have "(\\<^sup>+x. indicator A x\M) = (\\<^sup>+x. ennreal (D x) * indicator A x\M)" by (intro nn_integral_cong_AE) (auto simp: one_ennreal_def[symmetric]) also have "\ = density M D A" using \A \ sets M\ D by (simp add: emeasure_density) finally show False using \A \ sets M\ \emeasure (density M D) A \ emeasure M A\ by simp qed show "{x\space M. D x \ 1 \ D x \ 0} \ sets M" using D(1) by (auto intro: sets.sets_Collect_conj) show "AE t in M. t \ {x\space M. D x \ 1 \ D x \ 0} \ D t - indicator ?D_set t \ D t * (ln b * log b (D t))" using D(2) proof (eventually_elim, safe) fix t assume Dt: "t \ space M" "D t \ 1" "D t \ 0" "0 \ D t" and eq: "D t - indicator ?D_set t = D t * (ln b * log b (D t))" have "D t - 1 = D t - indicator ?D_set t" using Dt by simp also note eq also have "D t * (ln b * log b (D t)) = - D t * ln (1 / D t)" using b_gt_1 \D t \ 0\ \0 \ D t\ by (simp add: log_def ln_div less_le) finally have "ln (1 / D t) = 1 / D t - 1" using \D t \ 0\ by (auto simp: field_simps) from ln_eq_minus_one[OF _ this] \D t \ 0\ \0 \ D t\ \D t \ 1\ show False by auto qed show "AE t in M. D t - indicator ?D_set t \ D t * (ln b * log b (D t))" using D(2) AE_space proof eventually_elim fix t assume "t \ space M" "0 \ D t" show "D t - indicator ?D_set t \ D t * (ln b * log b (D t))" proof cases assume asm: "D t \ 0" then have "0 < D t" using \0 \ D t\ by auto then have "0 < 1 / D t" by auto have "D t - indicator ?D_set t \ - D t * (1 / D t - 1)" using asm \t \ space M\ by (simp add: field_simps) also have "- D t * (1 / D t - 1) \ - D t * ln (1 / D t)" using ln_le_minus_one \0 < 1 / D t\ by (intro mult_left_mono_neg) auto also have "\ = D t * (ln b * log b (D t))" using \0 < D t\ b_gt_1 by (simp_all add: log_def ln_div) finally show ?thesis by simp qed simp qed qed also have "\ = (\ x. ln b * (D x * log b (D x)) \M)" by (simp add: ac_simps) also have "\ = ln b * (\ x. D x * log b (D x) \M)" using int by simp finally show ?thesis using b_gt_1 D by (subst KL_density) (auto simp: zero_less_mult_iff) qed lemma (in sigma_finite_measure) KL_same_eq_0: "KL_divergence b M M = 0" proof - have "AE x in M. 1 = RN_deriv M M x" proof (rule RN_deriv_unique) show "density M (\x. 1) = M" apply (auto intro!: measure_eqI emeasure_density) apply (subst emeasure_density) apply auto done qed auto then have "AE x in M. log b (enn2real (RN_deriv M M x)) = 0" by (elim AE_mp) simp from integral_cong_AE[OF _ _ this] have "integral\<^sup>L M (entropy_density b M M) = 0" by (simp add: entropy_density_def comp_def) then show "KL_divergence b M M = 0" unfolding KL_divergence_def by auto qed lemma (in information_space) KL_eq_0_iff_eq: fixes D :: "'a \ real" assumes "prob_space (density M D)" assumes D: "D \ borel_measurable M" "AE x in M. 0 \ D x" assumes int: "integrable M (\x. D x * log b (D x))" shows "KL_divergence b M (density M D) = 0 \ density M D = M" using KL_same_eq_0[of b] KL_gt_0[OF assms] by (auto simp: less_le) lemma (in information_space) KL_eq_0_iff_eq_ac: fixes D :: "'a \ real" assumes "prob_space N" assumes ac: "absolutely_continuous M N" "sets N = sets M" assumes int: "integrable N (entropy_density b M N)" shows "KL_divergence b M N = 0 \ N = M" proof - interpret N: prob_space N by fact have "finite_measure N" by unfold_locales - from real_RN_deriv[OF this ac] guess D . note D = this + from real_RN_deriv[OF this ac] obtain D + where D: + "random_variable borel D" + "AE x in M. RN_deriv M N x = ennreal (D x)" + "AE x in N. 0 < D x" + "\x. 0 \ D x" + by this auto have "N = density M (RN_deriv M N)" using ac by (rule density_RN_deriv[symmetric]) also have "\ = density M D" using D by (auto intro!: density_cong) finally have N: "N = density M D" . from absolutely_continuous_AE[OF ac(2,1) D(2)] D b_gt_1 ac measurable_entropy_density have "integrable N (\x. log b (D x))" by (intro integrable_cong_AE[THEN iffD2, OF _ _ _ int]) (auto simp: N entropy_density_def) with D b_gt_1 have "integrable M (\x. D x * log b (D x))" by (subst integrable_real_density[symmetric]) (auto simp: N[symmetric] comp_def) with \prob_space N\ D show ?thesis unfolding N by (intro KL_eq_0_iff_eq) auto qed lemma (in information_space) KL_nonneg: assumes "prob_space (density M D)" assumes D: "D \ borel_measurable M" "AE x in M. 0 \ D x" assumes int: "integrable M (\x. D x * log b (D x))" shows "0 \ KL_divergence b M (density M D)" using KL_gt_0[OF assms] by (cases "density M D = M") (auto simp: KL_same_eq_0) lemma (in sigma_finite_measure) KL_density_density_nonneg: fixes f g :: "'a \ real" assumes "1 < b" assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" "prob_space (density M f)" assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" "prob_space (density M g)" assumes ac: "AE x in M. f x = 0 \ g x = 0" assumes int: "integrable M (\x. g x * log b (g x / f x))" shows "0 \ KL_divergence b (density M f) (density M g)" proof - interpret Mf: prob_space "density M f" by fact interpret Mf: information_space "density M f" b by standard fact have eq: "density (density M f) (\x. g x / f x) = density M g" (is "?DD = _") using f g ac by (subst density_density_divide) simp_all have "0 \ KL_divergence b (density M f) (density (density M f) (\x. g x / f x))" proof (rule Mf.KL_nonneg) show "prob_space ?DD" unfolding eq by fact from f g show "(\x. g x / f x) \ borel_measurable (density M f)" by auto show "AE x in density M f. 0 \ g x / f x" using f g by (auto simp: AE_density) show "integrable (density M f) (\x. g x / f x * log b (g x / f x))" using \1 < b\ f g ac by (subst integrable_density) (auto intro!: integrable_cong_AE[THEN iffD2, OF _ _ _ int] measurable_If) qed also have "\ = KL_divergence b (density M f) (density M g)" using f g ac by (subst density_density_divide) simp_all finally show ?thesis . qed subsection \Finite Entropy\ definition (in information_space) finite_entropy :: "'b measure \ ('a \ 'b) \ ('b \ real) \ bool" where "finite_entropy S X f \ distributed M S X f \ integrable S (\x. f x * log b (f x)) \ (\x\space S. 0 \ f x)" lemma (in information_space) finite_entropy_simple_function: assumes X: "simple_function M X" shows "finite_entropy (count_space (X`space M)) X (\a. measure M {x \ space M. X x = a})" unfolding finite_entropy_def proof safe have [simp]: "finite (X ` space M)" using X by (auto simp: simple_function_def) then show "integrable (count_space (X ` space M)) (\x. prob {xa \ space M. X xa = x} * log b (prob {xa \ space M. X xa = x}))" by (rule integrable_count_space) have d: "distributed M (count_space (X ` space M)) X (\x. ennreal (if x \ X`space M then prob {xa \ space M. X xa = x} else 0))" by (rule distributed_simple_function_superset[OF X]) (auto intro!: arg_cong[where f=prob]) show "distributed M (count_space (X ` space M)) X (\x. ennreal (prob {xa \ space M. X xa = x}))" by (rule distributed_cong_density[THEN iffD1, OF _ _ _ d]) auto qed (rule measure_nonneg) lemma ac_fst: assumes "sigma_finite_measure T" shows "absolutely_continuous S (distr (S \\<^sub>M T) S fst)" proof - interpret sigma_finite_measure T by fact { fix A assume A: "A \ sets S" "emeasure S A = 0" then have "fst -` A \ space (S \\<^sub>M T) = A \ space T" by (auto simp: space_pair_measure dest!: sets.sets_into_space) with A have "emeasure (S \\<^sub>M T) (fst -` A \ space (S \\<^sub>M T)) = 0" by (simp add: emeasure_pair_measure_Times) } then show ?thesis unfolding absolutely_continuous_def apply (auto simp: null_sets_distr_iff) apply (auto simp: null_sets_def intro!: measurable_sets) done qed lemma ac_snd: assumes "sigma_finite_measure T" shows "absolutely_continuous T (distr (S \\<^sub>M T) T snd)" proof - interpret sigma_finite_measure T by fact { fix A assume A: "A \ sets T" "emeasure T A = 0" then have "snd -` A \ space (S \\<^sub>M T) = space S \ A" by (auto simp: space_pair_measure dest!: sets.sets_into_space) with A have "emeasure (S \\<^sub>M T) (snd -` A \ space (S \\<^sub>M T)) = 0" by (simp add: emeasure_pair_measure_Times) } then show ?thesis unfolding absolutely_continuous_def apply (auto simp: null_sets_distr_iff) apply (auto simp: null_sets_def intro!: measurable_sets) done qed lemma (in information_space) finite_entropy_integrable: "finite_entropy S X Px \ integrable S (\x. Px x * log b (Px x))" unfolding finite_entropy_def by auto lemma (in information_space) finite_entropy_distributed: "finite_entropy S X Px \ distributed M S X Px" unfolding finite_entropy_def by auto lemma (in information_space) finite_entropy_nn: "finite_entropy S X Px \ x \ space S \ 0 \ Px x" by (auto simp: finite_entropy_def) lemma (in information_space) finite_entropy_measurable: "finite_entropy S X Px \ Px \ S \\<^sub>M borel" using distributed_real_measurable[of S Px M X] finite_entropy_nn[of S X Px] finite_entropy_distributed[of S X Px] by auto lemma (in information_space) subdensity_finite_entropy: fixes g :: "'b \ real" and f :: "'c \ real" assumes T: "T \ measurable P Q" assumes f: "finite_entropy P X f" assumes g: "finite_entropy Q Y g" assumes Y: "Y = T \ X" shows "AE x in P. g (T x) = 0 \ f x = 0" using subdensity[OF T, of M X "\x. ennreal (f x)" Y "\x. ennreal (g x)"] finite_entropy_distributed[OF f] finite_entropy_distributed[OF g] finite_entropy_nn[OF f] finite_entropy_nn[OF g] assms by auto lemma (in information_space) finite_entropy_integrable_transform: "finite_entropy S X Px \ distributed M T Y Py \ (\x. x \ space T \ 0 \ Py x) \ X = (\x. f (Y x)) \ f \ measurable T S \ integrable T (\x. Py x * log b (Px (f x)))" using distributed_transform_integrable[of M T Y Py S X Px f "\x. log b (Px x)"] using distributed_real_measurable[of S Px M X] by (auto simp: finite_entropy_def) subsection \Mutual Information\ definition (in prob_space) "mutual_information b S T X Y = KL_divergence b (distr M S X \\<^sub>M distr M T Y) (distr M (S \\<^sub>M T) (\x. (X x, Y x)))" lemma (in information_space) mutual_information_indep_vars: fixes S T X Y defines "P \ distr M S X \\<^sub>M distr M T Y" defines "Q \ distr M (S \\<^sub>M T) (\x. (X x, Y x))" shows "indep_var S X T Y \ (random_variable S X \ random_variable T Y \ absolutely_continuous P Q \ integrable Q (entropy_density b P Q) \ mutual_information b S T X Y = 0)" unfolding indep_var_distribution_eq proof safe assume rv[measurable]: "random_variable S X" "random_variable T Y" interpret X: prob_space "distr M S X" by (rule prob_space_distr) fact interpret Y: prob_space "distr M T Y" by (rule prob_space_distr) fact interpret XY: pair_prob_space "distr M S X" "distr M T Y" by standard interpret P: information_space P b unfolding P_def by standard (rule b_gt_1) interpret Q: prob_space Q unfolding Q_def by (rule prob_space_distr) simp { assume "distr M S X \\<^sub>M distr M T Y = distr M (S \\<^sub>M T) (\x. (X x, Y x))" then have [simp]: "Q = P" unfolding Q_def P_def by simp show ac: "absolutely_continuous P Q" by (simp add: absolutely_continuous_def) then have ed: "entropy_density b P Q \ borel_measurable P" by simp have "AE x in P. 1 = RN_deriv P Q x" proof (rule P.RN_deriv_unique) show "density P (\x. 1) = Q" unfolding \Q = P\ by (intro measure_eqI) (auto simp: emeasure_density) qed auto then have ae_0: "AE x in P. entropy_density b P Q x = 0" by eventually_elim (auto simp: entropy_density_def) then have "integrable P (entropy_density b P Q) \ integrable Q (\x. 0::real)" using ed unfolding \Q = P\ by (intro integrable_cong_AE) auto then show "integrable Q (entropy_density b P Q)" by simp from ae_0 have "mutual_information b S T X Y = (\x. 0 \P)" unfolding mutual_information_def KL_divergence_def P_def[symmetric] Q_def[symmetric] \Q = P\ by (intro integral_cong_AE) auto then show "mutual_information b S T X Y = 0" by simp } { assume ac: "absolutely_continuous P Q" assume int: "integrable Q (entropy_density b P Q)" assume I_eq_0: "mutual_information b S T X Y = 0" have eq: "Q = P" proof (rule P.KL_eq_0_iff_eq_ac[THEN iffD1]) show "prob_space Q" by unfold_locales show "absolutely_continuous P Q" by fact show "integrable Q (entropy_density b P Q)" by fact show "sets Q = sets P" by (simp add: P_def Q_def sets_pair_measure) show "KL_divergence b P Q = 0" using I_eq_0 unfolding mutual_information_def by (simp add: P_def Q_def) qed then show "distr M S X \\<^sub>M distr M T Y = distr M (S \\<^sub>M T) (\x. (X x, Y x))" unfolding P_def Q_def .. } qed abbreviation (in information_space) mutual_information_Pow ("\'(_ ; _')") where "\(X ; Y) \ mutual_information b (count_space (X`space M)) (count_space (Y`space M)) X Y" lemma (in information_space) fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Fx: "finite_entropy S X Px" and Fy: "finite_entropy T Y Py" assumes Fxy: "finite_entropy (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" defines "f \ \x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" shows mutual_information_distr': "mutual_information b S T X Y = integral\<^sup>L (S \\<^sub>M T) f" (is "?M = ?R") and mutual_information_nonneg': "0 \ mutual_information b S T X Y" proof - have Px: "distributed M S X Px" and Px_nn: "\x. x \ space S \ 0 \ Px x" using Fx by (auto simp: finite_entropy_def) have Py: "distributed M T Y Py" and Py_nn: "\x. x \ space T \ 0 \ Py x" using Fy by (auto simp: finite_entropy_def) have Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" and Pxy_nn: "\x. x \ space (S \\<^sub>M T) \ 0 \ Pxy x" "\x y. x \ space S \ y \ space T \ 0 \ Pxy (x, y)" using Fxy by (auto simp: finite_entropy_def space_pair_measure) have [measurable]: "Px \ S \\<^sub>M borel" using Px Px_nn by (intro distributed_real_measurable) have [measurable]: "Py \ T \\<^sub>M borel" using Py Py_nn by (intro distributed_real_measurable) have measurable_Pxy[measurable]: "Pxy \ (S \\<^sub>M T) \\<^sub>M borel" using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) have X[measurable]: "random_variable S X" using Px by auto have Y[measurable]: "random_variable T Y" using Py by auto interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. interpret X: prob_space "distr M S X" using X by (rule prob_space_distr) interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr) interpret XY: pair_prob_space "distr M S X" "distr M T Y" .. let ?P = "S \\<^sub>M T" let ?D = "distr M ?P (\x. (X x, Y x))" { fix A assume "A \ sets S" with X[THEN measurable_space] Y[THEN measurable_space] have "emeasure (distr M S X) A = emeasure ?D (A \ space T)" by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } note marginal_eq1 = this { fix A assume "A \ sets T" with X[THEN measurable_space] Y[THEN measurable_space] have "emeasure (distr M T Y) A = emeasure ?D (space S \ A)" by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } note marginal_eq2 = this have distr_eq: "distr M S X \\<^sub>M distr M T Y = density ?P (\x. ennreal (Px (fst x) * Py (snd x)))" unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] proof (subst pair_measure_density) show "(\x. ennreal (Px x)) \ borel_measurable S" "(\y. ennreal (Py y)) \ borel_measurable T" using Px Py by (auto simp: distributed_def) show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] .. show "density (S \\<^sub>M T) (\(x, y). ennreal (Px x) * ennreal (Py y)) = density (S \\<^sub>M T) (\x. ennreal (Px (fst x) * Py (snd x)))" using Px_nn Py_nn by (auto intro!: density_cong simp: distributed_def ennreal_mult space_pair_measure) qed fact have M: "?M = KL_divergence b (density ?P (\x. ennreal (Px (fst x) * Py (snd x)))) (density ?P (\x. ennreal (Pxy x)))" unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] .. from Px Py have f: "(\x. Px (fst x) * Py (snd x)) \ borel_measurable ?P" by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'') have PxPy_nonneg: "AE x in ?P. 0 \ Px (fst x) * Py (snd x)" using Px_nn Py_nn by (auto simp: space_pair_measure) have A: "(AE x in ?P. Px (fst x) = 0 \ Pxy x = 0)" by (rule subdensity_real[OF measurable_fst Pxy Px]) (insert Px_nn Pxy_nn, auto simp: space_pair_measure) moreover have B: "(AE x in ?P. Py (snd x) = 0 \ Pxy x = 0)" by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure) ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \ Pxy x = 0" by eventually_elim auto show "?M = ?R" unfolding M f_def using Pxy_nn Px_nn Py_nn by (intro ST.KL_density_density b_gt_1 f PxPy_nonneg ac) (auto simp: space_pair_measure) have X: "X = fst \ (\x. (X x, Y x))" and Y: "Y = snd \ (\x. (X x, Y x))" by auto have "integrable (S \\<^sub>M T) (\x. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)))" using finite_entropy_integrable[OF Fxy] using finite_entropy_integrable_transform[OF Fx Pxy, of fst] using finite_entropy_integrable_transform[OF Fy Pxy, of snd] by (simp add: Pxy_nn) moreover have "f \ borel_measurable (S \\<^sub>M T)" unfolding f_def using Px Py Pxy by (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'' intro!: borel_measurable_times borel_measurable_log borel_measurable_divide) ultimately have int: "integrable (S \\<^sub>M T) f" apply (rule integrable_cong_AE_imp) using A B AE_space by eventually_elim (auto simp: f_def log_divide_eq log_mult_eq field_simps space_pair_measure Px_nn Py_nn Pxy_nn less_le) show "0 \ ?M" unfolding M proof (intro ST.KL_density_density_nonneg) show "prob_space (density (S \\<^sub>M T) (\x. ennreal (Pxy x))) " unfolding distributed_distr_eq_density[OF Pxy, symmetric] using distributed_measurable[OF Pxy] by (rule prob_space_distr) show "prob_space (density (S \\<^sub>M T) (\x. ennreal (Px (fst x) * Py (snd x))))" unfolding distr_eq[symmetric] by unfold_locales show "integrable (S \\<^sub>M T) (\x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))))" using int unfolding f_def . qed (insert ac, auto simp: b_gt_1 Px_nn Py_nn Pxy_nn space_pair_measure) qed lemma (in information_space) fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" assumes "sigma_finite_measure S" "sigma_finite_measure T" assumes Px: "distributed M S X Px" and Px_nn: "\x. x \ space S \ 0 \ Px x" and Py: "distributed M T Y Py" and Py_nn: "\y. y \ space T \ 0 \ Py y" and Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" and Pxy_nn: "\x y. x \ space S \ y \ space T \ 0 \ Pxy (x, y)" defines "f \ \x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" shows mutual_information_distr: "mutual_information b S T X Y = integral\<^sup>L (S \\<^sub>M T) f" (is "?M = ?R") and mutual_information_nonneg: "integrable (S \\<^sub>M T) f \ 0 \ mutual_information b S T X Y" proof - have X[measurable]: "random_variable S X" using Px by (auto simp: distributed_def) have Y[measurable]: "random_variable T Y" using Py by (auto simp: distributed_def) have [measurable]: "Px \ S \\<^sub>M borel" using Px Px_nn by (intro distributed_real_measurable) have [measurable]: "Py \ T \\<^sub>M borel" using Py Py_nn by (intro distributed_real_measurable) have measurable_Pxy[measurable]: "Pxy \ (S \\<^sub>M T) \\<^sub>M borel" using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. interpret X: prob_space "distr M S X" using X by (rule prob_space_distr) interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr) interpret XY: pair_prob_space "distr M S X" "distr M T Y" .. let ?P = "S \\<^sub>M T" let ?D = "distr M ?P (\x. (X x, Y x))" { fix A assume "A \ sets S" with X[THEN measurable_space] Y[THEN measurable_space] have "emeasure (distr M S X) A = emeasure ?D (A \ space T)" by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } note marginal_eq1 = this { fix A assume "A \ sets T" with X[THEN measurable_space] Y[THEN measurable_space] have "emeasure (distr M T Y) A = emeasure ?D (space S \ A)" by (auto simp: emeasure_distr intro!: arg_cong[where f="emeasure M"]) } note marginal_eq2 = this have distr_eq: "distr M S X \\<^sub>M distr M T Y = density ?P (\x. ennreal (Px (fst x) * Py (snd x)))" unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] proof (subst pair_measure_density) show "(\x. ennreal (Px x)) \ borel_measurable S" "(\y. ennreal (Py y)) \ borel_measurable T" using Px Py by (auto simp: distributed_def) show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] .. show "density (S \\<^sub>M T) (\(x, y). ennreal (Px x) * ennreal (Py y)) = density (S \\<^sub>M T) (\x. ennreal (Px (fst x) * Py (snd x)))" using Px_nn Py_nn by (auto intro!: density_cong simp: distributed_def ennreal_mult space_pair_measure) qed fact have M: "?M = KL_divergence b (density ?P (\x. ennreal (Px (fst x) * Py (snd x)))) (density ?P (\x. ennreal (Pxy x)))" unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] .. from Px Py have f: "(\x. Px (fst x) * Py (snd x)) \ borel_measurable ?P" by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'') have PxPy_nonneg: "AE x in ?P. 0 \ Px (fst x) * Py (snd x)" using Px_nn Py_nn by (auto simp: space_pair_measure) have "(AE x in ?P. Px (fst x) = 0 \ Pxy x = 0)" by (rule subdensity_real[OF measurable_fst Pxy Px]) (insert Px_nn Pxy_nn, auto simp: space_pair_measure) moreover have "(AE x in ?P. Py (snd x) = 0 \ Pxy x = 0)" by (rule subdensity_real[OF measurable_snd Pxy Py]) (insert Py_nn Pxy_nn, auto simp: space_pair_measure) ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \ Pxy x = 0" by eventually_elim auto show "?M = ?R" unfolding M f_def using b_gt_1 f PxPy_nonneg ac Pxy_nn by (intro ST.KL_density_density) (auto simp: space_pair_measure) assume int: "integrable (S \\<^sub>M T) f" show "0 \ ?M" unfolding M proof (intro ST.KL_density_density_nonneg) show "prob_space (density (S \\<^sub>M T) (\x. ennreal (Pxy x))) " unfolding distributed_distr_eq_density[OF Pxy, symmetric] using distributed_measurable[OF Pxy] by (rule prob_space_distr) show "prob_space (density (S \\<^sub>M T) (\x. ennreal (Px (fst x) * Py (snd x))))" unfolding distr_eq[symmetric] by unfold_locales show "integrable (S \\<^sub>M T) (\x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))))" using int unfolding f_def . qed (insert ac, auto simp: b_gt_1 Px_nn Py_nn Pxy_nn space_pair_measure) qed lemma (in information_space) fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" assumes "sigma_finite_measure S" "sigma_finite_measure T" assumes Px[measurable]: "distributed M S X Px" and Px_nn: "\x. x \ space S \ 0 \ Px x" and Py[measurable]: "distributed M T Y Py" and Py_nn: "\x. x \ space T \ 0 \ Py x" and Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" and Pxy_nn: "\x. x \ space (S \\<^sub>M T) \ 0 \ Pxy x" assumes ae: "AE x in S. AE y in T. Pxy (x, y) = Px x * Py y" shows mutual_information_eq_0: "mutual_information b S T X Y = 0" proof - interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. note distributed_real_measurable[OF Px_nn Px, measurable] distributed_real_measurable[OF Py_nn Py, measurable] distributed_real_measurable[OF Pxy_nn Pxy, measurable] have "AE x in S \\<^sub>M T. Px (fst x) = 0 \ Pxy x = 0" by (rule subdensity_real[OF measurable_fst Pxy Px]) (auto simp: Px_nn Pxy_nn space_pair_measure) moreover have "AE x in S \\<^sub>M T. Py (snd x) = 0 \ Pxy x = 0" by (rule subdensity_real[OF measurable_snd Pxy Py]) (auto simp: Py_nn Pxy_nn space_pair_measure) moreover have "AE x in S \\<^sub>M T. Pxy x = Px (fst x) * Py (snd x)" by (intro ST.AE_pair_measure) (auto simp: ae intro!: measurable_snd'' measurable_fst'') ultimately have "AE x in S \\<^sub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0" by eventually_elim simp then have "(\x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \(S \\<^sub>M T)) = (\x. 0 \(S \\<^sub>M T))" by (intro integral_cong_AE) auto then show ?thesis by (subst mutual_information_distr[OF assms(1-8)]) (auto simp add: space_pair_measure) qed lemma (in information_space) mutual_information_simple_distributed: assumes X: "simple_distributed M X Px" and Y: "simple_distributed M Y Py" assumes XY: "simple_distributed M (\x. (X x, Y x)) Pxy" shows "\(X ; Y) = (\(x, y)\(\x. (X x, Y x))`space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))" proof (subst mutual_information_distr[OF _ _ simple_distributed[OF X] _ simple_distributed[OF Y] _ simple_distributed_joint[OF XY]]) note fin = simple_distributed_joint_finite[OF XY, simp] show "sigma_finite_measure (count_space (X ` space M))" by (simp add: sigma_finite_measure_count_space_finite) show "sigma_finite_measure (count_space (Y ` space M))" by (simp add: sigma_finite_measure_count_space_finite) let ?Pxy = "\x. (if x \ (\x. (X x, Y x)) ` space M then Pxy x else 0)" let ?f = "\x. ?Pxy x * log b (?Pxy x / (Px (fst x) * Py (snd x)))" have "\x. ?f x = (if x \ (\x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) else 0)" by auto with fin show "(\ x. ?f x \(count_space (X ` space M) \\<^sub>M count_space (Y ` space M))) = (\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))" by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite sum.If_cases split_beta' intro!: sum.cong) qed (insert X Y XY, auto simp: simple_distributed_def) lemma (in information_space) fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" assumes Px: "simple_distributed M X Px" and Py: "simple_distributed M Y Py" assumes Pxy: "simple_distributed M (\x. (X x, Y x)) Pxy" assumes ae: "\x\space M. Pxy (X x, Y x) = Px (X x) * Py (Y x)" shows mutual_information_eq_0_simple: "\(X ; Y) = 0" proof (subst mutual_information_simple_distributed[OF Px Py Pxy]) have "(\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = (\(x, y)\(\x. (X x, Y x)) ` space M. 0)" by (intro sum.cong) (auto simp: ae) then show "(\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = 0" by simp qed subsection \Entropy\ definition (in prob_space) entropy :: "real \ 'b measure \ ('a \ 'b) \ real" where "entropy b S X = - KL_divergence b S (distr M S X)" abbreviation (in information_space) entropy_Pow ("\'(_')") where "\(X) \ entropy b (count_space (X`space M)) X" lemma (in prob_space) distributed_RN_deriv: assumes X: "distributed M S X Px" shows "AE x in S. RN_deriv S (density S Px) x = Px x" proof - note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] interpret X: prob_space "distr M S X" using D(1) by (rule prob_space_distr) have sf: "sigma_finite_measure (distr M S X)" by standard show ?thesis using D apply (subst eq_commute) apply (intro RN_deriv_unique_sigma_finite) apply (auto simp: distributed_distr_eq_density[symmetric, OF X] sf) done qed lemma (in information_space) fixes X :: "'a \ 'b" assumes X[measurable]: "distributed M MX X f" and nn: "\x. x \ space MX \ 0 \ f x" shows entropy_distr: "entropy b MX X = - (\x. f x * log b (f x) \MX)" (is ?eq) proof - note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] note ae = distributed_RN_deriv[OF X] note distributed_real_measurable[OF nn X, measurable] have ae_eq: "AE x in distr M MX X. log b (enn2real (RN_deriv MX (distr M MX X) x)) = log b (f x)" unfolding distributed_distr_eq_density[OF X] apply (subst AE_density) using D apply simp using ae apply eventually_elim apply auto done have int_eq: "(\ x. f x * log b (f x) \MX) = (\ x. log b (f x) \distr M MX X)" unfolding distributed_distr_eq_density[OF X] using D by (subst integral_density) (auto simp: nn) show ?eq unfolding entropy_def KL_divergence_def entropy_density_def comp_def int_eq neg_equal_iff_equal using ae_eq by (intro integral_cong_AE) (auto simp: nn) qed lemma (in information_space) entropy_le: fixes Px :: "'b \ real" and MX :: "'b measure" assumes X[measurable]: "distributed M MX X Px" and Px_nn[simp]: "\x. x \ space MX \ 0 \ Px x" and fin: "emeasure MX {x \ space MX. Px x \ 0} \ top" and int: "integrable MX (\x. - Px x * log b (Px x))" shows "entropy b MX X \ log b (measure MX {x \ space MX. Px x \ 0})" proof - note Px = distributed_borel_measurable[OF X] interpret X: prob_space "distr M MX X" using distributed_measurable[OF X] by (rule prob_space_distr) have " - log b (measure MX {x \ space MX. Px x \ 0}) = - log b (\ x. indicator {x \ space MX. Px x \ 0} x \MX)" using Px Px_nn fin by (auto simp: measure_def) also have "- log b (\ x. indicator {x \ space MX. Px x \ 0} x \MX) = - log b (\ x. 1 / Px x \distr M MX X)" proof - have "integral\<^sup>L MX (indicator {x \ space MX. Px x \ 0}) = LINT x|MX. Px x *\<^sub>R (1 / Px x)" by (rule Bochner_Integration.integral_cong) auto also have "... = LINT x|density MX (\x. ennreal (Px x)). 1 / Px x" by (rule integral_density [symmetric]) (use Px Px_nn in auto) finally show ?thesis unfolding distributed_distr_eq_density[OF X] by simp qed also have "\ \ (\ x. - log b (1 / Px x) \distr M MX X)" proof (rule X.jensens_inequality[of "\x. 1 / Px x" "{0<..}" 0 1 "\x. - log b x"]) show "AE x in distr M MX X. 1 / Px x \ {0<..}" unfolding distributed_distr_eq_density[OF X] using Px by (auto simp: AE_density) have [simp]: "\x. x \ space MX \ ennreal (if Px x = 0 then 0 else 1) = indicator {x \ space MX. Px x \ 0} x" by (auto simp: one_ennreal_def) have "(\\<^sup>+ x. ennreal (- (if Px x = 0 then 0 else 1)) \MX) = (\\<^sup>+ x. 0 \MX)" by (intro nn_integral_cong) (auto simp: ennreal_neg) then show "integrable (distr M MX X) (\x. 1 / Px x)" unfolding distributed_distr_eq_density[OF X] using Px by (auto simp: nn_integral_density real_integrable_def fin ennreal_neg ennreal_mult[symmetric] cong: nn_integral_cong) have "integrable MX (\x. Px x * log b (1 / Px x)) = integrable MX (\x. - Px x * log b (Px x))" using Px by (intro integrable_cong_AE) (auto simp: log_divide_eq less_le) then show "integrable (distr M MX X) (\x. - log b (1 / Px x))" unfolding distributed_distr_eq_density[OF X] using Px int by (subst integrable_real_density) auto qed (auto simp: minus_log_convex[OF b_gt_1]) also have "\ = (\ x. log b (Px x) \distr M MX X)" unfolding distributed_distr_eq_density[OF X] using Px by (intro integral_cong_AE) (auto simp: AE_density log_divide_eq) also have "\ = - entropy b MX X" unfolding distributed_distr_eq_density[OF X] using Px by (subst entropy_distr[OF X]) (auto simp: integral_density) finally show ?thesis by simp qed lemma (in information_space) entropy_le_space: fixes Px :: "'b \ real" and MX :: "'b measure" assumes X: "distributed M MX X Px" and Px_nn[simp]: "\x. x \ space MX \ 0 \ Px x" and fin: "finite_measure MX" and int: "integrable MX (\x. - Px x * log b (Px x))" shows "entropy b MX X \ log b (measure MX (space MX))" proof - note Px = distributed_borel_measurable[OF X] interpret finite_measure MX by fact have "entropy b MX X \ log b (measure MX {x \ space MX. Px x \ 0})" using int X by (intro entropy_le) auto also have "\ \ log b (measure MX (space MX))" using Px distributed_imp_emeasure_nonzero[OF X] by (intro log_le) (auto intro!: finite_measure_mono b_gt_1 less_le[THEN iffD2] simp: emeasure_eq_measure cong: conj_cong) finally show ?thesis . qed lemma (in information_space) entropy_uniform: assumes X: "distributed M MX X (\x. indicator A x / measure MX A)" (is "distributed _ _ _ ?f") shows "entropy b MX X = log b (measure MX A)" proof (subst entropy_distr[OF X]) have [simp]: "emeasure MX A \ \" using uniform_distributed_params[OF X] by (auto simp add: measure_def) have eq: "(\ x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \MX) = (\ x. (- log b (measure MX A) / measure MX A) * indicator A x \MX)" using uniform_distributed_params[OF X] by (intro Bochner_Integration.integral_cong) (auto split: split_indicator simp: log_divide_eq zero_less_measure_iff) show "- (\ x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \MX) = log b (measure MX A)" unfolding eq using uniform_distributed_params[OF X] by (subst Bochner_Integration.integral_mult_right) (auto simp: measure_def less_top[symmetric] intro!: integrable_real_indicator) qed simp lemma (in information_space) entropy_simple_distributed: "simple_distributed M X f \ \(X) = - (\x\X`space M. f x * log b (f x))" by (subst entropy_distr[OF simple_distributed]) (auto simp add: lebesgue_integral_count_space_finite) lemma (in information_space) entropy_le_card_not_0: assumes X: "simple_distributed M X f" shows "\(X) \ log b (card (X ` space M \ {x. f x \ 0}))" proof - let ?X = "count_space (X`space M)" have "\(X) \ log b (measure ?X {x \ space ?X. f x \ 0})" by (rule entropy_le[OF simple_distributed[OF X]]) (insert X, auto simp: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space) also have "measure ?X {x \ space ?X. f x \ 0} = card (X ` space M \ {x. f x \ 0})" by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def Int_def) finally show ?thesis . qed lemma (in information_space) entropy_le_card: assumes X: "simple_distributed M X f" shows "\(X) \ log b (real (card (X ` space M)))" proof - let ?X = "count_space (X`space M)" have "\(X) \ log b (measure ?X (space ?X))" by (rule entropy_le_space[OF simple_distributed[OF X]]) (insert X, auto simp: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space finite_measure_count_space) also have "measure ?X (space ?X) = card (X ` space M)" by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def) finally show ?thesis . qed subsection \Conditional Mutual Information\ definition (in prob_space) "conditional_mutual_information b MX MY MZ X Y Z \ mutual_information b MX (MY \\<^sub>M MZ) X (\x. (Y x, Z x)) - mutual_information b MX MZ X Z" abbreviation (in information_space) conditional_mutual_information_Pow ("\'( _ ; _ | _ ')") where "\(X ; Y | Z) \ conditional_mutual_information b (count_space (X ` space M)) (count_space (Y ` space M)) (count_space (Z ` space M)) X Y Z" lemma (in information_space) assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P" assumes Px[measurable]: "distributed M S X Px" and Px_nn[simp]: "\x. x \ space S \ 0 \ Px x" assumes Pz[measurable]: "distributed M P Z Pz" and Pz_nn[simp]: "\z. z \ space P \ 0 \ Pz z" assumes Pyz[measurable]: "distributed M (T \\<^sub>M P) (\x. (Y x, Z x)) Pyz" and Pyz_nn[simp]: "\y z. y \ space T \ z \ space P \ 0 \ Pyz (y, z)" assumes Pxz[measurable]: "distributed M (S \\<^sub>M P) (\x. (X x, Z x)) Pxz" and Pxz_nn[simp]: "\x z. x \ space S \ z \ space P \ 0 \ Pxz (x, z)" assumes Pxyz[measurable]: "distributed M (S \\<^sub>M T \\<^sub>M P) (\x. (X x, Y x, Z x)) Pxyz" and Pxyz_nn[simp]: "\x y z. x \ space S \ y \ space T \ z \ space P \ 0 \ Pxyz (x, y, z)" assumes I1: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))" assumes I2: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))" shows conditional_mutual_information_generic_eq: "conditional_mutual_information b S T P X Y Z = (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \(S \\<^sub>M T \\<^sub>M P))" (is "?eq") and conditional_mutual_information_generic_nonneg: "0 \ conditional_mutual_information b S T P X Y Z" (is "?nonneg") proof - have [measurable]: "Px \ S \\<^sub>M borel" using Px Px_nn by (intro distributed_real_measurable) have [measurable]: "Pz \ P \\<^sub>M borel" using Pz Pz_nn by (intro distributed_real_measurable) have measurable_Pyz[measurable]: "Pyz \ (T \\<^sub>M P) \\<^sub>M borel" using Pyz Pyz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) have measurable_Pxz[measurable]: "Pxz \ (S \\<^sub>M P) \\<^sub>M borel" using Pxz Pxz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) have measurable_Pxyz[measurable]: "Pxyz \ (S \\<^sub>M T \\<^sub>M P) \\<^sub>M borel" using Pxyz Pxyz_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret P: sigma_finite_measure P by fact interpret TP: pair_sigma_finite T P .. interpret SP: pair_sigma_finite S P .. interpret ST: pair_sigma_finite S T .. interpret SPT: pair_sigma_finite "S \\<^sub>M P" T .. interpret STP: pair_sigma_finite S "T \\<^sub>M P" .. interpret TPS: pair_sigma_finite "T \\<^sub>M P" S .. have TP: "sigma_finite_measure (T \\<^sub>M P)" .. have SP: "sigma_finite_measure (S \\<^sub>M P)" .. have YZ: "random_variable (T \\<^sub>M P) (\x. (Y x, Z x))" using Pyz by (simp add: distributed_measurable) from Pxz Pxyz have distr_eq: "distr M (S \\<^sub>M P) (\x. (X x, Z x)) = distr (distr M (S \\<^sub>M T \\<^sub>M P) (\x. (X x, Y x, Z x))) (S \\<^sub>M P) (\(x, y, z). (x, z))" by (simp add: comp_def distr_distr) have "mutual_information b S P X Z = (\x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \(S \\<^sub>M P))" by (rule mutual_information_distr[OF S P Px Px_nn Pz Pz_nn Pxz Pxz_nn]) also have "\ = (\(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \(S \\<^sub>M T \\<^sub>M P))" using b_gt_1 Pxz Px Pz by (subst distributed_transform_integral[OF Pxyz _ Pxz _, where T="\(x, y, z). (x, z)"]) (auto simp: split_beta' space_pair_measure) finally have mi_eq: "mutual_information b S P X Z = (\(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \(S \\<^sub>M T \\<^sub>M P))" . have ae1: "AE x in S \\<^sub>M T \\<^sub>M P. Px (fst x) = 0 \ Pxyz x = 0" by (intro subdensity_real[of fst, OF _ Pxyz Px]) (auto simp: space_pair_measure) moreover have ae2: "AE x in S \\<^sub>M T \\<^sub>M P. Pz (snd (snd x)) = 0 \ Pxyz x = 0" by (intro subdensity_real[of "\x. snd (snd x)", OF _ Pxyz Pz]) (auto simp: space_pair_measure) moreover have ae3: "AE x in S \\<^sub>M T \\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \ Pxyz x = 0" by (intro subdensity_real[of "\x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) (auto simp: space_pair_measure) moreover have ae4: "AE x in S \\<^sub>M T \\<^sub>M P. Pyz (snd x) = 0 \ Pxyz x = 0" by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) (auto simp: space_pair_measure) ultimately have ae: "AE x in S \\<^sub>M T \\<^sub>M P. Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) - Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) = Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) " using AE_space proof eventually_elim case (elim x) show ?case proof cases assume "Pxyz x \ 0" with elim have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))" "0 < Pyz (snd x)" "0 < Pxyz x" by (auto simp: space_pair_measure less_le) then show ?thesis using b_gt_1 by (simp add: log_simps less_imp_le field_simps) qed simp qed with I1 I2 show ?eq unfolding conditional_mutual_information_def apply (subst mi_eq) apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz _ Pxyz]) apply (auto simp: space_pair_measure) apply (subst Bochner_Integration.integral_diff[symmetric]) apply (auto intro!: integral_cong_AE simp: split_beta' simp del: Bochner_Integration.integral_diff) done let ?P = "density (S \\<^sub>M T \\<^sub>M P) Pxyz" interpret P: prob_space ?P unfolding distributed_distr_eq_density[OF Pxyz, symmetric] by (rule prob_space_distr) simp let ?Q = "density (T \\<^sub>M P) Pyz" interpret Q: prob_space ?Q unfolding distributed_distr_eq_density[OF Pyz, symmetric] by (rule prob_space_distr) simp let ?f = "\(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)" from subdensity_real[of snd, OF _ Pyz Pz _ AE_I2 AE_I2] have aeX1: "AE x in T \\<^sub>M P. Pz (snd x) = 0 \ Pyz x = 0" by (auto simp: comp_def space_pair_measure) have aeX2: "AE x in T \\<^sub>M P. 0 \ Pz (snd x)" using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def) have aeX3: "AE y in T \\<^sub>M P. (\\<^sup>+ x. ennreal (Pxz (x, snd y)) \S) = ennreal (Pz (snd y))" using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz] by (intro TP.AE_pair_measure) auto have "(\\<^sup>+ x. ?f x \?P) \ (\\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \(S \\<^sub>M T \\<^sub>M P))" by (subst nn_integral_density) (auto intro!: nn_integral_mono simp: space_pair_measure ennreal_mult[symmetric]) also have "\ = (\\<^sup>+(y, z). (\\<^sup>+ x. ennreal (Pxz (x, z)) * ennreal (Pyz (y, z) / Pz z) \S) \(T \\<^sub>M P))" by (subst STP.nn_integral_snd[symmetric]) (auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong) also have "\ = (\\<^sup>+x. ennreal (Pyz x) * 1 \T \\<^sub>M P)" apply (rule nn_integral_cong_AE) using aeX1 aeX2 aeX3 AE_space apply eventually_elim proof (case_tac x, simp add: space_pair_measure) fix a b assume "Pz b = 0 \ Pyz (a, b) = 0" "a \ space T \ b \ space P" "(\\<^sup>+ x. ennreal (Pxz (x, b)) \S) = ennreal (Pz b)" then show "(\\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \S) = ennreal (Pyz (a, b))" by (subst nn_integral_multc) (auto split: prod.split simp: ennreal_mult[symmetric]) qed also have "\ = 1" using Q.emeasure_space_1 distributed_distr_eq_density[OF Pyz] by (subst nn_integral_density[symmetric]) auto finally have le1: "(\\<^sup>+ x. ?f x \?P) \ 1" . also have "\ < \" by simp finally have fin: "(\\<^sup>+ x. ?f x \?P) \ \" by simp have pos: "(\\<^sup>+x. ?f x \?P) \ 0" apply (subst nn_integral_density) apply (simp_all add: split_beta') proof let ?g = "\x. ennreal (Pxyz x) * (Pxz (fst x, snd (snd x)) * Pyz (snd x) / (Pz (snd (snd x)) * Pxyz x))" assume "(\\<^sup>+x. ?g x \(S \\<^sub>M T \\<^sub>M P)) = 0" then have "AE x in S \\<^sub>M T \\<^sub>M P. ?g x = 0" by (intro nn_integral_0_iff_AE[THEN iffD1]) auto then have "AE x in S \\<^sub>M T \\<^sub>M P. Pxyz x = 0" using ae1 ae2 ae3 ae4 AE_space by eventually_elim (auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure) then have "(\\<^sup>+ x. ennreal (Pxyz x) \S \\<^sub>M T \\<^sub>M P) = 0" by (subst nn_integral_cong_AE[of _ "\x. 0"]) auto with P.emeasure_space_1 show False by (subst (asm) emeasure_density) (auto cong: nn_integral_cong) qed have neg: "(\\<^sup>+ x. - ?f x \?P) = 0" apply (rule nn_integral_0_iff_AE[THEN iffD2]) apply simp apply (subst AE_density) apply (auto simp: space_pair_measure ennreal_neg) done have I3: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ Bochner_Integration.integrable_diff[OF I1 I2]]) using ae apply (auto simp: split_beta') done have "- log b 1 \ - log b (integral\<^sup>L ?P ?f)" proof (intro le_imp_neg_le log_le[OF b_gt_1]) have If: "integrable ?P ?f" unfolding real_integrable_def proof (intro conjI) from neg show "(\\<^sup>+ x. - ?f x \?P) \ \" by simp from fin show "(\\<^sup>+ x. ?f x \?P) \ \" by simp qed simp then have "(\\<^sup>+ x. ?f x \?P) = (\x. ?f x \?P)" apply (rule nn_integral_eq_integral) apply (subst AE_density) apply simp apply (auto simp: space_pair_measure ennreal_neg) done with pos le1 show "0 < (\x. ?f x \?P)" "(\x. ?f x \?P) \ 1" by (simp_all add: one_ennreal.rep_eq zero_less_iff_neq_zero[symmetric]) qed also have "- log b (integral\<^sup>L ?P ?f) \ (\ x. - log b (?f x) \?P)" proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) show "AE x in ?P. ?f x \ {0<..}" unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]] using ae1 ae2 ae3 ae4 AE_space by eventually_elim (auto simp: space_pair_measure less_le) show "integrable ?P ?f" unfolding real_integrable_def using fin neg by (auto simp: split_beta') show "integrable ?P (\x. - log b (?f x))" apply (subst integrable_real_density) apply simp apply (auto simp: space_pair_measure) [] apply simp apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3]) apply simp apply simp using ae1 ae2 ae3 ae4 AE_space apply eventually_elim apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps less_le space_pair_measure) done qed (auto simp: b_gt_1 minus_log_convex) also have "\ = conditional_mutual_information b S T P X Y Z" unfolding \?eq\ apply (subst integral_real_density) apply simp apply (auto simp: space_pair_measure) [] apply simp apply (intro integral_cong_AE) using ae1 ae2 ae3 ae4 apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps space_pair_measure less_le) done finally show ?nonneg by simp qed lemma (in information_space) fixes Px :: "_ \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P" assumes Fx: "finite_entropy S X Px" assumes Fz: "finite_entropy P Z Pz" assumes Fyz: "finite_entropy (T \\<^sub>M P) (\x. (Y x, Z x)) Pyz" assumes Fxz: "finite_entropy (S \\<^sub>M P) (\x. (X x, Z x)) Pxz" assumes Fxyz: "finite_entropy (S \\<^sub>M T \\<^sub>M P) (\x. (X x, Y x, Z x)) Pxyz" shows conditional_mutual_information_generic_eq': "conditional_mutual_information b S T P X Y Z = (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \(S \\<^sub>M T \\<^sub>M P))" (is "?eq") and conditional_mutual_information_generic_nonneg': "0 \ conditional_mutual_information b S T P X Y Z" (is "?nonneg") proof - note Px = Fx[THEN finite_entropy_distributed, measurable] note Pz = Fz[THEN finite_entropy_distributed, measurable] note Pyz = Fyz[THEN finite_entropy_distributed, measurable] note Pxz = Fxz[THEN finite_entropy_distributed, measurable] note Pxyz = Fxyz[THEN finite_entropy_distributed, measurable] note Px_nn = Fx[THEN finite_entropy_nn] note Pz_nn = Fz[THEN finite_entropy_nn] note Pyz_nn = Fyz[THEN finite_entropy_nn] note Pxz_nn = Fxz[THEN finite_entropy_nn] note Pxyz_nn = Fxyz[THEN finite_entropy_nn] note Px' = Fx[THEN finite_entropy_measurable, measurable] note Pz' = Fz[THEN finite_entropy_measurable, measurable] note Pyz' = Fyz[THEN finite_entropy_measurable, measurable] note Pxz' = Fxz[THEN finite_entropy_measurable, measurable] note Pxyz' = Fxyz[THEN finite_entropy_measurable, measurable] interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret P: sigma_finite_measure P by fact interpret TP: pair_sigma_finite T P .. interpret SP: pair_sigma_finite S P .. interpret ST: pair_sigma_finite S T .. interpret SPT: pair_sigma_finite "S \\<^sub>M P" T .. interpret STP: pair_sigma_finite S "T \\<^sub>M P" .. interpret TPS: pair_sigma_finite "T \\<^sub>M P" S .. have TP: "sigma_finite_measure (T \\<^sub>M P)" .. have SP: "sigma_finite_measure (S \\<^sub>M P)" .. from Pxz Pxyz have distr_eq: "distr M (S \\<^sub>M P) (\x. (X x, Z x)) = distr (distr M (S \\<^sub>M T \\<^sub>M P) (\x. (X x, Y x, Z x))) (S \\<^sub>M P) (\(x, y, z). (x, z))" by (simp add: distr_distr comp_def) have "mutual_information b S P X Z = (\x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \(S \\<^sub>M P))" using Px Px_nn Pz Pz_nn Pxz Pxz_nn by (rule mutual_information_distr[OF S P]) (auto simp: space_pair_measure) also have "\ = (\(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \(S \\<^sub>M T \\<^sub>M P))" using b_gt_1 Pxz Pxz_nn Pxyz Pxyz_nn by (subst distributed_transform_integral[OF Pxyz _ Pxz, where T="\(x, y, z). (x, z)"]) (auto simp: split_beta') finally have mi_eq: "mutual_information b S P X Z = (\(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \(S \\<^sub>M T \\<^sub>M P))" . have ae1: "AE x in S \\<^sub>M T \\<^sub>M P. Px (fst x) = 0 \ Pxyz x = 0" by (intro subdensity_finite_entropy[of fst, OF _ Fxyz Fx]) auto moreover have ae2: "AE x in S \\<^sub>M T \\<^sub>M P. Pz (snd (snd x)) = 0 \ Pxyz x = 0" by (intro subdensity_finite_entropy[of "\x. snd (snd x)", OF _ Fxyz Fz]) auto moreover have ae3: "AE x in S \\<^sub>M T \\<^sub>M P. Pxz (fst x, snd (snd x)) = 0 \ Pxyz x = 0" by (intro subdensity_finite_entropy[of "\x. (fst x, snd (snd x))", OF _ Fxyz Fxz]) auto moreover have ae4: "AE x in S \\<^sub>M T \\<^sub>M P. Pyz (snd x) = 0 \ Pxyz x = 0" by (intro subdensity_finite_entropy[of snd, OF _ Fxyz Fyz]) auto ultimately have ae: "AE x in S \\<^sub>M T \\<^sub>M P. Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) - Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) = Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) " using AE_space proof eventually_elim case (elim x) show ?case proof cases assume "Pxyz x \ 0" with elim have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))" "0 < Pyz (snd x)" "0 < Pxyz x" using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn by (auto simp: space_pair_measure less_le) then show ?thesis using b_gt_1 by (simp add: log_simps less_imp_le field_simps) qed simp qed have "integrable (S \\<^sub>M T \\<^sub>M P) (\x. Pxyz x * log b (Pxyz x) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pyz (snd x)))" using finite_entropy_integrable[OF Fxyz] using finite_entropy_integrable_transform[OF Fx Pxyz Pxyz_nn, of fst] using finite_entropy_integrable_transform[OF Fyz Pxyz Pxyz_nn, of snd] by simp moreover have "(\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z)))) \ borel_measurable (S \\<^sub>M T \\<^sub>M P)" using Pxyz Px Pyz by simp ultimately have I1: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))" apply (rule integrable_cong_AE_imp) using ae1 ae4 AE_space by eventually_elim (insert Px_nn Pyz_nn Pxyz_nn, auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff space_pair_measure less_le) have "integrable (S \\<^sub>M T \\<^sub>M P) (\x. Pxyz x * log b (Pxz (fst x, snd (snd x))) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pz (snd (snd x))))" using finite_entropy_integrable_transform[OF Fxz Pxyz Pxyz_nn, of "\x. (fst x, snd (snd x))"] using finite_entropy_integrable_transform[OF Fx Pxyz Pxyz_nn, of fst] using finite_entropy_integrable_transform[OF Fz Pxyz Pxyz_nn, of "snd \ snd"] by simp moreover have "(\(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z))) \ borel_measurable (S \\<^sub>M T \\<^sub>M P)" using Pxyz Px Pz by auto ultimately have I2: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))" apply (rule integrable_cong_AE_imp) using ae1 ae2 ae3 ae4 AE_space by eventually_elim (insert Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn, auto simp: log_divide_eq log_mult_eq field_simps zero_less_mult_iff less_le space_pair_measure) from ae I1 I2 show ?eq unfolding conditional_mutual_information_def apply (subst mi_eq) apply (subst mutual_information_distr[OF S TP Px Px_nn Pyz Pyz_nn Pxyz Pxyz_nn]) apply simp apply simp apply (simp add: space_pair_measure) apply (subst Bochner_Integration.integral_diff[symmetric]) apply (auto intro!: integral_cong_AE simp: split_beta' simp del: Bochner_Integration.integral_diff) done let ?P = "density (S \\<^sub>M T \\<^sub>M P) Pxyz" interpret P: prob_space ?P unfolding distributed_distr_eq_density[OF Pxyz, symmetric] by (rule prob_space_distr) simp let ?Q = "density (T \\<^sub>M P) Pyz" interpret Q: prob_space ?Q unfolding distributed_distr_eq_density[OF Pyz, symmetric] by (rule prob_space_distr) simp let ?f = "\(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)" from subdensity_finite_entropy[of snd, OF _ Fyz Fz] have aeX1: "AE x in T \\<^sub>M P. Pz (snd x) = 0 \ Pyz x = 0" by (auto simp: comp_def) have aeX2: "AE x in T \\<^sub>M P. 0 \ Pz (snd x)" using Pz by (intro TP.AE_pair_measure) (auto intro: Pz_nn) have aeX3: "AE y in T \\<^sub>M P. (\\<^sup>+ x. ennreal (Pxz (x, snd y)) \S) = ennreal (Pz (snd y))" using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz] by (intro TP.AE_pair_measure) (auto ) have "(\\<^sup>+ x. ?f x \?P) \ (\\<^sup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \(S \\<^sub>M T \\<^sub>M P))" using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn by (subst nn_integral_density) (auto intro!: nn_integral_mono simp: space_pair_measure ennreal_mult[symmetric]) also have "\ = (\\<^sup>+(y, z). \\<^sup>+ x. ennreal (Pxz (x, z)) * ennreal (Pyz (y, z) / Pz z) \S \T \\<^sub>M P)" using Px_nn Pz_nn Pxz_nn Pyz_nn Pxyz_nn by (subst STP.nn_integral_snd[symmetric]) (auto simp add: split_beta' ennreal_mult[symmetric] space_pair_measure intro!: nn_integral_cong) also have "\ = (\\<^sup>+x. ennreal (Pyz x) * 1 \T \\<^sub>M P)" apply (rule nn_integral_cong_AE) using aeX1 aeX2 aeX3 AE_space apply eventually_elim proof (case_tac x, simp add: space_pair_measure) fix a b assume "Pz b = 0 \ Pyz (a, b) = 0" "0 \ Pz b" "a \ space T \ b \ space P" "(\\<^sup>+ x. ennreal (Pxz (x, b)) \S) = ennreal (Pz b)" then show "(\\<^sup>+ x. ennreal (Pxz (x, b)) * ennreal (Pyz (a, b) / Pz b) \S) = ennreal (Pyz (a, b))" using Pyz_nn[of "(a,b)"] by (subst nn_integral_multc) (auto simp: space_pair_measure ennreal_mult[symmetric]) qed also have "\ = 1" using Q.emeasure_space_1 Pyz_nn distributed_distr_eq_density[OF Pyz] by (subst nn_integral_density[symmetric]) auto finally have le1: "(\\<^sup>+ x. ?f x \?P) \ 1" . also have "\ < \" by simp finally have fin: "(\\<^sup>+ x. ?f x \?P) \ \" by simp have "(\\<^sup>+ x. ?f x \?P) \ 0" using Pxyz_nn apply (subst nn_integral_density) apply (simp_all add: split_beta' ennreal_mult'[symmetric] cong: nn_integral_cong) proof let ?g = "\x. ennreal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))" assume "(\\<^sup>+ x. ?g x \(S \\<^sub>M T \\<^sub>M P)) = 0" then have "AE x in S \\<^sub>M T \\<^sub>M P. ?g x = 0" by (intro nn_integral_0_iff_AE[THEN iffD1]) auto then have "AE x in S \\<^sub>M T \\<^sub>M P. Pxyz x = 0" using ae1 ae2 ae3 ae4 AE_space by eventually_elim (insert Px_nn Pz_nn Pxz_nn Pyz_nn, auto split: if_split_asm simp: mult_le_0_iff divide_le_0_iff space_pair_measure) then have "(\\<^sup>+ x. ennreal (Pxyz x) \S \\<^sub>M T \\<^sub>M P) = 0" by (subst nn_integral_cong_AE[of _ "\x. 0"]) auto with P.emeasure_space_1 show False by (subst (asm) emeasure_density) (auto cong: nn_integral_cong) qed then have pos: "0 < (\\<^sup>+ x. ?f x \?P)" by (simp add: zero_less_iff_neq_zero) have neg: "(\\<^sup>+ x. - ?f x \?P) = 0" using Pz_nn Pxz_nn Pyz_nn Pxyz_nn by (intro nn_integral_0_iff_AE[THEN iffD2]) (auto simp: split_beta' AE_density space_pair_measure intro!: AE_I2 ennreal_neg) have I3: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ Bochner_Integration.integrable_diff[OF I1 I2]]) using ae apply (auto simp: split_beta') done have "- log b 1 \ - log b (integral\<^sup>L ?P ?f)" proof (intro le_imp_neg_le log_le[OF b_gt_1]) have If: "integrable ?P ?f" unfolding real_integrable_def proof (intro conjI) from neg show "(\\<^sup>+ x. - ?f x \?P) \ \" by simp from fin show "(\\<^sup>+ x. ?f x \?P) \ \" by simp qed simp then have "(\\<^sup>+ x. ?f x \?P) = (\x. ?f x \?P)" using Pz_nn Pxz_nn Pyz_nn Pxyz_nn by (intro nn_integral_eq_integral) (auto simp: AE_density space_pair_measure) with pos le1 show "0 < (\x. ?f x \?P)" "(\x. ?f x \?P) \ 1" by (simp_all add: ) qed also have "- log b (integral\<^sup>L ?P ?f) \ (\ x. - log b (?f x) \?P)" proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) show "AE x in ?P. ?f x \ {0<..}" unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]] using ae1 ae2 ae3 ae4 AE_space by eventually_elim (insert Pxyz_nn Pyz_nn Pz_nn Pxz_nn, auto simp: space_pair_measure less_le) show "integrable ?P ?f" unfolding real_integrable_def using fin neg by (auto simp: split_beta') show "integrable ?P (\x. - log b (?f x))" using Pz_nn Pxz_nn Pyz_nn Pxyz_nn apply (subst integrable_real_density) apply simp apply simp apply simp apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3]) apply simp apply simp using ae1 ae2 ae3 ae4 AE_space apply eventually_elim apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps space_pair_measure less_le) done qed (auto simp: b_gt_1 minus_log_convex) also have "\ = conditional_mutual_information b S T P X Y Z" unfolding \?eq\ using Pz_nn Pxz_nn Pyz_nn Pxyz_nn apply (subst integral_real_density) apply simp apply simp apply simp apply (intro integral_cong_AE) using ae1 ae2 ae3 ae4 AE_space apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps space_pair_measure less_le) done finally show ?nonneg by simp qed lemma (in information_space) conditional_mutual_information_eq: assumes Pz: "simple_distributed M Z Pz" assumes Pyz: "simple_distributed M (\x. (Y x, Z x)) Pyz" assumes Pxz: "simple_distributed M (\x. (X x, Z x)) Pxz" assumes Pxyz: "simple_distributed M (\x. (X x, Y x, Z x)) Pxyz" shows "\(X ; Y | Z) = (\(x, y, z)\(\x. (X x, Y x, Z x))`space M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" proof (subst conditional_mutual_information_generic_eq[OF _ _ _ _ _ simple_distributed[OF Pz] _ simple_distributed_joint[OF Pyz] _ simple_distributed_joint[OF Pxz] _ simple_distributed_joint2[OF Pxyz]]) note simple_distributed_joint2_finite[OF Pxyz, simp] show "sigma_finite_measure (count_space (X ` space M))" by (simp add: sigma_finite_measure_count_space_finite) show "sigma_finite_measure (count_space (Y ` space M))" by (simp add: sigma_finite_measure_count_space_finite) show "sigma_finite_measure (count_space (Z ` space M))" by (simp add: sigma_finite_measure_count_space_finite) have "count_space (X ` space M) \\<^sub>M count_space (Y ` space M) \\<^sub>M count_space (Z ` space M) = count_space (X`space M \ Y`space M \ Z`space M)" (is "?P = ?C") by (simp add: pair_measure_count_space) let ?Px = "\x. measure M (X -` {x} \ space M)" have "(\x. (X x, Z x)) \ measurable M (count_space (X ` space M) \\<^sub>M count_space (Z ` space M))" using simple_distributed_joint[OF Pxz] by (rule distributed_measurable) from measurable_comp[OF this measurable_fst] have "random_variable (count_space (X ` space M)) X" by (simp add: comp_def) then have "simple_function M X" unfolding simple_function_def by (auto simp: measurable_count_space_eq2) then have "simple_distributed M X ?Px" by (rule simple_distributedI) (auto simp: measure_nonneg) then show "distributed M (count_space (X ` space M)) X ?Px" by (rule simple_distributed) let ?f = "(\x. if x \ (\x. (X x, Y x, Z x)) ` space M then Pxyz x else 0)" let ?g = "(\x. if x \ (\x. (Y x, Z x)) ` space M then Pyz x else 0)" let ?h = "(\x. if x \ (\x. (X x, Z x)) ` space M then Pxz x else 0)" show "integrable ?P (\(x, y, z). ?f (x, y, z) * log b (?f (x, y, z) / (?Px x * ?g (y, z))))" "integrable ?P (\(x, y, z). ?f (x, y, z) * log b (?h (x, z) / (?Px x * Pz z)))" by (auto intro!: integrable_count_space simp: pair_measure_count_space) let ?i = "\x y z. ?f (x, y, z) * log b (?f (x, y, z) / (?h (x, z) * (?g (y, z) / Pz z)))" let ?j = "\x y z. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z)))" have "(\(x, y, z). ?i x y z) = (\x. if x \ (\x. (X x, Y x, Z x)) ` space M then ?j (fst x) (fst (snd x)) (snd (snd x)) else 0)" by (auto intro!: ext) then show "(\ (x, y, z). ?i x y z \?P) = (\(x, y, z)\(\x. (X x, Y x, Z x)) ` space M. ?j x y z)" by (auto intro!: sum.cong simp add: \?P = ?C\ lebesgue_integral_count_space_finite simple_distributed_finite sum.If_cases split_beta') qed (insert Pz Pyz Pxz Pxyz, auto intro: measure_nonneg) lemma (in information_space) conditional_mutual_information_nonneg: assumes X: "simple_function M X" and Y: "simple_function M Y" and Z: "simple_function M Z" shows "0 \ \(X ; Y | Z)" proof - have [simp]: "count_space (X ` space M) \\<^sub>M count_space (Y ` space M) \\<^sub>M count_space (Z ` space M) = count_space (X`space M \ Y`space M \ Z`space M)" by (simp add: pair_measure_count_space X Y Z simple_functionD) note sf = sigma_finite_measure_count_space_finite[OF simple_functionD(1)] note sd = simple_distributedI[OF _ _ refl] note sp = simple_function_Pair show ?thesis apply (rule conditional_mutual_information_generic_nonneg[OF sf[OF X] sf[OF Y] sf[OF Z]]) apply (rule simple_distributed[OF sd[OF X]]) apply simp apply simp apply (rule simple_distributed[OF sd[OF Z]]) apply simp apply simp apply (rule simple_distributed_joint[OF sd[OF sp[OF Y Z]]]) apply simp apply simp apply (rule simple_distributed_joint[OF sd[OF sp[OF X Z]]]) apply simp apply simp apply (rule simple_distributed_joint2[OF sd[OF sp[OF X sp[OF Y Z]]]]) apply simp apply simp apply (auto intro!: integrable_count_space simp: X Y Z simple_functionD) done qed subsection \Conditional Entropy\ definition (in prob_space) "conditional_entropy b S T X Y = - (\(x, y). log b (enn2real (RN_deriv (S \\<^sub>M T) (distr M (S \\<^sub>M T) (\x. (X x, Y x))) (x, y)) / enn2real (RN_deriv T (distr M T Y) y)) \distr M (S \\<^sub>M T) (\x. (X x, Y x)))" abbreviation (in information_space) conditional_entropy_Pow ("\'(_ | _')") where "\(X | Y) \ conditional_entropy b (count_space (X`space M)) (count_space (Y`space M)) X Y" lemma (in information_space) conditional_entropy_generic_eq: fixes Pxy :: "_ \ real" and Py :: "'c \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Py[measurable]: "distributed M T Y Py" and Py_nn[simp]: "\x. x \ space T \ 0 \ Py x" assumes Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" and Pxy_nn[simp]: "\x y. x \ space S \ y \ space T \ 0 \ Pxy (x, y)" shows "conditional_entropy b S T X Y = - (\(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \(S \\<^sub>M T))" proof - interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. have [measurable]: "Py \ T \\<^sub>M borel" using Py Py_nn by (intro distributed_real_measurable) have measurable_Pxy[measurable]: "Pxy \ (S \\<^sub>M T) \\<^sub>M borel" using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) have "AE x in density (S \\<^sub>M T) (\x. ennreal (Pxy x)). Pxy x = enn2real (RN_deriv (S \\<^sub>M T) (distr M (S \\<^sub>M T) (\x. (X x, Y x))) x)" unfolding AE_density[OF distributed_borel_measurable, OF Pxy] unfolding distributed_distr_eq_density[OF Pxy] using distributed_RN_deriv[OF Pxy] by auto moreover have "AE x in density (S \\<^sub>M T) (\x. ennreal (Pxy x)). Py (snd x) = enn2real (RN_deriv T (distr M T Y) (snd x))" unfolding AE_density[OF distributed_borel_measurable, OF Pxy] unfolding distributed_distr_eq_density[OF Py] apply (rule ST.AE_pair_measure) apply auto using distributed_RN_deriv[OF Py] apply auto done ultimately have "conditional_entropy b S T X Y = - (\x. Pxy x * log b (Pxy x / Py (snd x)) \(S \\<^sub>M T))" unfolding conditional_entropy_def neg_equal_iff_equal apply (subst integral_real_density[symmetric]) apply (auto simp: distributed_distr_eq_density[OF Pxy] space_pair_measure intro!: integral_cong_AE) done then show ?thesis by (simp add: split_beta') qed lemma (in information_space) conditional_entropy_eq_entropy: fixes Px :: "'b \ real" and Py :: "'c \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Py[measurable]: "distributed M T Y Py" and Py_nn[simp]: "\x. x \ space T \ 0 \ Py x" assumes Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" and Pxy_nn[simp]: "\x y. x \ space S \ y \ space T \ 0 \ Pxy (x, y)" assumes I1: "integrable (S \\<^sub>M T) (\x. Pxy x * log b (Pxy x))" assumes I2: "integrable (S \\<^sub>M T) (\x. Pxy x * log b (Py (snd x)))" shows "conditional_entropy b S T X Y = entropy b (S \\<^sub>M T) (\x. (X x, Y x)) - entropy b T Y" proof - interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. have [measurable]: "Py \ T \\<^sub>M borel" using Py Py_nn by (intro distributed_real_measurable) have measurable_Pxy[measurable]: "Pxy \ (S \\<^sub>M T) \\<^sub>M borel" using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) have "entropy b T Y = - (\y. Py y * log b (Py y) \T)" by (rule entropy_distr[OF Py Py_nn]) also have "\ = - (\(x,y). Pxy (x,y) * log b (Py y) \(S \\<^sub>M T))" using b_gt_1 by (subst distributed_transform_integral[OF Pxy _ Py, where T=snd]) (auto intro!: Bochner_Integration.integral_cong simp: space_pair_measure) finally have e_eq: "entropy b T Y = - (\(x,y). Pxy (x,y) * log b (Py y) \(S \\<^sub>M T))" . have **: "\x. x \ space (S \\<^sub>M T) \ 0 \ Pxy x" by (auto simp: space_pair_measure) have ae2: "AE x in S \\<^sub>M T. Py (snd x) = 0 \ Pxy x = 0" by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair simp: space_pair_measure) moreover have ae4: "AE x in S \\<^sub>M T. 0 \ Py (snd x)" using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'') ultimately have "AE x in S \\<^sub>M T. 0 \ Pxy x \ 0 \ Py (snd x) \ (Pxy x = 0 \ (Pxy x \ 0 \ 0 < Pxy x \ 0 < Py (snd x)))" using AE_space by eventually_elim (auto simp: space_pair_measure less_le) then have ae: "AE x in S \\<^sub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))" by eventually_elim (auto simp: log_simps field_simps b_gt_1) have "conditional_entropy b S T X Y = - (\x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \(S \\<^sub>M T))" unfolding conditional_entropy_generic_eq[OF S T Py Py_nn Pxy Pxy_nn, simplified] neg_equal_iff_equal apply (intro integral_cong_AE) using ae apply auto done also have "\ = - (\x. Pxy x * log b (Pxy x) \(S \\<^sub>M T)) - - (\x. Pxy x * log b (Py (snd x)) \(S \\<^sub>M T))" by (simp add: Bochner_Integration.integral_diff[OF I1 I2]) finally show ?thesis using conditional_entropy_generic_eq[OF S T Py Py_nn Pxy Pxy_nn, simplified] entropy_distr[OF Pxy **, simplified] e_eq by (simp add: split_beta') qed lemma (in information_space) conditional_entropy_eq_entropy_simple: assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(X | Y) = entropy b (count_space (X`space M) \\<^sub>M count_space (Y`space M)) (\x. (X x, Y x)) - \(Y)" proof - have "count_space (X ` space M) \\<^sub>M count_space (Y ` space M) = count_space (X`space M \ Y`space M)" (is "?P = ?C") using X Y by (simp add: simple_functionD pair_measure_count_space) show ?thesis by (rule conditional_entropy_eq_entropy sigma_finite_measure_count_space_finite simple_functionD X Y simple_distributed simple_distributedI[OF _ _ refl] simple_distributed_joint simple_function_Pair integrable_count_space measure_nonneg)+ (auto simp: \?P = ?C\ measure_nonneg intro!: integrable_count_space simple_functionD X Y) qed lemma (in information_space) conditional_entropy_eq: assumes Y: "simple_distributed M Y Py" assumes XY: "simple_distributed M (\x. (X x, Y x)) Pxy" shows "\(X | Y) = - (\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))" proof (subst conditional_entropy_generic_eq[OF _ _ simple_distributed[OF Y] _ simple_distributed_joint[OF XY]]) have "finite ((\x. (X x, Y x))`space M)" using XY unfolding simple_distributed_def by auto from finite_imageI[OF this, of fst] have [simp]: "finite (X`space M)" by (simp add: image_comp comp_def) note Y[THEN simple_distributed_finite, simp] show "sigma_finite_measure (count_space (X ` space M))" by (simp add: sigma_finite_measure_count_space_finite) show "sigma_finite_measure (count_space (Y ` space M))" by (simp add: sigma_finite_measure_count_space_finite) let ?f = "(\x. if x \ (\x. (X x, Y x)) ` space M then Pxy x else 0)" have "count_space (X ` space M) \\<^sub>M count_space (Y ` space M) = count_space (X`space M \ Y`space M)" (is "?P = ?C") using Y by (simp add: simple_distributed_finite pair_measure_count_space) have eq: "(\(x, y). ?f (x, y) * log b (?f (x, y) / Py y)) = (\x. if x \ (\x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / Py (snd x)) else 0)" by auto from Y show "- (\ (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \?P) = - (\(x, y)\(\x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))" by (auto intro!: sum.cong simp add: \?P = ?C\ lebesgue_integral_count_space_finite simple_distributed_finite eq sum.If_cases split_beta') qed (insert Y XY, auto) lemma (in information_space) conditional_mutual_information_eq_conditional_entropy: assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(X ; X | Y) = \(X | Y)" proof - define Py where "Py x = (if x \ Y`space M then measure M (Y -` {x} \ space M) else 0)" for x define Pxy where "Pxy x = (if x \ (\x. (X x, Y x))`space M then measure M ((\x. (X x, Y x)) -` {x} \ space M) else 0)" for x define Pxxy where "Pxxy x = (if x \ (\x. (X x, X x, Y x))`space M then measure M ((\x. (X x, X x, Y x)) -` {x} \ space M) else 0)" for x let ?M = "X`space M \ X`space M \ Y`space M" note XY = simple_function_Pair[OF X Y] note XXY = simple_function_Pair[OF X XY] have Py: "simple_distributed M Y Py" using Y by (rule simple_distributedI) (auto simp: Py_def measure_nonneg) have Pxy: "simple_distributed M (\x. (X x, Y x)) Pxy" using XY by (rule simple_distributedI) (auto simp: Pxy_def measure_nonneg) have Pxxy: "simple_distributed M (\x. (X x, X x, Y x)) Pxxy" using XXY by (rule simple_distributedI) (auto simp: Pxxy_def measure_nonneg) have eq: "(\x. (X x, X x, Y x)) ` space M = (\(x, y). (x, x, y)) ` (\x. (X x, Y x)) ` space M" by auto have inj: "\A. inj_on (\(x, y). (x, x, y)) A" by (auto simp: inj_on_def) have Pxxy_eq: "\x y. Pxxy (x, x, y) = Pxy (x, y)" by (auto simp: Pxxy_def Pxy_def intro!: arg_cong[where f=prob]) have "AE x in count_space ((\x. (X x, Y x))`space M). Py (snd x) = 0 \ Pxy x = 0" using Py Pxy by (intro subdensity_real[of snd, OF _ Pxy[THEN simple_distributed] Py[THEN simple_distributed]]) (auto intro: measurable_Pair simp: AE_count_space) then show ?thesis apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy]) apply (subst conditional_entropy_eq[OF Py Pxy]) apply (auto intro!: sum.cong simp: Pxxy_eq sum_negf[symmetric] eq sum.reindex[OF inj] log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space) using Py[THEN simple_distributed] Pxy[THEN simple_distributed] apply (auto simp add: not_le AE_count_space less_le antisym simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy]) done qed lemma (in information_space) conditional_entropy_nonneg: assumes X: "simple_function M X" and Y: "simple_function M Y" shows "0 \ \(X | Y)" using conditional_mutual_information_eq_conditional_entropy[OF X Y] conditional_mutual_information_nonneg[OF X X Y] by simp subsection \Equalities\ lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr: fixes Px :: "'b \ real" and Py :: "'c \ real" and Pxy :: "('b \ 'c) \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Px[measurable]: "distributed M S X Px" and Px_nn[simp]: "\x. x \ space S \ 0 \ Px x" and Py[measurable]: "distributed M T Y Py" and Py_nn[simp]: "\x. x \ space T \ 0 \ Py x" and Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" and Pxy_nn[simp]: "\x y. x \ space S \ y \ space T \ 0 \ Pxy (x, y)" assumes Ix: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Px (fst x)))" assumes Iy: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Py (snd x)))" assumes Ixy: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Pxy x))" shows "mutual_information b S T X Y = entropy b S X + entropy b T Y - entropy b (S \\<^sub>M T) (\x. (X x, Y x))" proof - have [measurable]: "Px \ S \\<^sub>M borel" using Px Px_nn by (intro distributed_real_measurable) have [measurable]: "Py \ T \\<^sub>M borel" using Py Py_nn by (intro distributed_real_measurable) have measurable_Pxy[measurable]: "Pxy \ (S \\<^sub>M T) \\<^sub>M borel" using Pxy Pxy_nn by (intro distributed_real_measurable) (auto simp: space_pair_measure) have X: "entropy b S X = - (\x. Pxy x * log b (Px (fst x)) \(S \\<^sub>M T))" using b_gt_1 apply (subst entropy_distr[OF Px Px_nn], simp) apply (subst distributed_transform_integral[OF Pxy _ Px, where T=fst]) apply (auto intro!: integral_cong simp: space_pair_measure) done have Y: "entropy b T Y = - (\x. Pxy x * log b (Py (snd x)) \(S \\<^sub>M T))" using b_gt_1 apply (subst entropy_distr[OF Py Py_nn], simp) apply (subst distributed_transform_integral[OF Pxy _ Py, where T=snd]) apply (auto intro!: integral_cong simp: space_pair_measure) done interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. have ST: "sigma_finite_measure (S \\<^sub>M T)" .. have XY: "entropy b (S \\<^sub>M T) (\x. (X x, Y x)) = - (\x. Pxy x * log b (Pxy x) \(S \\<^sub>M T))" by (subst entropy_distr[OF Pxy]) (auto intro!: integral_cong simp: space_pair_measure) have "AE x in S \\<^sub>M T. Px (fst x) = 0 \ Pxy x = 0" by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair simp: space_pair_measure) moreover have "AE x in S \\<^sub>M T. Py (snd x) = 0 \ Pxy x = 0" by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair simp: space_pair_measure) moreover have "AE x in S \\<^sub>M T. 0 \ Px (fst x)" using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'') moreover have "AE x in S \\<^sub>M T. 0 \ Py (snd x)" using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'') ultimately have "AE x in S \\<^sub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))" (is "AE x in _. ?f x = ?g x") using AE_space proof eventually_elim case (elim x) show ?case proof cases assume "Pxy x \ 0" with elim have "0 < Px (fst x)" "0 < Py (snd x)" "0 < Pxy x" by (auto simp: space_pair_measure less_le) then show ?thesis using b_gt_1 by (simp add: log_simps less_imp_le field_simps) qed simp qed have "entropy b S X + entropy b T Y - entropy b (S \\<^sub>M T) (\x. (X x, Y x)) = integral\<^sup>L (S \\<^sub>M T) ?f" unfolding X Y XY apply (subst Bochner_Integration.integral_diff) apply (intro Bochner_Integration.integrable_diff Ixy Ix Iy)+ apply (subst Bochner_Integration.integral_diff) apply (intro Ixy Ix Iy)+ apply (simp add: field_simps) done also have "\ = integral\<^sup>L (S \\<^sub>M T) ?g" using \AE x in _. ?f x = ?g x\ by (intro integral_cong_AE) auto also have "\ = mutual_information b S T X Y" by (rule mutual_information_distr[OF S T Px _ Py _ Pxy _ , symmetric]) (auto simp: space_pair_measure) finally show ?thesis .. qed lemma (in information_space) mutual_information_eq_entropy_conditional_entropy': fixes Px :: "'b \ real" and Py :: "'c \ real" and Pxy :: "('b \ 'c) \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Px: "distributed M S X Px" "\x. x \ space S \ 0 \ Px x" and Py: "distributed M T Y Py" "\x. x \ space T \ 0 \ Py x" assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" "\x. x \ space (S \\<^sub>M T) \ 0 \ Pxy x" assumes Ix: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Px (fst x)))" assumes Iy: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Py (snd x)))" assumes Ixy: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Pxy x))" shows "mutual_information b S T X Y = entropy b S X - conditional_entropy b S T X Y" using mutual_information_eq_entropy_conditional_entropy_distr[OF S T Px Py Pxy Ix Iy Ixy] conditional_entropy_eq_entropy[OF S T Py Pxy Ixy Iy] by (simp add: space_pair_measure) lemma (in information_space) mutual_information_eq_entropy_conditional_entropy: assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y" shows "\(X ; Y) = \(X) - \(X | Y)" proof - have X: "simple_distributed M X (\x. measure M (X -` {x} \ space M))" using sf_X by (rule simple_distributedI) (auto simp: measure_nonneg) have Y: "simple_distributed M Y (\x. measure M (Y -` {x} \ space M))" using sf_Y by (rule simple_distributedI) (auto simp: measure_nonneg) have sf_XY: "simple_function M (\x. (X x, Y x))" using sf_X sf_Y by (rule simple_function_Pair) then have XY: "simple_distributed M (\x. (X x, Y x)) (\x. measure M ((\x. (X x, Y x)) -` {x} \ space M))" by (rule simple_distributedI) (auto simp: measure_nonneg) from simple_distributed_joint_finite[OF this, simp] have eq: "count_space (X ` space M) \\<^sub>M count_space (Y ` space M) = count_space (X ` space M \ Y ` space M)" by (simp add: pair_measure_count_space) have "\(X ; Y) = \(X) + \(Y) - entropy b (count_space (X`space M) \\<^sub>M count_space (Y`space M)) (\x. (X x, Y x))" using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite simple_distributed[OF X] measure_nonneg simple_distributed[OF Y] measure_nonneg simple_distributed_joint[OF XY] by (rule mutual_information_eq_entropy_conditional_entropy_distr) (auto simp: eq integrable_count_space measure_nonneg) then show ?thesis unfolding conditional_entropy_eq_entropy_simple[OF sf_X sf_Y] by simp qed lemma (in information_space) mutual_information_nonneg_simple: assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y" shows "0 \ \(X ; Y)" proof - have X: "simple_distributed M X (\x. measure M (X -` {x} \ space M))" using sf_X by (rule simple_distributedI) (auto simp: measure_nonneg) have Y: "simple_distributed M Y (\x. measure M (Y -` {x} \ space M))" using sf_Y by (rule simple_distributedI) (auto simp: measure_nonneg) have sf_XY: "simple_function M (\x. (X x, Y x))" using sf_X sf_Y by (rule simple_function_Pair) then have XY: "simple_distributed M (\x. (X x, Y x)) (\x. measure M ((\x. (X x, Y x)) -` {x} \ space M))" by (rule simple_distributedI) (auto simp: measure_nonneg) from simple_distributed_joint_finite[OF this, simp] have eq: "count_space (X ` space M) \\<^sub>M count_space (Y ` space M) = count_space (X ` space M \ Y ` space M)" by (simp add: pair_measure_count_space) show ?thesis by (rule mutual_information_nonneg[OF _ _ simple_distributed[OF X] _ simple_distributed[OF Y] _ simple_distributed_joint[OF XY]]) (simp_all add: eq integrable_count_space sigma_finite_measure_count_space_finite measure_nonneg) qed lemma (in information_space) conditional_entropy_less_eq_entropy: assumes X: "simple_function M X" and Z: "simple_function M Z" shows "\(X | Z) \ \(X)" proof - have "0 \ \(X ; Z)" using X Z by (rule mutual_information_nonneg_simple) also have "\(X ; Z) = \(X) - \(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] . finally show ?thesis by auto qed lemma (in information_space) fixes Px :: "'b \ real" and Py :: "'c \ real" and Pxy :: "('b \ 'c) \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Px: "finite_entropy S X Px" and Py: "finite_entropy T Y Py" assumes Pxy: "finite_entropy (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" shows "conditional_entropy b S T X Y \ entropy b S X" proof - have "0 \ mutual_information b S T X Y" by (rule mutual_information_nonneg') fact+ also have "\ = entropy b S X - conditional_entropy b S T X Y" apply (rule mutual_information_eq_entropy_conditional_entropy') using assms by (auto intro!: finite_entropy_integrable finite_entropy_distributed finite_entropy_integrable_transform[OF Px] finite_entropy_integrable_transform[OF Py] intro: finite_entropy_nn) finally show ?thesis by auto qed lemma (in information_space) entropy_chain_rule: assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(\x. (X x, Y x)) = \(X) + \(Y|X)" proof - note XY = simple_distributedI[OF simple_function_Pair[OF X Y] measure_nonneg refl] note YX = simple_distributedI[OF simple_function_Pair[OF Y X] measure_nonneg refl] note simple_distributed_joint_finite[OF this, simp] let ?f = "\x. prob ((\x. (X x, Y x)) -` {x} \ space M)" let ?g = "\x. prob ((\x. (Y x, X x)) -` {x} \ space M)" let ?h = "\x. if x \ (\x. (Y x, X x)) ` space M then prob ((\x. (Y x, X x)) -` {x} \ space M) else 0" have "\(\x. (X x, Y x)) = - (\x\(\x. (X x, Y x)) ` space M. ?f x * log b (?f x))" using XY by (rule entropy_simple_distributed) also have "\ = - (\x\(\(x, y). (y, x)) ` (\x. (X x, Y x)) ` space M. ?g x * log b (?g x))" by (subst (2) sum.reindex) (auto simp: inj_on_def intro!: sum.cong arg_cong[where f="\A. prob A * log b (prob A)"]) also have "\ = - (\x\(\x. (Y x, X x)) ` space M. ?h x * log b (?h x))" by (auto intro!: sum.cong) also have "\ = entropy b (count_space (Y ` space M) \\<^sub>M count_space (X ` space M)) (\x. (Y x, X x))" by (subst entropy_distr[OF simple_distributed_joint[OF YX]]) (auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite cong del: sum.cong_simp intro!: sum.mono_neutral_left measure_nonneg) finally have "\(\x. (X x, Y x)) = entropy b (count_space (Y ` space M) \\<^sub>M count_space (X ` space M)) (\x. (Y x, X x))" . then show ?thesis unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp qed lemma (in information_space) entropy_partition: assumes X: "simple_function M X" shows "\(X) = \(f \ X) + \(X|f \ X)" proof - note fX = simple_function_compose[OF X, of f] have eq: "(\x. ((f \ X) x, X x)) ` space M = (\x. (f x, x)) ` X ` space M" by auto have inj: "\A. inj_on (\x. (f x, x)) A" by (auto simp: inj_on_def) show ?thesis apply (subst entropy_chain_rule[symmetric, OF fX X]) apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] measure_nonneg refl]]) apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]]) unfolding eq apply (subst sum.reindex[OF inj]) apply (auto intro!: sum.cong arg_cong[where f="\A. prob A * log b (prob A)"]) done qed corollary (in information_space) entropy_data_processing: assumes X: "simple_function M X" shows "\(f \ X) \ \(X)" proof - note fX = simple_function_compose[OF X, of f] from X have "\(X) = \(f\X) + \(X|f\X)" by (rule entropy_partition) then show "\(f \ X) \ \(X)" by (simp only: conditional_entropy_nonneg [OF X fX] le_add_same_cancel1) qed corollary (in information_space) entropy_of_inj: assumes X: "simple_function M X" and inj: "inj_on f (X`space M)" shows "\(f \ X) = \(X)" proof (rule antisym) show "\(f \ X) \ \(X)" using entropy_data_processing[OF X] . next have sf: "simple_function M (f \ X)" using X by auto have "\(X) = \(the_inv_into (X`space M) f \ (f \ X))" unfolding o_assoc apply (subst entropy_simple_distributed[OF simple_distributedI[OF X measure_nonneg refl]]) apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_compose[OF X]], where f="\x. prob (X -` {x} \ space M)"]) apply (auto intro!: sum.cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def measure_nonneg) done also have "... \ \(f \ X)" using entropy_data_processing[OF sf] . finally show "\(X) \ \(f \ X)" . qed end diff --git a/src/HOL/Probability/Levy.thy b/src/HOL/Probability/Levy.thy --- a/src/HOL/Probability/Levy.thy +++ b/src/HOL/Probability/Levy.thy @@ -1,527 +1,529 @@ (* Title: HOL/Probability/Levy.thy Authors: Jeremy Avigad (CMU) *) section \The Levy inversion theorem, and the Levy continuity theorem.\ theory Levy imports Characteristic_Functions Helly_Selection Sinc_Integral begin subsection \The Levy inversion theorem\ (* Actually, this is not needed for us -- but it is useful for other purposes. (See Billingsley.) *) lemma Levy_Inversion_aux1: fixes a b :: real assumes "a \ b" shows "((\t. (iexp (-(t * a)) - iexp (-(t * b))) / (\ * t)) \ b - a) (at 0)" (is "(?F \ _) (at _)") proof - have 1: "cmod (?F t - (b - a)) \ a^2 / 2 * abs t + b^2 / 2 * abs t" if "t \ 0" for t proof - have "cmod (?F t - (b - a)) = cmod ( (iexp (-(t * a)) - (1 + \ * -(t * a))) / (\ * t) - (iexp (-(t * b)) - (1 + \ * -(t * b))) / (\ * t))" (is "_ = cmod (?one / (\ * t) - ?two / (\ * t))") using \t \ 0\ by (intro arg_cong[where f=norm]) (simp add: field_simps) also have "\ \ cmod (?one / (\ * t)) + cmod (?two / (\ * t))" by (rule norm_triangle_ineq4) also have "cmod (?one / (\ * t)) = cmod ?one / abs t" by (simp add: norm_divide norm_mult) also have "cmod (?two / (\ * t)) = cmod ?two / abs t" by (simp add: norm_divide norm_mult) also have "cmod ?one / abs t + cmod ?two / abs t \ ((- (a * t))^2 / 2) / abs t + ((- (b * t))^2 / 2) / abs t" apply (rule add_mono) apply (rule divide_right_mono) using iexp_approx1 [of "-(t * a)" 1] apply (simp add: field_simps eval_nat_numeral) apply force apply (rule divide_right_mono) using iexp_approx1 [of "-(t * b)" 1] apply (simp add: field_simps eval_nat_numeral) by force also have "\ = a^2 / 2 * abs t + b^2 / 2 * abs t" using \t \ 0\ apply (case_tac "t \ 0", simp add: field_simps power2_eq_square) using \t \ 0\ by (subst (1 2) abs_of_neg, auto simp add: field_simps power2_eq_square) finally show "cmod (?F t - (b - a)) \ a^2 / 2 * abs t + b^2 / 2 * abs t" . qed show ?thesis apply (rule LIM_zero_cancel) apply (rule tendsto_norm_zero_cancel) apply (rule real_LIM_sandwich_zero [OF _ _ 1]) apply (auto intro!: tendsto_eq_intros) done qed lemma Levy_Inversion_aux2: fixes a b t :: real assumes "a \ b" and "t \ 0" shows "cmod ((iexp (t * b) - iexp (t * a)) / (\ * t)) \ b - a" (is "?F \ _") proof - have "?F = cmod (iexp (t * a) * (iexp (t * (b - a)) - 1) / (\ * t))" using \t \ 0\ by (intro arg_cong[where f=norm]) (simp add: field_simps exp_diff exp_minus) also have "\ = cmod (iexp (t * (b - a)) - 1) / abs t" unfolding norm_divide norm_mult norm_exp_i_times using \t \ 0\ by (simp add: complex_eq_iff norm_mult) also have "\ \ abs (t * (b - a)) / abs t" using iexp_approx1 [of "t * (b - a)" 0] by (intro divide_right_mono) (auto simp add: field_simps eval_nat_numeral) also have "\ = b - a" using assms by (auto simp add: abs_mult) finally show ?thesis . qed (* TODO: refactor! *) theorem (in real_distribution) Levy_Inversion: fixes a b :: real assumes "a \ b" defines "\ \ measure M" and "\ \ char M" assumes "\ {a} = 0" and "\ {b} = 0" shows "(\T. 1 / (2 * pi) * (CLBINT t=-T..T. (iexp (-(t * a)) - iexp (-(t * b))) / (\ * t) * \ t)) \ \ {a<..b}" (is "(\T. 1 / (2 * pi) * (CLBINT t=-T..T. ?F t * \ t)) \ of_real (\ {a<..b})") proof - interpret P: pair_sigma_finite lborel M .. from bounded_Si obtain B where Bprop: "\T. abs (Si T) \ B" by auto from Bprop [of 0] have [simp]: "B \ 0" by auto let ?f = "\t x :: real. (iexp (t * (x - a)) - iexp(t * (x - b))) / (\ * t)" { fix T :: real assume "T \ 0" let ?f' = "\(t, x). indicator {-T<..R ?f t x" { fix x have 1: "complex_interval_lebesgue_integrable lborel u v (\t. ?f t x)" for u v :: real using Levy_Inversion_aux2[of "x - b" "x - a"] apply (simp add: interval_lebesgue_integrable_def set_integrable_def del: times_divide_eq_left) apply (intro integrableI_bounded_set_indicator[where B="b - a"] conjI impI) apply (auto intro!: AE_I [of _ _ "{0}"] simp: assms) done have "(CLBINT t. ?f' (t, x)) = (CLBINT t=-T..T. ?f t x)" using \T \ 0\ by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def) also have "\ = (CLBINT t=-T..(0 :: real). ?f t x) + (CLBINT t=(0 :: real)..T. ?f t x)" (is "_ = _ + ?t") using 1 by (intro interval_integral_sum[symmetric]) (simp add: min_absorb1 max_absorb2 \T \ 0\) also have "(CLBINT t=-T..(0 :: real). ?f t x) = (CLBINT t=(0::real)..T. ?f (-t) x)" by (subst interval_integral_reflect) auto also have "\ + ?t = (CLBINT t=(0::real)..T. ?f (-t) x + ?f t x)" using 1 by (intro interval_lebesgue_integral_add(2) [symmetric] interval_integrable_mirror[THEN iffD2]) simp_all also have "\ = (CLBINT t=(0::real)..T. ((iexp(t * (x - a)) - iexp (-(t * (x - a)))) - (iexp(t * (x - b)) - iexp (-(t * (x - b))))) / (\ * t))" using \T \ 0\ by (intro interval_integral_cong) (auto simp add: field_split_simps) also have "\ = (CLBINT t=(0::real)..T. complex_of_real( 2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t)))" using \T \ 0\ apply (intro interval_integral_cong) apply (simp add: field_simps cis.ctr Im_divide Re_divide Im_exp Re_exp complex_eq_iff) unfolding minus_diff_eq[symmetric, of "y * x" "y * a" for y a] sin_minus cos_minus apply (simp add: field_simps power2_eq_square) done also have "\ = complex_of_real (LBINT t=(0::real)..T. 2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t))" by (rule interval_lebesgue_integral_of_real) also have "\ = complex_of_real (2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))" apply (subst interval_lebesgue_integral_diff) apply (rule interval_lebesgue_integrable_mult_right, rule integrable_sinc')+ apply (subst interval_lebesgue_integral_mult_right)+ apply (simp add: zero_ereal_def[symmetric] LBINT_I0c_sin_scale_divide[OF \T \ 0\]) done finally have "(CLBINT t. ?f' (t, x)) = 2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))" . } note main_eq = this have "(CLBINT t=-T..T. ?F t * \ t) = (CLBINT t. (CLINT x | M. ?F t * iexp (t * x) * indicator {-T<..T \ 0\ unfolding \_def char_def interval_lebesgue_integral_def set_lebesgue_integral_def by (auto split: split_indicator intro!: Bochner_Integration.integral_cong) also have "\ = (CLBINT t. (CLINT x | M. ?f' (t, x)))" by (auto intro!: Bochner_Integration.integral_cong simp: field_simps exp_diff exp_minus split: split_indicator) also have "\ = (CLINT x | M. (CLBINT t. ?f' (t, x)))" proof (intro P.Fubini_integral [symmetric] integrableI_bounded_set [where B="b - a"]) show "emeasure (lborel \\<^sub>M M) ({- T<.. space M) < \" using \T \ 0\ by (subst emeasure_pair_measure_Times) (auto simp: ennreal_mult_less_top not_less top_unique) show "AE x\{- T<.. space M in lborel \\<^sub>M M. cmod (case x of (t, x) \ ?f' (t, x)) \ b - a" using Levy_Inversion_aux2[of "x - b" "x - a" for x] \a \ b\ by (intro AE_I [of _ _ "{0} \ UNIV"]) (force simp: emeasure_pair_measure_Times)+ qed (auto split: split_indicator split_indicator_asm) also have "\ = (CLINT x | M. (complex_of_real (2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))))" using main_eq by (intro Bochner_Integration.integral_cong, auto) also have "\ = complex_of_real (LINT x | M. (2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))))" by (rule integral_complex_of_real) finally have "(CLBINT t=-T..T. ?F t * \ t) = complex_of_real (LINT x | M. (2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))))" . } note main_eq2 = this have "(\T :: nat. LINT x | M. (2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \ (LINT x | M. 2 * pi * indicator {a<..b} x)" proof (rule integral_dominated_convergence [where w="\x. 4 * B"]) show "integrable M (\x. 4 * B)" by (rule integrable_const_bound [of _ "4 * B"]) auto next let ?S = "\n::nat. \x. sgn (x - a) * Si (n * \x - a\) - sgn (x - b) * Si (n * \x - b\)" { fix n x have "norm (?S n x) \ norm (sgn (x - a) * Si (n * \x - a\)) + norm (sgn (x - b) * Si (n * \x - b\))" by (rule norm_triangle_ineq4) also have "\ \ B + B" using Bprop by (intro add_mono) (auto simp: abs_mult abs_sgn_eq) finally have "norm (2 * ?S n x) \ 4 * B" by simp } then show "\n. AE x in M. norm (2 * ?S n x) \ 4 * B" by auto have "AE x in M. x \ a" "AE x in M. x \ b" using prob_eq_0[of "{a}"] prob_eq_0[of "{b}"] \\ {a} = 0\ \\ {b} = 0\ by (auto simp: \_def) then show "AE x in M. (\n. 2 * ?S n x) \ 2 * pi * indicator {a<..b} x" proof eventually_elim fix x assume x: "x \ a" "x \ b" then have "(\n. 2 * (sgn (x - a) * Si (\x - a\ * n) - sgn (x - b) * Si (\x - b\ * n))) \ 2 * (sgn (x - a) * (pi / 2) - sgn (x - b) * (pi / 2))" by (intro tendsto_intros filterlim_compose[OF Si_at_top] filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_real_sequentially) auto also have "(\n. 2 * (sgn (x - a) * Si (\x - a\ * n) - sgn (x - b) * Si (\x - b\ * n))) = (\n. 2 * ?S n x)" by (auto simp: ac_simps) also have "2 * (sgn (x - a) * (pi / 2) - sgn (x - b) * (pi / 2)) = 2 * pi * indicator {a<..b} x" using x \a \ b\ by (auto split: split_indicator) finally show "(\n. 2 * ?S n x) \ 2 * pi * indicator {a<..b} x" . qed qed simp_all also have "(LINT x | M. 2 * pi * indicator {a<..b} x) = 2 * pi * \ {a<..b}" by (simp add: \_def) finally have "(\T. LINT x | M. (2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \ 2 * pi * \ {a<..b}" . with main_eq2 show ?thesis by (auto intro!: tendsto_eq_intros) qed theorem Levy_uniqueness: fixes M1 M2 :: "real measure" assumes "real_distribution M1" "real_distribution M2" and "char M1 = char M2" shows "M1 = M2" proof - interpret M1: real_distribution M1 by (rule assms) interpret M2: real_distribution M2 by (rule assms) have "countable ({x. measure M1 {x} \ 0} \ {x. measure M2 {x} \ 0})" by (intro countable_Un M2.countable_support M1.countable_support) then have count: "countable {x. measure M1 {x} \ 0 \ measure M2 {x} \ 0}" by (simp add: Un_def) have "cdf M1 = cdf M2" proof (rule ext) fix x let ?D = "\x. \cdf M1 x - cdf M2 x\" { fix \ :: real assume "\ > 0" have "(?D \ 0) at_bot" using M1.cdf_lim_at_bot M2.cdf_lim_at_bot by (intro tendsto_eq_intros) auto then have "eventually (\y. ?D y < \ / 2 \ y \ x) at_bot" using \\ > 0\ by (simp only: tendsto_iff dist_real_def diff_0_right eventually_conj eventually_le_at_bot abs_idempotent) then obtain M where "\y. y \ M \ ?D y < \ / 2" "M \ x" unfolding eventually_at_bot_linorder by auto with open_minus_countable[OF count, of "{..< M}"] obtain a where "measure M1 {a} = 0" "measure M2 {a} = 0" "a < M" "a \ x" and 1: "?D a < \ / 2" by auto have "(?D \ ?D x) (at_right x)" using M1.cdf_is_right_cont [of x] M2.cdf_is_right_cont [of x] by (intro tendsto_intros) (auto simp add: continuous_within) then have "eventually (\y. \?D y - ?D x\ < \ / 2) (at_right x)" using \\ > 0\ by (simp only: tendsto_iff dist_real_def eventually_conj eventually_at_right_less) then obtain N where "N > x" "\y. x < y \ y < N \ \?D y - ?D x\ < \ / 2" by (auto simp add: eventually_at_right[OF less_add_one]) with open_minus_countable[OF count, of "{x <..< N}"] obtain b where "x < b" "b < N" "measure M1 {b} = 0" "measure M2 {b} = 0" and 2: "\?D b - ?D x\ < \ / 2" by (auto simp: abs_minus_commute) from \a \ x\ \x < b\ have "a < b" "a \ b" by auto from \char M1 = char M2\ M1.Levy_Inversion [OF \a \ b\ \measure M1 {a} = 0\ \measure M1 {b} = 0\] M2.Levy_Inversion [OF \a \ b\ \measure M2 {a} = 0\ \measure M2 {b} = 0\] have "complex_of_real (measure M1 {a<..b}) = complex_of_real (measure M2 {a<..b})" by (intro LIMSEQ_unique) auto then have "?D a = ?D b" unfolding of_real_eq_iff M1.cdf_diff_eq [OF \a < b\, symmetric] M2.cdf_diff_eq [OF \a < b\, symmetric] by simp then have "?D x = \(?D b - ?D x) - ?D a\" by simp also have "\ \ \?D b - ?D x\ + \?D a\" by (rule abs_triangle_ineq4) also have "\ \ \ / 2 + \ / 2" using 1 2 by (intro add_mono) auto finally have "?D x \ \" by simp } then show "cdf M1 x = cdf M2 x" by (metis abs_le_zero_iff dense_ge eq_iff_diff_eq_0) qed thus ?thesis by (rule cdf_unique [OF \real_distribution M1\ \real_distribution M2\]) qed subsection \The Levy continuity theorem\ theorem levy_continuity1: fixes M :: "nat \ real measure" and M' :: "real measure" assumes "\n. real_distribution (M n)" "real_distribution M'" "weak_conv_m M M'" shows "(\n. char (M n) t) \ char M' t" unfolding char_def using assms by (rule weak_conv_imp_integral_bdd_continuous_conv) auto theorem levy_continuity: fixes M :: "nat \ real measure" and M' :: "real measure" assumes real_distr_M : "\n. real_distribution (M n)" and real_distr_M': "real_distribution M'" and char_conv: "\t. (\n. char (M n) t) \ char M' t" shows "weak_conv_m M M'" proof - interpret Mn: real_distribution "M n" for n by fact interpret M': real_distribution M' by fact have *: "\u x. u > 0 \ x \ 0 \ (CLBINT t:{-u..u}. 1 - iexp (t * x)) = 2 * (u - sin (u * x) / x)" proof - fix u :: real and x :: real assume "u > 0" and "x \ 0" hence "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = (CLBINT t=-u..u. 1 - iexp (t * x))" by (subst interval_integral_Icc, auto) also have "\ = (CLBINT t=-u..0. 1 - iexp (t * x)) + (CLBINT t=0..u. 1 - iexp (t * x))" using \u > 0\ apply (subst interval_integral_sum) apply (simp add: min_absorb1 min_absorb2 max_absorb1 max_absorb2) apply (rule interval_integrable_isCont) apply auto done also have "\ = (CLBINT t=ereal 0..u. 1 - iexp (t * -x)) + (CLBINT t=ereal 0..u. 1 - iexp (t * x))" apply (subgoal_tac "0 = ereal 0", erule ssubst) by (subst interval_integral_reflect, auto) also have "\ = (LBINT t=ereal 0..u. 2 - 2 * cos (t * x))" apply (subst interval_lebesgue_integral_add (2) [symmetric]) apply ((rule interval_integrable_isCont, auto)+) [2] unfolding exp_Euler cos_of_real apply (simp add: of_real_mult interval_lebesgue_integral_of_real[symmetric]) done also have "\ = 2 * u - 2 * sin (u * x) / x" by (subst interval_lebesgue_integral_diff) (auto intro!: interval_integrable_isCont simp: interval_lebesgue_integral_of_real integral_cos [OF \x \ 0\] mult.commute[of _ x]) finally show "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = 2 * (u - sin (u * x) / x)" by (simp add: field_simps) qed have main_bound: "\u n. u > 0 \ Re (CLBINT t:{-u..u}. 1 - char (M n) t) \ u * measure (M n) {x. abs x \ 2 / u}" proof - fix u :: real and n assume "u > 0" interpret P: pair_sigma_finite "M n" lborel .. (* TODO: put this in the real_distribution locale as a simp rule? *) have Mn1 [simp]: "measure (M n) UNIV = 1" by (metis Mn.prob_space Mn.space_eq_univ) (* TODO: make this automatic somehow? *) have Mn2 [simp]: "\x. complex_integrable (M n) (\t. exp (\ * complex_of_real (x * t)))" by (rule Mn.integrable_const_bound [where B = 1], auto) have Mn3: "set_integrable (M n \\<^sub>M lborel) (UNIV \ {- u..u}) (\a. 1 - exp (\ * complex_of_real (snd a * fst a)))" using \0 < u\ unfolding set_integrable_def by (intro integrableI_bounded_set_indicator [where B="2"]) (auto simp: lborel.emeasure_pair_measure_Times ennreal_mult_less_top not_less top_unique split: split_indicator intro!: order_trans [OF norm_triangle_ineq4]) have "(CLBINT t:{-u..u}. 1 - char (M n) t) = (CLBINT t:{-u..u}. (CLINT x | M n. 1 - iexp (t * x)))" unfolding char_def by (rule set_lebesgue_integral_cong, auto simp del: of_real_mult) also have "\ = (CLBINT t. (CLINT x | M n. indicator {-u..u} t *\<^sub>R (1 - iexp (t * x))))" unfolding set_lebesgue_integral_def by (rule Bochner_Integration.integral_cong) (auto split: split_indicator) also have "\ = (CLINT x | M n. (CLBINT t:{-u..u}. 1 - iexp (t * x)))" using Mn3 by (subst P.Fubini_integral) (auto simp: indicator_times split_beta' set_integrable_def set_lebesgue_integral_def) also have "\ = (CLINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))" using \u > 0\ by (intro Bochner_Integration.integral_cong, auto simp add: * simp del: of_real_mult) also have "\ = (LINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))" by (rule integral_complex_of_real) finally have "Re (CLBINT t:{-u..u}. 1 - char (M n) t) = (LINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))" by simp also have "\ \ (LINT x : {x. abs x \ 2 / u} | M n. u)" proof - have "complex_integrable (M n) (\x. CLBINT t:{-u..u}. 1 - iexp (snd (x, t) * fst (x, t)))" using Mn3 unfolding set_integrable_def set_lebesgue_integral_def by (intro P.integrable_fst) (simp add: indicator_times split_beta') hence "complex_integrable (M n) (\x. if x = 0 then 0 else 2 * (u - sin (u * x) / x))" using \u > 0\ unfolding set_integrable_def by (subst integrable_cong) (auto simp add: * simp del: of_real_mult) hence **: "integrable (M n) (\x. if x = 0 then 0 else 2 * (u - sin (u * x) / x))" unfolding complex_of_real_integrable_eq . have "2 * sin x \ x" if "2 \ x" for x :: real by (rule order_trans[OF _ \2 \ x\]) auto moreover have "x \ 2 * sin x" if "x \ - 2" for x :: real by (rule order_trans[OF \x \ - 2\]) auto moreover have "x < 0 \ x \ sin x" for x :: real using sin_x_le_x[of "-x"] by simp ultimately show ?thesis using \u > 0\ unfolding set_lebesgue_integral_def by (intro integral_mono [OF _ **]) (auto simp: divide_simps sin_x_le_x mult.commute[of u] mult_neg_pos top_unique less_top[symmetric] split: split_indicator) qed also (xtrans) have "(LINT x : {x. abs x \ 2 / u} | M n. u) = u * measure (M n) {x. abs x \ 2 / u}" unfolding set_lebesgue_integral_def by (simp add: Mn.emeasure_eq_measure) finally show "Re (CLBINT t:{-u..u}. 1 - char (M n) t) \ u * measure (M n) {x. abs x \ 2 / u}" . qed have tight_aux: "\\. \ > 0 \ \a b. a < b \ (\n. 1 - \ < measure (M n) {a<..b})" proof - fix \ :: real assume "\ > 0" note M'.isCont_char [of 0] hence "\d>0. \t. abs t < d \ cmod (char M' t - 1) < \ / 4" apply (subst (asm) continuous_at_eps_delta) apply (drule_tac x = "\ / 4" in spec) using \\ > 0\ by (auto simp add: dist_real_def dist_complex_def M'.char_zero) then obtain d where "d > 0 \ (\t. (abs t < d \ cmod (char M' t - 1) < \ / 4))" .. hence d0: "d > 0" and d1: "\t. abs t < d \ cmod (char M' t - 1) < \ / 4" by auto have 1: "\x. cmod (1 - char M' x) \ 2" by (rule order_trans [OF norm_triangle_ineq4], auto simp add: M'.cmod_char_le_1) then have 2: "\u v. complex_set_integrable lborel {u..v} (\x. 1 - char M' x)" unfolding set_integrable_def by (intro integrableI_bounded_set_indicator[where B=2]) (auto simp: emeasure_lborel_Icc_eq) have 3: "\u v. integrable lborel (\x. indicat_real {u..v} x *\<^sub>R cmod (1 - char M' x))" by (intro borel_integrable_compact[OF compact_Icc] continuous_at_imp_continuous_on continuous_intros ballI M'.isCont_char continuous_intros) have "cmod (CLBINT t:{-d/2..d/2}. 1 - char M' t) \ LBINT t:{-d/2..d/2}. cmod (1 - char M' t)" unfolding set_lebesgue_integral_def using integral_norm_bound[of _ "\x. indicator {u..v} x *\<^sub>R (1 - char M' x)" for u v] by simp also have 4: "\ \ LBINT t:{-d/2..d/2}. \ / 4" unfolding set_lebesgue_integral_def apply (rule integral_mono [OF 3]) apply (simp add: emeasure_lborel_Icc_eq) apply (case_tac "x \ {-d/2..d/2}") apply auto apply (subst norm_minus_commute) apply (rule less_imp_le) apply (rule d1 [simplified]) using d0 apply auto done also from d0 4 have "\ = d * \ / 4" unfolding set_lebesgue_integral_def by simp finally have bound: "cmod (CLBINT t:{-d/2..d/2}. 1 - char M' t) \ d * \ / 4" . have "cmod (1 - char (M n) x) \ 2" for n x by (rule order_trans [OF norm_triangle_ineq4], auto simp add: Mn.cmod_char_le_1) then have "(\n. CLBINT t:{-d/2..d/2}. 1 - char (M n) t) \ (CLBINT t:{-d/2..d/2}. 1 - char M' t)" unfolding set_lebesgue_integral_def apply (intro integral_dominated_convergence[where w="\x. indicator {-d/2..d/2} x *\<^sub>R 2"]) apply (auto intro!: char_conv tendsto_intros simp: emeasure_lborel_Icc_eq split: split_indicator) done hence "eventually (\n. cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - (CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \ / 4) sequentially" using d0 \\ > 0\ apply (subst (asm) tendsto_iff) by (subst (asm) dist_complex_def, drule spec, erule mp, auto) hence "\N. \n \ N. cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - (CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \ / 4" by (simp add: eventually_sequentially) - then guess N .. + then obtain N + where "\n\N. cmod ((CLBINT t:{- d / 2..d / 2}. 1 - char (M n) t) - + (CLBINT t:{- d / 2..d / 2}. 1 - char M' t)) < d * \ / 4" .. hence N: "\n. n \ N \ cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - (CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \ / 4" by auto { fix n assume "n \ N" have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) = cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - (CLBINT t:{-d/2..d/2}. 1 - char M' t) + (CLBINT t:{-d/2..d/2}. 1 - char M' t))" by simp also have "\ \ cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - (CLBINT t:{-d/2..d/2}. 1 - char M' t)) + cmod(CLBINT t:{-d/2..d/2}. 1 - char M' t)" by (rule norm_triangle_ineq) also have "\ < d * \ / 4 + d * \ / 4" by (rule add_less_le_mono [OF N [OF \n \ N\] bound]) also have "\ = d * \ / 2" by auto finally have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) < d * \ / 2" . hence "d * \ / 2 > Re (CLBINT t:{-d/2..d/2}. 1 - char (M n) t)" by (rule order_le_less_trans [OF complex_Re_le_cmod]) hence "d * \ / 2 > Re (CLBINT t:{-(d/2)..d/2}. 1 - char (M n) t)" (is "_ > ?lhs") by simp also have "?lhs \ (d / 2) * measure (M n) {x. abs x \ 2 / (d / 2)}" using d0 by (intro main_bound, simp) finally (xtrans) have "d * \ / 2 > (d / 2) * measure (M n) {x. abs x \ 2 / (d / 2)}" . with d0 \\ > 0\ have "\ > measure (M n) {x. abs x \ 2 / (d / 2)}" by (simp add: field_simps) hence "\ > 1 - measure (M n) (UNIV - {x. abs x \ 2 / (d / 2)})" apply (subst Mn.borel_UNIV [symmetric]) by (subst Mn.prob_compl, auto) also have "UNIV - {x. abs x \ 2 / (d / 2)} = {x. -(4 / d) < x \ x < (4 / d)}" using d0 apply (auto simp add: field_simps) (* very annoying -- this should be automatic *) apply (case_tac "x \ 0", auto simp add: field_simps) apply (subgoal_tac "0 \ x * d", arith, rule mult_nonneg_nonneg, auto) apply (case_tac "x \ 0", auto simp add: field_simps) apply (subgoal_tac "x * d \ 0", arith) apply (rule mult_nonpos_nonneg, auto) by (case_tac "x \ 0", auto simp add: field_simps) finally have "measure (M n) {x. -(4 / d) < x \ x < (4 / d)} > 1 - \" by auto } note 6 = this { fix n :: nat have *: "(UN (k :: nat). {- real k<..real k}) = UNIV" by (auto, metis leI le_less_trans less_imp_le minus_less_iff reals_Archimedean2) have "(\k. measure (M n) {- real k<..real k}) \ measure (M n) (UN (k :: nat). {- real k<..real k})" by (rule Mn.finite_Lim_measure_incseq, auto simp add: incseq_def) hence "(\k. measure (M n) {- real k<..real k}) \ 1" using Mn.prob_space unfolding * Mn.borel_UNIV by simp hence "eventually (\k. measure (M n) {- real k<..real k} > 1 - \) sequentially" apply (elim order_tendstoD (1)) using \\ > 0\ by auto } note 7 = this { fix n :: nat have "eventually (\k. \m < n. measure (M m) {- real k<..real k} > 1 - \) sequentially" (is "?P n") proof (induct n) case (Suc n) with 7[of n] show ?case by eventually_elim (auto simp add: less_Suc_eq) qed simp } note 8 = this from 8 [of N] have "\K :: nat. \k \ K. \m < Sigma_Algebra.measure (M m) {- real k<..real k}" by (auto simp add: eventually_sequentially) hence "\K :: nat. \m < Sigma_Algebra.measure (M m) {- real K<..real K}" by auto then obtain K :: nat where "\m < Sigma_Algebra.measure (M m) {- real K<..real K}" .. hence K: "\m. m < N \ 1 - \ < Sigma_Algebra.measure (M m) {- real K<..real K}" by auto let ?K' = "max K (4 / d)" have "-?K' < ?K' \ (\n. 1 - \ < measure (M n) {-?K'<..?K'})" using d0 apply auto apply (rule max.strict_coboundedI2, auto) proof - fix n show " 1 - \ < measure (M n) {- max (real K) (4 / d)<..max (real K) (4 / d)}" apply (case_tac "n < N") apply (rule order_less_le_trans) apply (erule K) apply (rule Mn.finite_measure_mono, auto) apply (rule order_less_le_trans) apply (rule 6, erule leI) by (rule Mn.finite_measure_mono, auto) qed thus "\a b. a < b \ (\n. 1 - \ < measure (M n) {a<..b})" by (intro exI) qed have tight: "tight M" by (auto simp: tight_def intro: assms tight_aux) show ?thesis proof (rule tight_subseq_weak_converge [OF real_distr_M real_distr_M' tight]) fix s \ assume s: "strict_mono (s :: nat \ nat)" assume nu: "weak_conv_m (M \ s) \" assume *: "real_distribution \" have 2: "\n. real_distribution ((M \ s) n)" unfolding comp_def by (rule assms) have 3: "\t. (\n. char ((M \ s) n) t) \ char \ t" by (intro levy_continuity1 [OF 2 * nu]) have 4: "\t. (\n. char ((M \ s) n) t) = ((\n. char (M n) t) \ s)" by (rule ext, simp) have 5: "\t. (\n. char ((M \ s) n) t) \ char M' t" by (subst 4, rule LIMSEQ_subseq_LIMSEQ [OF _ s], rule assms) hence "char \ = char M'" by (intro ext, intro LIMSEQ_unique [OF 3 5]) hence "\ = M'" by (rule Levy_uniqueness [OF * \real_distribution M'\]) thus "weak_conv_m (M \ s) M'" by (elim subst) (rule nu) qed qed end diff --git a/src/HOL/Probability/Probability_Measure.thy b/src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy +++ b/src/HOL/Probability/Probability_Measure.thy @@ -1,1305 +1,1320 @@ (* Title: HOL/Probability/Probability_Measure.thy Author: Johannes Hölzl, TU München Author: Armin Heller, TU München *) section \Probability measure\ theory Probability_Measure imports "HOL-Analysis.Analysis" begin locale prob_space = finite_measure + assumes emeasure_space_1: "emeasure M (space M) = 1" lemma prob_spaceI[Pure.intro!]: assumes *: "emeasure M (space M) = 1" shows "prob_space M" proof - interpret finite_measure M proof show "emeasure M (space M) \ \" using * by simp qed show "prob_space M" by standard fact qed lemma prob_space_imp_sigma_finite: "prob_space M \ sigma_finite_measure M" unfolding prob_space_def finite_measure_def by simp abbreviation (in prob_space) "events \ sets M" abbreviation (in prob_space) "prob \ measure M" abbreviation (in prob_space) "random_variable M' X \ X \ measurable M M'" abbreviation (in prob_space) "expectation \ integral\<^sup>L M" abbreviation (in prob_space) "variance X \ integral\<^sup>L M (\x. (X x - expectation X)\<^sup>2)" lemma (in prob_space) finite_measure [simp]: "finite_measure M" by unfold_locales lemma (in prob_space) prob_space_distr: assumes f: "f \ measurable M M'" shows "prob_space (distr M M' f)" proof (rule prob_spaceI) have "f -` space M' \ space M = space M" using f by (auto dest: measurable_space) with f show "emeasure (distr M M' f) (space (distr M M' f)) = 1" by (auto simp: emeasure_distr emeasure_space_1) qed lemma prob_space_distrD: assumes f: "f \ measurable M N" and M: "prob_space (distr M N f)" shows "prob_space M" proof interpret M: prob_space "distr M N f" by fact have "f -` space N \ space M = space M" using f[THEN measurable_space] by auto then show "emeasure M (space M) = 1" using M.emeasure_space_1 by (simp add: emeasure_distr[OF f]) qed lemma (in prob_space) prob_space: "prob (space M) = 1" using emeasure_space_1 unfolding measure_def by (simp add: one_ennreal.rep_eq) lemma (in prob_space) prob_le_1[simp, intro]: "prob A \ 1" using bounded_measure[of A] by (simp add: prob_space) lemma (in prob_space) not_empty: "space M \ {}" using prob_space by auto lemma (in prob_space) emeasure_eq_1_AE: "S \ sets M \ AE x in M. x \ S \ emeasure M S = 1" by (subst emeasure_eq_AE[where B="space M"]) (auto simp: emeasure_space_1) lemma (in prob_space) emeasure_le_1: "emeasure M S \ 1" unfolding ennreal_1[symmetric] emeasure_eq_measure by (subst ennreal_le_iff) auto lemma (in prob_space) emeasure_ge_1_iff: "emeasure M A \ 1 \ emeasure M A = 1" by (rule iffI, intro antisym emeasure_le_1) simp_all lemma (in prob_space) AE_iff_emeasure_eq_1: assumes [measurable]: "Measurable.pred M P" shows "(AE x in M. P x) \ emeasure M {x\space M. P x} = 1" proof - have *: "{x \ space M. \ P x} = space M - {x\space M. P x}" by auto show ?thesis by (auto simp add: ennreal_minus_eq_0 * emeasure_compl emeasure_space_1 AE_iff_measurable[OF _ refl] intro: antisym emeasure_le_1) qed lemma (in prob_space) measure_le_1: "emeasure M X \ 1" using emeasure_space[of M X] by (simp add: emeasure_space_1) lemma (in prob_space) measure_ge_1_iff: "measure M A \ 1 \ measure M A = 1" by (auto intro!: antisym) lemma (in prob_space) AE_I_eq_1: assumes "emeasure M {x\space M. P x} = 1" "{x\space M. P x} \ sets M" shows "AE x in M. P x" proof (rule AE_I) show "emeasure M (space M - {x \ space M. P x}) = 0" using assms emeasure_space_1 by (simp add: emeasure_compl) qed (insert assms, auto) lemma prob_space_restrict_space: "S \ sets M \ emeasure M S = 1 \ prob_space (restrict_space M S)" by (intro prob_spaceI) (simp add: emeasure_restrict_space space_restrict_space) lemma (in prob_space) prob_compl: assumes A: "A \ events" shows "prob (space M - A) = 1 - prob A" using finite_measure_compl[OF A] by (simp add: prob_space) lemma (in prob_space) AE_in_set_eq_1: assumes A[measurable]: "A \ events" shows "(AE x in M. x \ A) \ prob A = 1" proof - have *: "{x\space M. x \ A} = A" using A[THEN sets.sets_into_space] by auto show ?thesis by (subst AE_iff_emeasure_eq_1) (auto simp: emeasure_eq_measure *) qed lemma (in prob_space) AE_False: "(AE x in M. False) \ False" proof assume "AE x in M. False" then have "AE x in M. x \ {}" by simp then show False by (subst (asm) AE_in_set_eq_1) auto qed simp lemma (in prob_space) AE_prob_1: assumes "prob A = 1" shows "AE x in M. x \ A" proof - from \prob A = 1\ have "A \ events" by (metis measure_notin_sets zero_neq_one) with AE_in_set_eq_1 assms show ?thesis by simp qed lemma (in prob_space) AE_const[simp]: "(AE x in M. P) \ P" by (cases P) (auto simp: AE_False) lemma (in prob_space) ae_filter_bot: "ae_filter M \ bot" by (simp add: trivial_limit_def) lemma (in prob_space) AE_contr: assumes ae: "AE \ in M. P \" "AE \ in M. \ P \" shows False proof - from ae have "AE \ in M. False" by eventually_elim auto then show False by auto qed lemma (in prob_space) integral_ge_const: fixes c :: real shows "integrable M f \ (AE x in M. c \ f x) \ c \ (\x. f x \M)" using integral_mono_AE[of M "\x. c" f] prob_space by simp lemma (in prob_space) integral_le_const: fixes c :: real shows "integrable M f \ (AE x in M. f x \ c) \ (\x. f x \M) \ c" using integral_mono_AE[of M f "\x. c"] prob_space by simp lemma (in prob_space) nn_integral_ge_const: "(AE x in M. c \ f x) \ c \ (\\<^sup>+x. f x \M)" using nn_integral_mono_AE[of "\x. c" f M] emeasure_space_1 by (simp split: if_split_asm) lemma (in prob_space) expectation_less: fixes X :: "_ \ real" assumes [simp]: "integrable M X" assumes gt: "AE x in M. X x < b" shows "expectation X < b" proof - have "expectation X < expectation (\x. b)" using gt emeasure_space_1 by (intro integral_less_AE_space) auto then show ?thesis using prob_space by simp qed lemma (in prob_space) expectation_greater: fixes X :: "_ \ real" assumes [simp]: "integrable M X" assumes gt: "AE x in M. a < X x" shows "a < expectation X" proof - have "expectation (\x. a) < expectation X" using gt emeasure_space_1 by (intro integral_less_AE_space) auto then show ?thesis using prob_space by simp qed lemma (in prob_space) jensens_inequality: fixes q :: "real \ real" assumes X: "integrable M X" "AE x in M. X x \ I" assumes I: "I = {a <..< b} \ I = {a <..} \ I = {..< b} \ I = UNIV" assumes q: "integrable M (\x. q (X x))" "convex_on I q" shows "q (expectation X) \ expectation (\x. q (X x))" proof - let ?F = "\x. Inf ((\t. (q x - q t) / (x - t)) ` ({x<..} \ I))" from X(2) AE_False have "I \ {}" by auto from I have "open I" by auto note I moreover { assume "I \ {a <..}" with X have "a < expectation X" by (intro expectation_greater) auto } moreover { assume "I \ {..< b}" with X have "expectation X < b" by (intro expectation_less) auto } ultimately have "expectation X \ I" by (elim disjE) (auto simp: subset_eq) moreover { fix y assume y: "y \ I" with q(2) \open I\ have "Sup ((\x. q x + ?F x * (y - x)) ` I) = q y" by (auto intro!: cSup_eq_maximum convex_le_Inf_differential image_eqI [OF _ y] simp: interior_open) } ultimately have "q (expectation X) = Sup ((\x. q x + ?F x * (expectation X - x)) ` I)" by simp also have "\ \ expectation (\w. q (X w))" proof (rule cSup_least) show "(\x. q x + ?F x * (expectation X - x)) ` I \ {}" using \I \ {}\ by auto next fix k assume "k \ (\x. q x + ?F x * (expectation X - x)) ` I" - then guess x .. note x = this + then obtain x + where x: "k = q x + (INF t\{x<..} \ I. (q x - q t) / (x - t)) * (expectation X - x)" "x \ I" .. have "q x + ?F x * (expectation X - x) = expectation (\w. q x + ?F x * (X w - x))" using prob_space by (simp add: X) also have "\ \ expectation (\w. q (X w))" using \x \ I\ \open I\ X(2) apply (intro integral_mono_AE Bochner_Integration.integrable_add Bochner_Integration.integrable_mult_right Bochner_Integration.integrable_diff integrable_const X q) apply (elim eventually_mono) apply (intro convex_le_Inf_differential) apply (auto simp: interior_open q) done finally show "k \ expectation (\w. q (X w))" using x by auto qed finally show "q (expectation X) \ expectation (\x. q (X x))" . qed subsection \Introduce binder for probability\ syntax "_prob" :: "pttrn \ logic \ logic \ logic" (\('\

'((/_ in _./ _)'))\) translations "\

(x in M. P)" => "CONST measure M {x \ CONST space M. P}" print_translation \ let fun to_pattern (Const (\<^const_syntax>\Pair\, _) $ l $ r) = Syntax.const \<^const_syntax>\Pair\ :: to_pattern l @ to_pattern r | to_pattern (t as (Const (\<^syntax_const>\_bound\, _)) $ _) = [t] fun mk_pattern ((t, n) :: xs) = mk_patterns n xs |>> curry list_comb t and mk_patterns 0 xs = ([], xs) | mk_patterns n xs = let val (t, xs') = mk_pattern xs val (ts, xs'') = mk_patterns (n - 1) xs' in (t :: ts, xs'') end fun unnest_tuples (Const (\<^syntax_const>\_pattern\, _) $ t1 $ (t as (Const (\<^syntax_const>\_pattern\, _) $ _ $ _))) = let val (_ $ t2 $ t3) = unnest_tuples t in Syntax.const \<^syntax_const>\_pattern\ $ unnest_tuples t1 $ (Syntax.const \<^syntax_const>\_patterns\ $ t2 $ t3) end | unnest_tuples pat = pat fun tr' [sig_alg, Const (\<^const_syntax>\Collect\, _) $ t] = let val bound_dummyT = Const (\<^syntax_const>\_bound\, dummyT) fun go pattern elem (Const (\<^const_syntax>\conj\, _) $ (Const (\<^const_syntax>\Set.member\, _) $ elem' $ (Const (\<^const_syntax>\space\, _) $ sig_alg')) $ u) = let val _ = if sig_alg aconv sig_alg' andalso to_pattern elem' = rev elem then () else raise Match; val (pat, rest) = mk_pattern (rev pattern); val _ = case rest of [] => () | _ => raise Match in Syntax.const \<^syntax_const>\_prob\ $ unnest_tuples pat $ sig_alg $ u end | go pattern elem (Abs abs) = let val (x as (_ $ tx), t) = Syntax_Trans.atomic_abs_tr' abs in go ((x, 0) :: pattern) (bound_dummyT $ tx :: elem) t end | go pattern elem (Const (\<^const_syntax>\case_prod\, _) $ t) = go ((Syntax.const \<^syntax_const>\_pattern\, 2) :: pattern) (Syntax.const \<^const_syntax>\Pair\ :: elem) t in go [] [] t end in [(\<^const_syntax>\Sigma_Algebra.measure\, K tr')] end \ definition "cond_prob M P Q = \

(\ in M. P \ \ Q \) / \

(\ in M. Q \)" syntax "_conditional_prob" :: "pttrn \ logic \ logic \ logic \ logic" (\('\

'(_ in _. _ \/ _'))\) translations "\

(x in M. P \ Q)" => "CONST cond_prob M (\x. P) (\x. Q)" lemma (in prob_space) AE_E_prob: assumes ae: "AE x in M. P x" obtains S where "S \ {x \ space M. P x}" "S \ events" "prob S = 1" proof - - from ae[THEN AE_E] guess N . + from ae[THEN AE_E] obtain N + where "{x \ space M. \ P x} \ N" "emeasure M N = 0" "N \ events" by auto then show thesis by (intro that[of "space M - N"]) (auto simp: prob_compl prob_space emeasure_eq_measure measure_nonneg) qed lemma (in prob_space) prob_neg: "{x\space M. P x} \ events \ \

(x in M. \ P x) = 1 - \

(x in M. P x)" by (auto intro!: arg_cong[where f=prob] simp add: prob_compl[symmetric]) lemma (in prob_space) prob_eq_AE: "(AE x in M. P x \ Q x) \ {x\space M. P x} \ events \ {x\space M. Q x} \ events \ \

(x in M. P x) = \

(x in M. Q x)" by (rule finite_measure_eq_AE) auto lemma (in prob_space) prob_eq_0_AE: assumes not: "AE x in M. \ P x" shows "\

(x in M. P x) = 0" proof cases assume "{x\space M. P x} \ events" with not have "\

(x in M. P x) = \

(x in M. False)" by (intro prob_eq_AE) auto then show ?thesis by simp qed (simp add: measure_notin_sets) lemma (in prob_space) prob_Collect_eq_0: "{x \ space M. P x} \ sets M \ \

(x in M. P x) = 0 \ (AE x in M. \ P x)" using AE_iff_measurable[OF _ refl, of M "\x. \ P x"] by (simp add: emeasure_eq_measure measure_nonneg) lemma (in prob_space) prob_Collect_eq_1: "{x \ space M. P x} \ sets M \ \

(x in M. P x) = 1 \ (AE x in M. P x)" using AE_in_set_eq_1[of "{x\space M. P x}"] by simp lemma (in prob_space) prob_eq_0: "A \ sets M \ prob A = 0 \ (AE x in M. x \ A)" using AE_iff_measurable[OF _ refl, of M "\x. x \ A"] by (auto simp add: emeasure_eq_measure Int_def[symmetric] measure_nonneg) lemma (in prob_space) prob_eq_1: "A \ sets M \ prob A = 1 \ (AE x in M. x \ A)" using AE_in_set_eq_1[of A] by simp lemma (in prob_space) prob_sums: assumes P: "\n. {x\space M. P n x} \ events" assumes Q: "{x\space M. Q x} \ events" assumes ae: "AE x in M. (\n. P n x \ Q x) \ (Q x \ (\!n. P n x))" shows "(\n. \

(x in M. P n x)) sums \

(x in M. Q x)" proof - - from ae[THEN AE_E_prob] guess S . note S = this + from ae[THEN AE_E_prob] obtain S + where S: + "S \ {x \ space M. (\n. P n x \ Q x) \ (Q x \ (\!n. P n x))}" + "S \ events" + "prob S = 1" + by auto then have disj: "disjoint_family (\n. {x\space M. P n x} \ S)" by (auto simp: disjoint_family_on_def) from S have ae_S: "AE x in M. x \ {x\space M. Q x} \ x \ (\n. {x\space M. P n x} \ S)" "\n. AE x in M. x \ {x\space M. P n x} \ x \ {x\space M. P n x} \ S" using ae by (auto dest!: AE_prob_1) from ae_S have *: "\

(x in M. Q x) = prob (\n. {x\space M. P n x} \ S)" using P Q S by (intro finite_measure_eq_AE) auto from ae_S have **: "\n. \

(x in M. P n x) = prob ({x\space M. P n x} \ S)" using P Q S by (intro finite_measure_eq_AE) auto show ?thesis unfolding * ** using S P disj by (intro finite_measure_UNION) auto qed lemma (in prob_space) prob_sum: assumes [simp, intro]: "finite I" assumes P: "\n. n \ I \ {x\space M. P n x} \ events" assumes Q: "{x\space M. Q x} \ events" assumes ae: "AE x in M. (\n\I. P n x \ Q x) \ (Q x \ (\!n\I. P n x))" shows "\

(x in M. Q x) = (\n\I. \

(x in M. P n x))" proof - - from ae[THEN AE_E_prob] guess S . note S = this + from ae[THEN AE_E_prob] obtain S + where S: + "S \ {x \ space M. (\n\I. P n x \ Q x) \ (Q x \ (\!n. n \ I \ P n x))}" + "S \ events" + "prob S = 1" + by auto then have disj: "disjoint_family_on (\n. {x\space M. P n x} \ S) I" by (auto simp: disjoint_family_on_def) from S have ae_S: "AE x in M. x \ {x\space M. Q x} \ x \ (\n\I. {x\space M. P n x} \ S)" "\n. n \ I \ AE x in M. x \ {x\space M. P n x} \ x \ {x\space M. P n x} \ S" using ae by (auto dest!: AE_prob_1) from ae_S have *: "\

(x in M. Q x) = prob (\n\I. {x\space M. P n x} \ S)" using P Q S by (intro finite_measure_eq_AE) (auto intro!: sets.Int) from ae_S have **: "\n. n \ I \ \

(x in M. P n x) = prob ({x\space M. P n x} \ S)" using P Q S by (intro finite_measure_eq_AE) auto show ?thesis using S P disj by (auto simp add: * ** simp del: UN_simps intro!: finite_measure_finite_Union) qed lemma (in prob_space) prob_EX_countable: assumes sets: "\i. i \ I \ {x\space M. P i x} \ sets M" and I: "countable I" assumes disj: "AE x in M. \i\I. \j\I. P i x \ P j x \ i = j" shows "\

(x in M. \i\I. P i x) = (\\<^sup>+i. \

(x in M. P i x) \count_space I)" proof - let ?N= "\x. \!i\I. P i x" have "ennreal (\

(x in M. \i\I. P i x)) = \

(x in M. (\i\I. P i x \ ?N x))" unfolding ennreal_inj[OF measure_nonneg measure_nonneg] proof (rule prob_eq_AE) show "AE x in M. (\i\I. P i x) = (\i\I. P i x \ ?N x)" using disj by eventually_elim blast qed (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets)+ also have "\

(x in M. (\i\I. P i x \ ?N x)) = emeasure M (\i\I. {x\space M. P i x \ ?N x})" unfolding emeasure_eq_measure by (auto intro!: arg_cong[where f=prob] simp: measure_nonneg) also have "\ = (\\<^sup>+i. emeasure M {x\space M. P i x \ ?N x} \count_space I)" by (rule emeasure_UN_countable) (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets simp: disjoint_family_on_def) also have "\ = (\\<^sup>+i. \

(x in M. P i x) \count_space I)" unfolding emeasure_eq_measure using disj by (intro nn_integral_cong ennreal_inj[THEN iffD2] prob_eq_AE) (auto intro!: sets.sets_Collect_countable_Ex' sets.sets_Collect_conj sets.sets_Collect_countable_Ex1' I sets measure_nonneg)+ finally show ?thesis . qed lemma (in prob_space) cond_prob_eq_AE: assumes P: "AE x in M. Q x \ P x \ P' x" "{x\space M. P x} \ events" "{x\space M. P' x} \ events" assumes Q: "AE x in M. Q x \ Q' x" "{x\space M. Q x} \ events" "{x\space M. Q' x} \ events" shows "cond_prob M P Q = cond_prob M P' Q'" using P Q by (auto simp: cond_prob_def intro!: arg_cong2[where f="(/)"] prob_eq_AE sets.sets_Collect_conj) lemma (in prob_space) joint_distribution_Times_le_fst: "random_variable MX X \ random_variable MY Y \ A \ sets MX \ B \ sets MY \ emeasure (distr M (MX \\<^sub>M MY) (\x. (X x, Y x))) (A \ B) \ emeasure (distr M MX X) A" by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) lemma (in prob_space) joint_distribution_Times_le_snd: "random_variable MX X \ random_variable MY Y \ A \ sets MX \ B \ sets MY \ emeasure (distr M (MX \\<^sub>M MY) (\x. (X x, Y x))) (A \ B) \ emeasure (distr M MY Y) B" by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets) lemma (in prob_space) variance_eq: fixes X :: "'a \ real" assumes [simp]: "integrable M X" assumes [simp]: "integrable M (\x. (X x)\<^sup>2)" shows "variance X = expectation (\x. (X x)\<^sup>2) - (expectation X)\<^sup>2" by (simp add: field_simps prob_space power2_diff power2_eq_square[symmetric]) lemma (in prob_space) variance_positive: "0 \ variance (X::'a \ real)" by (intro integral_nonneg_AE) (auto intro!: integral_nonneg_AE) lemma (in prob_space) variance_mean_zero: "expectation X = 0 \ variance X = expectation (\x. (X x)^2)" by simp theorem%important (in prob_space) Chebyshev_inequality: assumes [measurable]: "random_variable borel f" assumes "integrable M (\x. f x ^ 2)" defines "\ \ expectation f" assumes "a > 0" shows "prob {x\space M. \f x - \\ \ a} \ variance f / a\<^sup>2" unfolding \_def proof (rule second_moment_method) have integrable: "integrable M f" using assms by (blast dest: square_integrable_imp_integrable) show "integrable M (\x. (f x - expectation f)\<^sup>2)" using assms integrable unfolding power2_eq_square ring_distribs by (intro Bochner_Integration.integrable_diff) auto qed (use assms in auto) locale pair_prob_space = pair_sigma_finite M1 M2 + M1: prob_space M1 + M2: prob_space M2 for M1 M2 sublocale pair_prob_space \ P?: prob_space "M1 \\<^sub>M M2" proof show "emeasure (M1 \\<^sub>M M2) (space (M1 \\<^sub>M M2)) = 1" by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure) qed locale product_prob_space = product_sigma_finite M for M :: "'i \ 'a measure" + fixes I :: "'i set" assumes prob_space: "\i. prob_space (M i)" sublocale product_prob_space \ M?: prob_space "M i" for i by (rule prob_space) locale finite_product_prob_space = finite_product_sigma_finite M I + product_prob_space M I for M I sublocale finite_product_prob_space \ prob_space "\\<^sub>M i\I. M i" proof show "emeasure (\\<^sub>M i\I. M i) (space (\\<^sub>M i\I. M i)) = 1" by (simp add: measure_times M.emeasure_space_1 prod.neutral_const space_PiM) qed lemma (in finite_product_prob_space) prob_times: assumes X: "\i. i \ I \ X i \ sets (M i)" shows "prob (\\<^sub>E i\I. X i) = (\i\I. M.prob i (X i))" proof - have "ennreal (measure (\\<^sub>M i\I. M i) (\\<^sub>E i\I. X i)) = emeasure (\\<^sub>M i\I. M i) (\\<^sub>E i\I. X i)" using X by (simp add: emeasure_eq_measure) also have "\ = (\i\I. emeasure (M i) (X i))" using measure_times X by simp also have "\ = ennreal (\i\I. measure (M i) (X i))" using X by (simp add: M.emeasure_eq_measure prod_ennreal measure_nonneg) finally show ?thesis by (simp add: measure_nonneg prod_nonneg) qed lemma product_prob_spaceI: assumes "\i. prob_space (M i)" shows "product_prob_space M" unfolding product_prob_space_def product_prob_space_axioms_def product_sigma_finite_def proof safe fix i interpret prob_space "M i" by (rule assms) show "sigma_finite_measure (M i)" "prob_space (M i)" by unfold_locales qed subsection \Distributions\ definition distributed :: "'a measure \ 'b measure \ ('a \ 'b) \ ('b \ ennreal) \ bool" where "distributed M N X f \ distr M N X = density N f \ f \ borel_measurable N \ X \ measurable M N" lemma assumes "distributed M N X f" shows distributed_distr_eq_density: "distr M N X = density N f" and distributed_measurable: "X \ measurable M N" and distributed_borel_measurable: "f \ borel_measurable N" using assms by (simp_all add: distributed_def) lemma assumes D: "distributed M N X f" shows distributed_measurable'[measurable_dest]: "g \ measurable L M \ (\x. X (g x)) \ measurable L N" and distributed_borel_measurable'[measurable_dest]: "h \ measurable L N \ (\x. f (h x)) \ borel_measurable L" using distributed_measurable[OF D] distributed_borel_measurable[OF D] by simp_all lemma distributed_real_measurable: "(\x. x \ space N \ 0 \ f x) \ distributed M N X (\x. ennreal (f x)) \ f \ borel_measurable N" by (simp_all add: distributed_def) lemma distributed_real_measurable': "(\x. x \ space N \ 0 \ f x) \ distributed M N X (\x. ennreal (f x)) \ h \ measurable L N \ (\x. f (h x)) \ borel_measurable L" using distributed_real_measurable[measurable] by simp lemma joint_distributed_measurable1: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) f \ h1 \ measurable N M \ (\x. X (h1 x)) \ measurable N S" by simp lemma joint_distributed_measurable2: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) f \ h2 \ measurable N M \ (\x. Y (h2 x)) \ measurable N T" by simp lemma distributed_count_space: assumes X: "distributed M (count_space A) X P" and a: "a \ A" and A: "finite A" shows "P a = emeasure M (X -` {a} \ space M)" proof - have "emeasure M (X -` {a} \ space M) = emeasure (distr M (count_space A) X) {a}" using X a A by (simp add: emeasure_distr) also have "\ = emeasure (density (count_space A) P) {a}" using X by (simp add: distributed_distr_eq_density) also have "\ = (\\<^sup>+x. P a * indicator {a} x \count_space A)" using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: nn_integral_cong) also have "\ = P a" using X a by (subst nn_integral_cmult_indicator) (auto simp: distributed_def one_ennreal_def[symmetric] AE_count_space) finally show ?thesis .. qed lemma distributed_cong_density: "(AE x in N. f x = g x) \ g \ borel_measurable N \ f \ borel_measurable N \ distributed M N X f \ distributed M N X g" by (auto simp: distributed_def intro!: density_cong) lemma (in prob_space) distributed_imp_emeasure_nonzero: assumes X: "distributed M MX X Px" shows "emeasure MX {x \ space MX. Px x \ 0} \ 0" proof note Px = distributed_borel_measurable[OF X] interpret X: prob_space "distr M MX X" using distributed_measurable[OF X] by (rule prob_space_distr) assume "emeasure MX {x \ space MX. Px x \ 0} = 0" with Px have "AE x in MX. Px x = 0" by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ennreal_iff) moreover from X.emeasure_space_1 have "(\\<^sup>+x. Px x \MX) = 1" unfolding distributed_distr_eq_density[OF X] using Px by (subst (asm) emeasure_density) (auto simp: borel_measurable_ennreal_iff intro!: integral_cong cong: nn_integral_cong) ultimately show False by (simp add: nn_integral_cong_AE) qed lemma subdensity: assumes T: "T \ measurable P Q" assumes f: "distributed M P X f" assumes g: "distributed M Q Y g" assumes Y: "Y = T \ X" shows "AE x in P. g (T x) = 0 \ f x = 0" proof - have "{x\space Q. g x = 0} \ null_sets (distr M Q (T \ X))" using g Y by (auto simp: null_sets_density_iff distributed_def) also have "distr M Q (T \ X) = distr (distr M P X) Q T" using T f[THEN distributed_measurable] by (rule distr_distr[symmetric]) finally have "T -` {x\space Q. g x = 0} \ space P \ null_sets (distr M P X)" using T by (subst (asm) null_sets_distr_iff) auto also have "T -` {x\space Q. g x = 0} \ space P = {x\space P. g (T x) = 0}" using T by (auto dest: measurable_space) finally show ?thesis using f g by (auto simp add: null_sets_density_iff distributed_def) qed lemma subdensity_real: fixes g :: "'a \ real" and f :: "'b \ real" assumes T: "T \ measurable P Q" assumes f: "distributed M P X f" assumes g: "distributed M Q Y g" assumes Y: "Y = T \ X" shows "(AE x in P. 0 \ g (T x)) \ (AE x in P. 0 \ f x) \ AE x in P. g (T x) = 0 \ f x = 0" using subdensity[OF T, of M X "\x. ennreal (f x)" Y "\x. ennreal (g x)"] assms by auto lemma distributed_emeasure: "distributed M N X f \ A \ sets N \ emeasure M (X -` A \ space M) = (\\<^sup>+x. f x * indicator A x \N)" by (auto simp: distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr) lemma distributed_nn_integral: "distributed M N X f \ g \ borel_measurable N \ (\\<^sup>+x. f x * g x \N) = (\\<^sup>+x. g (X x) \M)" by (auto simp: distributed_distr_eq_density[symmetric] nn_integral_density[symmetric] nn_integral_distr) lemma distributed_integral: "distributed M N X f \ g \ borel_measurable N \ (\x. x \ space N \ 0 \ f x) \ (\x. f x * g x \N) = (\x. g (X x) \M)" supply distributed_real_measurable[measurable] by (auto simp: distributed_distr_eq_density[symmetric] integral_real_density[symmetric] integral_distr) lemma distributed_transform_integral: assumes Px: "distributed M N X Px" "\x. x \ space N \ 0 \ Px x" assumes "distributed M P Y Py" "\x. x \ space P \ 0 \ Py x" assumes Y: "Y = T \ X" and T: "T \ measurable N P" and f: "f \ borel_measurable P" shows "(\x. Py x * f x \P) = (\x. Px x * f (T x) \N)" proof - have "(\x. Py x * f x \P) = (\x. f (Y x) \M)" by (rule distributed_integral) fact+ also have "\ = (\x. f (T (X x)) \M)" using Y by simp also have "\ = (\x. Px x * f (T x) \N)" using measurable_comp[OF T f] Px by (intro distributed_integral[symmetric]) (auto simp: comp_def) finally show ?thesis . qed lemma (in prob_space) distributed_unique: assumes Px: "distributed M S X Px" assumes Py: "distributed M S X Py" shows "AE x in S. Px x = Py x" proof - interpret X: prob_space "distr M S X" using Px by (intro prob_space_distr) simp have "sigma_finite_measure (distr M S X)" .. with sigma_finite_density_unique[of Px S Py ] Px Py show ?thesis by (auto simp: distributed_def) qed lemma (in prob_space) distributed_jointI: assumes "sigma_finite_measure S" "sigma_finite_measure T" assumes X[measurable]: "X \ measurable M S" and Y[measurable]: "Y \ measurable M T" assumes [measurable]: "f \ borel_measurable (S \\<^sub>M T)" and f: "AE x in S \\<^sub>M T. 0 \ f x" assumes eq: "\A B. A \ sets S \ B \ sets T \ emeasure M {x \ space M. X x \ A \ Y x \ B} = (\\<^sup>+x. (\\<^sup>+y. f (x, y) * indicator B y \T) * indicator A x \S)" shows "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) f" unfolding distributed_def proof safe interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. - from ST.sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('b \ 'c) set" .. note F = this + from ST.sigma_finite_up_in_pair_measure_generator + obtain F :: "nat \ ('b \ 'c) set" + where F: "range F \ {A \ B |A B. A \ sets S \ B \ sets T} \ incseq F \ + \ (range F) = space S \ space T \ (\i. emeasure (S \\<^sub>M T) (F i) \ \)" .. let ?E = "{a \ b |a b. a \ sets S \ b \ sets T}" let ?P = "S \\<^sub>M T" show "distr M ?P (\x. (X x, Y x)) = density ?P f" (is "?L = ?R") proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of S T]]) show "?E \ Pow (space ?P)" using sets.space_closed[of S] sets.space_closed[of T] by (auto simp: space_pair_measure) show "sets ?L = sigma_sets (space ?P) ?E" by (simp add: sets_pair_measure space_pair_measure) then show "sets ?R = sigma_sets (space ?P) ?E" by simp next interpret L: prob_space ?L by (rule prob_space_distr) (auto intro!: measurable_Pair) show "range F \ ?E" "(\i. F i) = space ?P" "\i. emeasure ?L (F i) \ \" using F by (auto simp: space_pair_measure) next fix E assume "E \ ?E" then obtain A B where E[simp]: "E = A \ B" and A[measurable]: "A \ sets S" and B[measurable]: "B \ sets T" by auto have "emeasure ?L E = emeasure M {x \ space M. X x \ A \ Y x \ B}" by (auto intro!: arg_cong[where f="emeasure M"] simp add: emeasure_distr measurable_Pair) also have "\ = (\\<^sup>+x. (\\<^sup>+y. (f (x, y) * indicator B y) * indicator A x \T) \S)" using f by (auto simp add: eq nn_integral_multc intro!: nn_integral_cong) also have "\ = emeasure ?R E" by (auto simp add: emeasure_density T.nn_integral_fst[symmetric] intro!: nn_integral_cong split: split_indicator) finally show "emeasure ?L E = emeasure ?R E" . qed qed (auto simp: f) lemma (in prob_space) distributed_swap: assumes "sigma_finite_measure S" "sigma_finite_measure T" assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" shows "distributed M (T \\<^sub>M S) (\x. (Y x, X x)) (\(x, y). Pxy (y, x))" proof - interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. interpret TS: pair_sigma_finite T S .. note Pxy[measurable] show ?thesis apply (subst TS.distr_pair_swap) unfolding distributed_def proof safe let ?D = "distr (S \\<^sub>M T) (T \\<^sub>M S) (\(x, y). (y, x))" show 1: "(\(x, y). Pxy (y, x)) \ borel_measurable ?D" by auto show 2: "random_variable (distr (S \\<^sub>M T) (T \\<^sub>M S) (\(x, y). (y, x))) (\x. (Y x, X x))" using Pxy by auto { fix A assume A: "A \ sets (T \\<^sub>M S)" let ?B = "(\(x, y). (y, x)) -` A \ space (S \\<^sub>M T)" from sets.sets_into_space[OF A] have "emeasure M ((\x. (Y x, X x)) -` A \ space M) = emeasure M ((\x. (X x, Y x)) -` ?B \ space M)" by (auto intro!: arg_cong2[where f=emeasure] simp: space_pair_measure) also have "\ = (\\<^sup>+ x. Pxy x * indicator ?B x \(S \\<^sub>M T))" using Pxy A by (intro distributed_emeasure) auto finally have "emeasure M ((\x. (Y x, X x)) -` A \ space M) = (\\<^sup>+ x. Pxy x * indicator A (snd x, fst x) \(S \\<^sub>M T))" by (auto intro!: nn_integral_cong split: split_indicator) } note * = this show "distr M ?D (\x. (Y x, X x)) = density ?D (\(x, y). Pxy (y, x))" apply (intro measure_eqI) apply (simp_all add: emeasure_distr[OF 2] emeasure_density[OF 1]) apply (subst nn_integral_distr) apply (auto intro!: * simp: comp_def split_beta) done qed qed lemma (in prob_space) distr_marginal1: assumes "sigma_finite_measure S" "sigma_finite_measure T" assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" defines "Px \ \x. (\\<^sup>+z. Pxy (x, z) \T)" shows "distributed M S X Px" unfolding distributed_def proof safe interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. note Pxy[measurable] show X: "X \ measurable M S" by simp show borel: "Px \ borel_measurable S" by (auto intro!: T.nn_integral_fst simp: Px_def) interpret Pxy: prob_space "distr M (S \\<^sub>M T) (\x. (X x, Y x))" by (intro prob_space_distr) simp show "distr M S X = density S Px" proof (rule measure_eqI) fix A assume A: "A \ sets (distr M S X)" with X measurable_space[of Y M T] have "emeasure (distr M S X) A = emeasure (distr M (S \\<^sub>M T) (\x. (X x, Y x))) (A \ space T)" by (auto simp add: emeasure_distr intro!: arg_cong[where f="emeasure M"]) also have "\ = emeasure (density (S \\<^sub>M T) Pxy) (A \ space T)" using Pxy by (simp add: distributed_def) also have "\ = \\<^sup>+ x. \\<^sup>+ y. Pxy (x, y) * indicator (A \ space T) (x, y) \T \S" using A borel Pxy by (simp add: emeasure_density T.nn_integral_fst[symmetric]) also have "\ = \\<^sup>+ x. Px x * indicator A x \S" proof (rule nn_integral_cong) fix x assume "x \ space S" moreover have eq: "\y. y \ space T \ indicator (A \ space T) (x, y) = indicator A x" by (auto simp: indicator_def) ultimately have "(\\<^sup>+ y. Pxy (x, y) * indicator (A \ space T) (x, y) \T) = (\\<^sup>+ y. Pxy (x, y) \T) * indicator A x" by (simp add: eq nn_integral_multc cong: nn_integral_cong) also have "(\\<^sup>+ y. Pxy (x, y) \T) = Px x" by (simp add: Px_def) finally show "(\\<^sup>+ y. Pxy (x, y) * indicator (A \ space T) (x, y) \T) = Px x * indicator A x" . qed finally show "emeasure (distr M S X) A = emeasure (density S Px) A" using A borel Pxy by (simp add: emeasure_density) qed simp qed lemma (in prob_space) distr_marginal2: assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" shows "distributed M T Y (\y. (\\<^sup>+x. Pxy (x, y) \S))" using distr_marginal1[OF T S distributed_swap[OF S T]] Pxy by simp lemma (in prob_space) distributed_marginal_eq_joint1: assumes T: "sigma_finite_measure T" assumes S: "sigma_finite_measure S" assumes Px: "distributed M S X Px" assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" shows "AE x in S. Px x = (\\<^sup>+y. Pxy (x, y) \T)" using Px distr_marginal1[OF S T Pxy] by (rule distributed_unique) lemma (in prob_space) distributed_marginal_eq_joint2: assumes T: "sigma_finite_measure T" assumes S: "sigma_finite_measure S" assumes Py: "distributed M T Y Py" assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" shows "AE y in T. Py y = (\\<^sup>+x. Pxy (x, y) \S)" using Py distr_marginal2[OF S T Pxy] by (rule distributed_unique) lemma (in prob_space) distributed_joint_indep': assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes X[measurable]: "distributed M S X Px" and Y[measurable]: "distributed M T Y Py" assumes indep: "distr M S X \\<^sub>M distr M T Y = distr M (S \\<^sub>M T) (\x. (X x, Y x))" shows "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) (\(x, y). Px x * Py y)" unfolding distributed_def proof safe interpret S: sigma_finite_measure S by fact interpret T: sigma_finite_measure T by fact interpret ST: pair_sigma_finite S T .. interpret X: prob_space "density S Px" unfolding distributed_distr_eq_density[OF X, symmetric] by (rule prob_space_distr) simp have sf_X: "sigma_finite_measure (density S Px)" .. interpret Y: prob_space "density T Py" unfolding distributed_distr_eq_density[OF Y, symmetric] by (rule prob_space_distr) simp have sf_Y: "sigma_finite_measure (density T Py)" .. show "distr M (S \\<^sub>M T) (\x. (X x, Y x)) = density (S \\<^sub>M T) (\(x, y). Px x * Py y)" unfolding indep[symmetric] distributed_distr_eq_density[OF X] distributed_distr_eq_density[OF Y] using distributed_borel_measurable[OF X] using distributed_borel_measurable[OF Y] by (rule pair_measure_density[OF _ _ T sf_Y]) show "random_variable (S \\<^sub>M T) (\x. (X x, Y x))" by auto show Pxy: "(\(x, y). Px x * Py y) \ borel_measurable (S \\<^sub>M T)" by auto qed lemma distributed_integrable: "distributed M N X f \ g \ borel_measurable N \ (\x. x \ space N \ 0 \ f x) \ integrable N (\x. f x * g x) \ integrable M (\x. g (X x))" supply distributed_real_measurable[measurable] by (auto simp: distributed_distr_eq_density[symmetric] integrable_real_density[symmetric] integrable_distr_eq) lemma distributed_transform_integrable: assumes Px: "distributed M N X Px" "\x. x \ space N \ 0 \ Px x" assumes "distributed M P Y Py" "\x. x \ space P \ 0 \ Py x" assumes Y: "Y = (\x. T (X x))" and T: "T \ measurable N P" and f: "f \ borel_measurable P" shows "integrable P (\x. Py x * f x) \ integrable N (\x. Px x * f (T x))" proof - have "integrable P (\x. Py x * f x) \ integrable M (\x. f (Y x))" by (rule distributed_integrable) fact+ also have "\ \ integrable M (\x. f (T (X x)))" using Y by simp also have "\ \ integrable N (\x. Px x * f (T x))" using measurable_comp[OF T f] Px by (intro distributed_integrable[symmetric]) (auto simp: comp_def) finally show ?thesis . qed lemma distributed_integrable_var: fixes X :: "'a \ real" shows "distributed M lborel X (\x. ennreal (f x)) \ (\x. 0 \ f x) \ integrable lborel (\x. f x * x) \ integrable M X" using distributed_integrable[of M lborel X f "\x. x"] by simp lemma (in prob_space) distributed_variance: fixes f::"real \ real" assumes D: "distributed M lborel X f" and [simp]: "\x. 0 \ f x" shows "variance X = (\x. x\<^sup>2 * f (x + expectation X) \lborel)" proof (subst distributed_integral[OF D, symmetric]) show "(\ x. f x * (x - expectation X)\<^sup>2 \lborel) = (\ x. x\<^sup>2 * f (x + expectation X) \lborel)" by (subst lborel_integral_real_affine[where c=1 and t="expectation X"]) (auto simp: ac_simps) qed simp_all lemma (in prob_space) variance_affine: fixes f::"real \ real" assumes [arith]: "b \ 0" assumes D[intro]: "distributed M lborel X f" assumes [simp]: "prob_space (density lborel f)" assumes I[simp]: "integrable M X" assumes I2[simp]: "integrable M (\x. (X x)\<^sup>2)" shows "variance (\x. a + b * X x) = b\<^sup>2 * variance X" by (subst variance_eq) (auto simp: power2_sum power_mult_distrib prob_space variance_eq right_diff_distrib) definition "simple_distributed M X f \ (\x. 0 \ f x) \ distributed M (count_space (X`space M)) X (\x. ennreal (f x)) \ finite (X`space M)" lemma simple_distributed_nonneg[dest]: "simple_distributed M X f \ 0 \ f x" by (auto simp: simple_distributed_def) lemma simple_distributed: "simple_distributed M X Px \ distributed M (count_space (X`space M)) X Px" unfolding simple_distributed_def by auto lemma simple_distributed_finite[dest]: "simple_distributed M X P \ finite (X`space M)" by (simp add: simple_distributed_def) lemma (in prob_space) distributed_simple_function_superset: assumes X: "simple_function M X" "\x. x \ X ` space M \ P x = measure M (X -` {x} \ space M)" assumes A: "X`space M \ A" "finite A" defines "S \ count_space A" and "P' \ (\x. if x \ X`space M then P x else 0)" shows "distributed M S X P'" unfolding distributed_def proof safe show "(\x. ennreal (P' x)) \ borel_measurable S" unfolding S_def by simp show "distr M S X = density S P'" proof (rule measure_eqI_finite) show "sets (distr M S X) = Pow A" "sets (density S P') = Pow A" using A unfolding S_def by auto show "finite A" by fact fix a assume a: "a \ A" then have "a \ X`space M \ X -` {a} \ space M = {}" by auto with A a X have "emeasure (distr M S X) {a} = P' a" by (subst emeasure_distr) (auto simp add: S_def P'_def simple_functionD emeasure_eq_measure measurable_count_space_eq2 intro!: arg_cong[where f=prob]) also have "\ = (\\<^sup>+x. ennreal (P' a) * indicator {a} x \S)" using A X a by (subst nn_integral_cmult_indicator) (auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg) also have "\ = (\\<^sup>+x. ennreal (P' x) * indicator {a} x \S)" by (auto simp: indicator_def intro!: nn_integral_cong) also have "\ = emeasure (density S P') {a}" using a A by (intro emeasure_density[symmetric]) (auto simp: S_def) finally show "emeasure (distr M S X) {a} = emeasure (density S P') {a}" . qed show "random_variable S X" using X(1) A by (auto simp: measurable_def simple_functionD S_def) qed lemma (in prob_space) simple_distributedI: assumes X: "simple_function M X" "\x. 0 \ P x" "\x. x \ X ` space M \ P x = measure M (X -` {x} \ space M)" shows "simple_distributed M X P" unfolding simple_distributed_def proof (safe intro!: X) have "distributed M (count_space (X ` space M)) X (\x. ennreal (if x \ X`space M then P x else 0))" (is "?A") using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X(1,3)]) auto also have "?A \ distributed M (count_space (X ` space M)) X (\x. ennreal (P x))" by (rule distributed_cong_density) auto finally show "\" . qed (rule simple_functionD[OF X(1)]) lemma simple_distributed_joint_finite: assumes X: "simple_distributed M (\x. (X x, Y x)) Px" shows "finite (X ` space M)" "finite (Y ` space M)" proof - have "finite ((\x. (X x, Y x)) ` space M)" using X by (auto simp: simple_distributed_def simple_functionD) then have "finite (fst ` (\x. (X x, Y x)) ` space M)" "finite (snd ` (\x. (X x, Y x)) ` space M)" by auto then show fin: "finite (X ` space M)" "finite (Y ` space M)" by (auto simp: image_image) qed lemma simple_distributed_joint2_finite: assumes X: "simple_distributed M (\x. (X x, Y x, Z x)) Px" shows "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" proof - have "finite ((\x. (X x, Y x, Z x)) ` space M)" using X by (auto simp: simple_distributed_def simple_functionD) then have "finite (fst ` (\x. (X x, Y x, Z x)) ` space M)" "finite ((fst \ snd) ` (\x. (X x, Y x, Z x)) ` space M)" "finite ((snd \ snd) ` (\x. (X x, Y x, Z x)) ` space M)" by auto then show fin: "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)" by (auto simp: image_image) qed lemma simple_distributed_simple_function: "simple_distributed M X Px \ simple_function M X" unfolding simple_distributed_def distributed_def by (auto simp: simple_function_def measurable_count_space_eq2) lemma simple_distributed_measure: "simple_distributed M X P \ a \ X`space M \ P a = measure M (X -` {a} \ space M)" using distributed_count_space[of M "X`space M" X P a, symmetric] by (auto simp: simple_distributed_def measure_def) lemma (in prob_space) simple_distributed_joint: assumes X: "simple_distributed M (\x. (X x, Y x)) Px" defines "S \ count_space (X`space M) \\<^sub>M count_space (Y`space M)" defines "P \ (\x. if x \ (\x. (X x, Y x))`space M then Px x else 0)" shows "distributed M S (\x. (X x, Y x)) P" proof - from simple_distributed_joint_finite[OF X, simp] have S_eq: "S = count_space (X`space M \ Y`space M)" by (simp add: S_def pair_measure_count_space) show ?thesis unfolding S_eq P_def proof (rule distributed_simple_function_superset) show "simple_function M (\x. (X x, Y x))" using X by (rule simple_distributed_simple_function) fix x assume "x \ (\x. (X x, Y x)) ` space M" from simple_distributed_measure[OF X this] show "Px x = prob ((\x. (X x, Y x)) -` {x} \ space M)" . qed auto qed lemma (in prob_space) simple_distributed_joint2: assumes X: "simple_distributed M (\x. (X x, Y x, Z x)) Px" defines "S \ count_space (X`space M) \\<^sub>M count_space (Y`space M) \\<^sub>M count_space (Z`space M)" defines "P \ (\x. if x \ (\x. (X x, Y x, Z x))`space M then Px x else 0)" shows "distributed M S (\x. (X x, Y x, Z x)) P" proof - from simple_distributed_joint2_finite[OF X, simp] have S_eq: "S = count_space (X`space M \ Y`space M \ Z`space M)" by (simp add: S_def pair_measure_count_space) show ?thesis unfolding S_eq P_def proof (rule distributed_simple_function_superset) show "simple_function M (\x. (X x, Y x, Z x))" using X by (rule simple_distributed_simple_function) fix x assume "x \ (\x. (X x, Y x, Z x)) ` space M" from simple_distributed_measure[OF X this] show "Px x = prob ((\x. (X x, Y x, Z x)) -` {x} \ space M)" . qed auto qed lemma (in prob_space) simple_distributed_sum_space: assumes X: "simple_distributed M X f" shows "sum f (X`space M) = 1" proof - from X have "sum f (X`space M) = prob (\i\X`space M. X -` {i} \ space M)" by (subst finite_measure_finite_Union) (auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD intro!: sum.cong arg_cong[where f="prob"]) also have "\ = prob (space M)" by (auto intro!: arg_cong[where f=prob]) finally show ?thesis using emeasure_space_1 by (simp add: emeasure_eq_measure) qed lemma (in prob_space) distributed_marginal_eq_joint_simple: assumes Px: "simple_function M X" assumes Py: "simple_distributed M Y Py" assumes Pxy: "simple_distributed M (\x. (X x, Y x)) Pxy" assumes y: "y \ Y`space M" shows "Py y = (\x\X`space M. if (x, y) \ (\x. (X x, Y x)) ` space M then Pxy (x, y) else 0)" proof - note Px = simple_distributedI[OF Px measure_nonneg refl] have "AE y in count_space (Y ` space M). ennreal (Py y) = \\<^sup>+ x. ennreal (if (x, y) \ (\x. (X x, Y x)) ` space M then Pxy (x, y) else 0) \count_space (X ` space M)" using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite simple_distributed[OF Py] simple_distributed_joint[OF Pxy] by (rule distributed_marginal_eq_joint2) (auto intro: Py Px simple_distributed_finite) then have "ennreal (Py y) = (\x\X`space M. ennreal (if (x, y) \ (\x. (X x, Y x)) ` space M then Pxy (x, y) else 0))" using y Px[THEN simple_distributed_finite] by (auto simp: AE_count_space nn_integral_count_space_finite) also have "\ = (\x\X`space M. if (x, y) \ (\x. (X x, Y x)) ` space M then Pxy (x, y) else 0)" using Pxy by (intro sum_ennreal) auto finally show ?thesis using simple_distributed_nonneg[OF Py] simple_distributed_nonneg[OF Pxy] by (subst (asm) ennreal_inj) (auto intro!: sum_nonneg) qed lemma distributedI_real: fixes f :: "'a \ real" assumes gen: "sets M1 = sigma_sets (space M1) E" and "Int_stable E" and A: "range A \ E" "(\i::nat. A i) = space M1" "\i. emeasure (distr M M1 X) (A i) \ \" and X: "X \ measurable M M1" and f: "f \ borel_measurable M1" "AE x in M1. 0 \ f x" and eq: "\A. A \ E \ emeasure M (X -` A \ space M) = (\\<^sup>+ x. f x * indicator A x \M1)" shows "distributed M M1 X f" unfolding distributed_def proof (intro conjI) show "distr M M1 X = density M1 f" proof (rule measure_eqI_generator_eq[where A=A]) { fix A assume A: "A \ E" then have "A \ sigma_sets (space M1) E" by auto then have "A \ sets M1" using gen by simp with f A eq[of A] X show "emeasure (distr M M1 X) A = emeasure (density M1 f) A" by (auto simp add: emeasure_distr emeasure_density ennreal_indicator intro!: nn_integral_cong split: split_indicator) } note eq_E = this show "Int_stable E" by fact { fix e assume "e \ E" then have "e \ sigma_sets (space M1) E" by auto then have "e \ sets M1" unfolding gen . then have "e \ space M1" by (rule sets.sets_into_space) } then show "E \ Pow (space M1)" by auto show "sets (distr M M1 X) = sigma_sets (space M1) E" "sets (density M1 (\x. ennreal (f x))) = sigma_sets (space M1) E" unfolding gen[symmetric] by auto qed fact+ qed (insert X f, auto) lemma distributedI_borel_atMost: fixes f :: "real \ real" assumes [measurable]: "X \ borel_measurable M" and [measurable]: "f \ borel_measurable borel" and f[simp]: "AE x in lborel. 0 \ f x" and g_eq: "\a. (\\<^sup>+x. f x * indicator {..a} x \lborel) = ennreal (g a)" and M_eq: "\a. emeasure M {x\space M. X x \ a} = ennreal (g a)" shows "distributed M lborel X f" proof (rule distributedI_real) show "sets (lborel::real measure) = sigma_sets (space lborel) (range atMost)" by (simp add: borel_eq_atMost) show "Int_stable (range atMost :: real set set)" by (auto simp: Int_stable_def) have vimage_eq: "\a. (X -` {..a} \ space M) = {x\space M. X x \ a}" by auto define A where "A i = {.. real i}" for i :: nat then show "range A \ range atMost" "(\i. A i) = space lborel" "\i. emeasure (distr M lborel X) (A i) \ \" by (auto simp: real_arch_simple emeasure_distr vimage_eq M_eq) fix A :: "real set" assume "A \ range atMost" then obtain a where A: "A = {..a}" by auto show "emeasure M (X -` A \ space M) = (\\<^sup>+x. f x * indicator A x \lborel)" unfolding vimage_eq A M_eq g_eq .. qed auto lemma (in prob_space) uniform_distributed_params: assumes X: "distributed M MX X (\x. indicator A x / measure MX A)" shows "A \ sets MX" "measure MX A \ 0" proof - interpret X: prob_space "distr M MX X" using distributed_measurable[OF X] by (rule prob_space_distr) show "measure MX A \ 0" proof assume "measure MX A = 0" with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X] show False by (simp add: emeasure_density zero_ennreal_def[symmetric]) qed with measure_notin_sets[of A MX] show "A \ sets MX" by blast qed lemma prob_space_uniform_measure: assumes A: "emeasure M A \ 0" "emeasure M A \ \" shows "prob_space (uniform_measure M A)" proof show "emeasure (uniform_measure M A) (space (uniform_measure M A)) = 1" using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)], of "space M"] using sets.sets_into_space[OF emeasure_neq_0_sets[OF A(1)]] A by (simp add: Int_absorb2 less_top) qed lemma prob_space_uniform_count_measure: "finite A \ A \ {} \ prob_space (uniform_count_measure A)" by standard (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ennreal_def) lemma (in prob_space) measure_uniform_measure_eq_cond_prob: assumes [measurable]: "Measurable.pred M P" "Measurable.pred M Q" shows "\

(x in uniform_measure M {x\space M. Q x}. P x) = \

(x in M. P x \ Q x)" proof cases assume Q: "measure M {x\space M. Q x} = 0" then have *: "AE x in M. \ Q x" by (simp add: prob_eq_0) then have "density M (\x. indicator {x \ space M. Q x} x / emeasure M {x \ space M. Q x}) = density M (\x. 0)" by (intro density_cong) auto with * show ?thesis unfolding uniform_measure_def by (simp add: emeasure_density measure_def cond_prob_def emeasure_eq_0_AE) next assume Q: "measure M {x\space M. Q x} \ 0" then show "\

(x in uniform_measure M {x \ space M. Q x}. P x) = cond_prob M P Q" by (subst measure_uniform_measure) (auto simp: emeasure_eq_measure cond_prob_def measure_nonneg intro!: arg_cong[where f=prob]) qed lemma prob_space_point_measure: "finite S \ (\s. s \ S \ 0 \ p s) \ (\s\S. p s) = 1 \ prob_space (point_measure S p)" by (rule prob_spaceI) (simp add: space_point_measure emeasure_point_measure_finite) lemma (in prob_space) distr_pair_fst: "distr (N \\<^sub>M M) N fst = N" proof (intro measure_eqI) fix A assume A: "A \ sets (distr (N \\<^sub>M M) N fst)" from A have "emeasure (distr (N \\<^sub>M M) N fst) A = emeasure (N \\<^sub>M M) (A \ space M)" by (auto simp add: emeasure_distr space_pair_measure dest: sets.sets_into_space intro!: arg_cong2[where f=emeasure]) with A show "emeasure (distr (N \\<^sub>M M) N fst) A = emeasure N A" by (simp add: emeasure_pair_measure_Times emeasure_space_1) qed simp lemma (in product_prob_space) distr_reorder: assumes "inj_on t J" "t \ J \ K" "finite K" shows "distr (PiM K M) (Pi\<^sub>M J (\x. M (t x))) (\\. \n\J. \ (t n)) = PiM J (\x. M (t x))" proof (rule product_sigma_finite.PiM_eqI) show "product_sigma_finite (\x. M (t x))" .. have "t`J \ K" using assms by auto then show [simp]: "finite J" by (rule finite_imageD[OF finite_subset]) fact+ fix A assume A: "\i. i \ J \ A i \ sets (M (t i))" moreover have "((\\. \n\J. \ (t n)) -` Pi\<^sub>E J A \ space (Pi\<^sub>M K M)) = (\\<^sub>E i\K. if i \ t`J then A (the_inv_into J t i) else space (M i))" using A A[THEN sets.sets_into_space] \t \ J \ K\ \inj_on t J\ by (subst prod_emb_Pi[symmetric]) (auto simp: space_PiM PiE_iff the_inv_into_f_f prod_emb_def) ultimately show "distr (Pi\<^sub>M K M) (Pi\<^sub>M J (\x. M (t x))) (\\. \n\J. \ (t n)) (Pi\<^sub>E J A) = (\i\J. M (t i) (A i))" using assms apply (subst emeasure_distr) apply (auto intro!: sets_PiM_I_finite simp: Pi_iff) apply (subst emeasure_PiM) apply (auto simp: the_inv_into_f_f \inj_on t J\ prod.reindex[OF \inj_on t J\] if_distrib[where f="emeasure (M _)"] prod.If_cases emeasure_space_1 Int_absorb1 \t`J \ K\) done qed simp lemma (in product_prob_space) distr_restrict: "J \ K \ finite K \ (\\<^sub>M i\J. M i) = distr (\\<^sub>M i\K. M i) (\\<^sub>M i\J. M i) (\f. restrict f J)" using distr_reorder[of "\x. x" J K] by (simp add: Pi_iff subset_eq) lemma (in product_prob_space) emeasure_prod_emb[simp]: assumes L: "J \ L" "finite L" and X: "X \ sets (Pi\<^sub>M J M)" shows "emeasure (Pi\<^sub>M L M) (prod_emb L M J X) = emeasure (Pi\<^sub>M J M) X" by (subst distr_restrict[OF L]) (simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X) lemma emeasure_distr_restrict: assumes "I \ K" and Q[measurable_cong]: "sets Q = sets (PiM K M)" and A[measurable]: "A \ sets (PiM I M)" shows "emeasure (distr Q (PiM I M) (\\. restrict \ I)) A = emeasure Q (prod_emb K M I A)" using \I\K\ sets_eq_imp_space_eq[OF Q] by (subst emeasure_distr) (auto simp: measurable_cong_sets[OF Q] prod_emb_def space_PiM[symmetric] intro!: measurable_restrict) lemma (in prob_space) prob_space_completion: "prob_space (completion M)" by (rule prob_spaceI) (simp add: emeasure_space_1) lemma distr_PiM_finite_prob_space: assumes fin: "finite I" assumes "product_prob_space M" assumes "product_prob_space M'" assumes [measurable]: "\i. i \ I \ f \ measurable (M i) (M' i)" shows "distr (PiM I M) (PiM I M') (compose I f) = PiM I (\i. distr (M i) (M' i) f)" proof - interpret M: product_prob_space M by fact interpret M': product_prob_space M' by fact define N where "N = (\i. if i \ I then distr (M i) (M' i) f else M' i)" have [intro]: "prob_space (N i)" for i by (auto simp: N_def intro!: M.M.prob_space_distr M'.prob_space) interpret N: product_prob_space N by (intro product_prob_spaceI) (auto simp: N_def M'.prob_space intro: M.M.prob_space_distr) have "distr (PiM I M) (PiM I M') (compose I f) = PiM I N" proof (rule N.PiM_eqI) have N_events_eq: "sets (Pi\<^sub>M I N) = sets (Pi\<^sub>M I M')" unfolding N_def by (intro sets_PiM_cong) auto also have "\ = sets (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f))" by simp finally show "sets (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f)) = sets (Pi\<^sub>M I N)" .. fix A assume A: "\i. i \ I \ A i \ N.M.events i" have "emeasure (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f)) (Pi\<^sub>E I A) = emeasure (Pi\<^sub>M I M) (compose I f -` Pi\<^sub>E I A \ space (Pi\<^sub>M I M))" proof (intro emeasure_distr) show "compose I f \ Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M I M'" unfolding compose_def by measurable show "Pi\<^sub>E I A \ sets (Pi\<^sub>M I M')" unfolding N_events_eq [symmetric] by (intro sets_PiM_I_finite fin A) qed also have "compose I f -` Pi\<^sub>E I A \ space (Pi\<^sub>M I M) = Pi\<^sub>E I (\i. f -` A i \ space (M i))" using A by (auto simp: space_PiM PiE_def Pi_def extensional_def N_def compose_def) also have "emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I (\i. f -` A i \ space (M i))) = (\i\I. emeasure (M i) (f -` A i \ space (M i)))" using A by (intro M.emeasure_PiM fin) (auto simp: N_def) also have "\ = (\i\I. emeasure (distr (M i) (M' i) f) (A i))" using A by (intro prod.cong emeasure_distr [symmetric]) (auto simp: N_def) also have "\ = (\i\I. emeasure (N i) (A i))" unfolding N_def by (intro prod.cong) (auto simp: N_def) finally show "emeasure (distr (Pi\<^sub>M I M) (Pi\<^sub>M I M') (compose I f)) (Pi\<^sub>E I A) = \" . qed fact+ also have "PiM I N = PiM I (\i. distr (M i) (M' i) f)" by (intro PiM_cong) (auto simp: N_def) finally show ?thesis . qed end diff --git a/src/HOL/Probability/Projective_Limit.thy b/src/HOL/Probability/Projective_Limit.thy --- a/src/HOL/Probability/Projective_Limit.thy +++ b/src/HOL/Probability/Projective_Limit.thy @@ -1,483 +1,487 @@ (* Title: HOL/Probability/Projective_Limit.thy Author: Fabian Immler, TU München *) section \Projective Limit\ theory Projective_Limit imports Fin_Map Infinite_Product_Measure "HOL-Library.Diagonal_Subsequence" begin subsection \Sequences of Finite Maps in Compact Sets\ locale finmap_seqs_into_compact = fixes K::"nat \ (nat \\<^sub>F 'a::metric_space) set" and f::"nat \ (nat \\<^sub>F 'a)" and M assumes compact: "\n. compact (K n)" assumes f_in_K: "\n. K n \ {}" assumes domain_K: "\n. k \ K n \ domain k = domain (f n)" assumes proj_in_K: "\t n m. m \ n \ t \ domain (f n) \ (f m)\<^sub>F t \ (\k. (k)\<^sub>F t) ` K n" begin lemma proj_in_K': "(\n. \m \ n. (f m)\<^sub>F t \ (\k. (k)\<^sub>F t) ` K n)" using proj_in_K f_in_K proof cases obtain k where "k \ K (Suc 0)" using f_in_K by auto assume "\n. t \ domain (f n)" thus ?thesis by (auto intro!: exI[where x=1] image_eqI[OF _ \k \ K (Suc 0)\] simp: domain_K[OF \k \ K (Suc 0)\]) qed blast lemma proj_in_KE: obtains n where "\m. m \ n \ (f m)\<^sub>F t \ (\k. (k)\<^sub>F t) ` K n" using proj_in_K' by blast lemma compact_projset: shows "compact ((\k. (k)\<^sub>F i) ` K n)" using continuous_proj compact by (rule compact_continuous_image) end lemma compactE': fixes S :: "'a :: metric_space set" assumes "compact S" "\n\m. f n \ S" obtains l r where "l \ S" "strict_mono (r::nat\nat)" "((f \ r) \ l) sequentially" proof atomize_elim have "strict_mono ((+) m)" by (simp add: strict_mono_def) have "\n. (f o (\i. m + i)) n \ S" using assms by auto - from seq_compactE[OF \compact S\[unfolded compact_eq_seq_compact_metric] this] guess l r . + from seq_compactE[OF \compact S\[unfolded compact_eq_seq_compact_metric] this] + obtain l r where "l \ S" "strict_mono r" "(f \ (+) m \ r) \ l" by blast hence "l \ S" "strict_mono ((\i. m + i) o r) \ (f \ ((\i. m + i) o r)) \ l" using strict_mono_o[OF \strict_mono ((+) m)\ \strict_mono r\] by (auto simp: o_def) thus "\l r. l \ S \ strict_mono r \ (f \ r) \ l" by blast qed sublocale finmap_seqs_into_compact \ subseqs "\n s. (\l. (\i. ((f o s) i)\<^sub>F n) \ l)" proof fix n and s :: "nat \ nat" assume "strict_mono s" - from proj_in_KE[of n] guess n0 . note n0 = this + from proj_in_KE[of n] obtain n0 where n0: "\m. n0 \ m \ (f m)\<^sub>F n \ (\k. (k)\<^sub>F n) ` K n0" + by blast have "\i \ n0. ((f \ s) i)\<^sub>F n \ (\k. (k)\<^sub>F n) ` K n0" proof safe fix i assume "n0 \ i" also have "\ \ s i" by (rule seq_suble) fact finally have "n0 \ s i" . with n0 show "((f \ s) i)\<^sub>F n \ (\k. (k)\<^sub>F n) ` K n0 " by auto qed - from compactE'[OF compact_projset this] guess ls rs . + then obtain ls rs + where "ls \ (\k. (k)\<^sub>F n) ` K n0" "strict_mono rs" "((\i. ((f \ s) i)\<^sub>F n) \ rs) \ ls" + by (rule compactE'[OF compact_projset]) thus "\r'. strict_mono r' \ (\l. (\i. ((f \ (s \ r')) i)\<^sub>F n) \ l)" by (auto simp: o_def) qed lemma (in finmap_seqs_into_compact) diagonal_tendsto: "\l. (\i. (f (diagseq i))\<^sub>F n) \ l" proof - obtain l where "(\i. ((f o (diagseq o (+) (Suc n))) i)\<^sub>F n) \ l" proof (atomize_elim, rule diagseq_holds) fix r s n assume "strict_mono (r :: nat \ nat)" assume "\l. (\i. ((f \ s) i)\<^sub>F n) \ l" then obtain l where "((\i. (f i)\<^sub>F n) o s) \ l" by (auto simp: o_def) hence "((\i. (f i)\<^sub>F n) o s o r) \ l" using \strict_mono r\ by (rule LIMSEQ_subseq_LIMSEQ) thus "\l. (\i. ((f \ (s \ r)) i)\<^sub>F n) \ l" by (auto simp add: o_def) qed hence "(\i. ((f (diagseq (i + Suc n))))\<^sub>F n) \ l" by (simp add: ac_simps) hence "(\i. (f (diagseq i))\<^sub>F n) \ l" by (rule LIMSEQ_offset) thus ?thesis .. qed subsection \Daniell-Kolmogorov Theorem\ text \Existence of Projective Limit\ locale polish_projective = projective_family I P "\_. borel::'a::polish_space measure" for I::"'i set" and P begin lemma emeasure_lim_emb: assumes X: "J \ I" "finite J" "X \ sets (\\<^sub>M i\J. borel)" shows "lim (emb I J X) = P J X" proof (rule emeasure_lim) write mu_G ("\G") interpret generator: algebra "space (PiM I (\i. borel))" generator by (rule algebra_generator) fix J and B :: "nat \ ('i \ 'a) set" assume J: "\n. finite (J n)" "\n. J n \ I" "\n. B n \ sets (\\<^sub>M i\J n. borel)" "incseq J" and B: "decseq (\n. emb I (J n) (B n))" and "0 < (INF i. P (J i) (B i))" (is "0 < ?a") moreover have "?a \ 1" using J by (auto intro!: INF_lower2[of 0] prob_space_P[THEN prob_space.measure_le_1]) ultimately obtain r where r: "?a = ennreal r" "0 < r" "r \ 1" by (cases "?a") (auto simp: top_unique) define Z where "Z n = emb I (J n) (B n)" for n have Z_mono: "n \ m \ Z m \ Z n" for n m unfolding Z_def using B[THEN antimonoD, of n m] . have J_mono: "\n m. n \ m \ J n \ J m" using \incseq J\ by (force simp: incseq_def) note [simp] = \\n. finite (J n)\ interpret prob_space "P (J i)" for i using J prob_space_P by simp have P_eq[simp]: "sets (P (J i)) = sets (\\<^sub>M i\J i. borel)" "space (P (J i)) = space (\\<^sub>M i\J i. borel)" for i using J by (auto simp: sets_P space_P) have "Z i \ generator" for i unfolding Z_def by (auto intro!: generator.intros J) have countable_UN_J: "countable (\n. J n)" by (simp add: countable_finite) define Utn where "Utn = to_nat_on (\n. J n)" interpret function_to_finmap "J n" Utn "from_nat_into (\n. J n)" for n by unfold_locales (auto simp: Utn_def intro: from_nat_into_to_nat_on[OF countable_UN_J]) have inj_on_Utn: "inj_on Utn (\n. J n)" unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on) hence inj_on_Utn_J: "\n. inj_on Utn (J n)" by (rule subset_inj_on) auto define P' where "P' n = mapmeasure n (P (J n)) (\_. borel)" for n interpret P': prob_space "P' n" for n unfolding P'_def mapmeasure_def using J by (auto intro!: prob_space_distr fm_measurable simp: measurable_cong_sets[OF sets_P]) let ?SUP = "\n. SUP K \ {K. K \ fm n ` (B n) \ compact K}. emeasure (P' n) K" { fix n have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))" using J by (auto simp: P'_def mapmeasure_PiM space_P sets_P) also have "\ = ?SUP n" proof (rule inner_regular) show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def) next show "fm n ` B n \ sets borel" unfolding borel_eq_PiF_borel by (auto simp: P'_def fm_image_measurable_finite sets_P J(3)) qed simp finally have *: "emeasure (P (J n)) (B n) = ?SUP n" . have "?SUP n \ \" unfolding *[symmetric] by simp note * this } note R = this have "\n. \K. emeasure (P (J n)) (B n) - emeasure (P' n) K \ 2 powr (-n) * ?a \ compact K \ K \ fm n ` B n" proof fix n show "\K. emeasure (P (J n)) (B n) - emeasure (P' n) K \ ennreal (2 powr - real n) * ?a \ compact K \ K \ fm n ` B n" unfolding R[of n] proof (rule ccontr) assume H: "\K'. ?SUP n - emeasure (P' n) K' \ ennreal (2 powr - real n) * ?a \ compact K' \ K' \ fm n ` B n" have "?SUP n + 0 < ?SUP n + 2 powr (-n) * ?a" using R[of n] unfolding ennreal_add_left_cancel_less ennreal_zero_less_mult_iff by (auto intro: \0 < ?a\) also have "\ = (SUP K\{K. K \ fm n ` B n \ compact K}. emeasure (P' n) K + 2 powr (-n) * ?a)" by (rule ennreal_SUP_add_left[symmetric]) auto also have "\ \ ?SUP n" proof (intro SUP_least) fix K assume "K \ {K. K \ fm n ` B n \ compact K}" with H have "2 powr (-n) * ?a < ?SUP n - emeasure (P' n) K" by auto then show "emeasure (P' n) K + (2 powr (-n)) * ?a \ ?SUP n" by (subst (asm) less_diff_eq_ennreal) (auto simp: less_top[symmetric] R(1)[symmetric] ac_simps) qed finally show False by simp qed qed then obtain K' where K': "\n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \ ennreal (2 powr - real n) * ?a" "\n. compact (K' n)" "\n. K' n \ fm n ` B n" unfolding choice_iff by blast define K where "K n = fm n -` K' n \ space (Pi\<^sub>M (J n) (\_. borel))" for n have K_sets: "\n. K n \ sets (Pi\<^sub>M (J n) (\_. borel))" unfolding K_def using compact_imp_closed[OF \compact (K' _)\] by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"]) (auto simp: borel_eq_PiF_borel[symmetric]) have K_B: "\n. K n \ B n" proof fix x n assume "x \ K n" then have fm_in: "fm n x \ fm n ` B n" using K' by (force simp: K_def) show "x \ B n" using \x \ K n\ K_sets sets.sets_into_space J(1,2,3)[of n] inj_on_image_mem_iff[OF inj_on_fm] by (metis (no_types) Int_iff K_def fm_in space_borel) qed define Z' where "Z' n = emb I (J n) (K n)" for n have Z': "\n. Z' n \ Z n" unfolding Z'_def Z_def proof (rule prod_emb_mono, safe) fix n x assume "x \ K n" hence "fm n x \ K' n" "x \ space (Pi\<^sub>M (J n) (\_. borel))" by (simp_all add: K_def space_P) note this(1) also have "K' n \ fm n ` B n" by (simp add: K') finally have "fm n x \ fm n ` B n" . thus "x \ B n" proof safe fix y assume y: "y \ B n" hence "y \ space (Pi\<^sub>M (J n) (\_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"] by (auto simp add: space_P sets_P) assume "fm n x = fm n y" note inj_onD[OF inj_on_fm[OF space_borel], OF \fm n x = fm n y\ \x \ space _\ \y \ space _\] with y show "x \ B n" by simp qed qed have "\n. Z' n \ generator" using J K'(2) unfolding Z'_def by (auto intro!: generator.intros measurable_sets[OF fm_measurable[of _ "Collect finite"]] simp: K_def borel_eq_PiF_borel[symmetric] compact_imp_closed) define Y where "Y n = (\i\{1..n}. Z' i)" for n hence "\n k. Y (n + k) \ Y n" by (induct_tac k) (auto simp: Y_def) hence Y_mono: "\n m. n \ m \ Y m \ Y n" by (auto simp: le_iff_add) have Y_Z': "\n. n \ 1 \ Y n \ Z' n" by (auto simp: Y_def) hence Y_Z: "\n. n \ 1 \ Y n \ Z n" using Z' by auto have Y_notempty: "\n. n \ 1 \ (Y n) \ {}" proof - fix n::nat assume "n \ 1" hence "Y n \ Z n" by fact have "Y n = (\i\{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono by (auto simp: Y_def Z'_def) also have "\ = prod_emb I (\_. borel) (J n) (\i\{1..n}. emb (J n) (J i) (K i))" using \n \ 1\ by (subst prod_emb_INT) auto finally have Y_emb: "Y n = prod_emb I (\_. borel) (J n) (\i\{1..n}. prod_emb (J n) (\_. borel) (J i) (K i))" . hence "Y n \ generator" using J J_mono K_sets \n \ 1\ by (auto simp del: prod_emb_INT intro!: generator.intros) have *: "\G (Z n) = P (J n) (B n)" unfolding Z_def using J by (intro mu_G_spec) auto then have "\G (Z n) \ \" by auto note * moreover have *: "\G (Y n) = P (J n) (\i\{Suc 0..n}. prod_emb (J n) (\_. borel) (J i) (K i))" unfolding Y_emb using J J_mono K_sets \n \ 1\ by (subst mu_G_spec) auto then have "\G (Y n) \ \" by auto note * moreover have "\G (Z n - Y n) = P (J n) (B n - (\i\{Suc 0..n}. prod_emb (J n) (\_. borel) (J i) (K i)))" unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets \n \ 1\ by (subst mu_G_spec) (auto intro!: sets.Diff) ultimately have "\G (Z n) - \G (Y n) = \G (Z n - Y n)" using J J_mono K_sets \n \ 1\ by (simp only: emeasure_eq_measure Z_def) (auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] subsetD[OF K_B] intro!: arg_cong[where f=ennreal] simp: extensional_restrict emeasure_eq_measure prod_emb_iff sets_P space_P ennreal_minus measure_nonneg) also have subs: "Z n - Y n \ (\i\{1..n}. (Z i - Z' i))" using \n \ 1\ unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto have "Z n - Y n \ generator" "(\i\{1..n}. (Z i - Z' i)) \ generator" using \Z' _ \ generator\ \Z _ \ generator\ \Y _ \ generator\ by auto hence "\G (Z n - Y n) \ \G (\i\{1..n}. (Z i - Z' i))" using subs generator.additive_increasing[OF positive_mu_G additive_mu_G] unfolding increasing_def by auto also have "\ \ (\ i\{1..n}. \G (Z i - Z' i))" using \Z _ \ generator\ \Z' _ \ generator\ by (intro generator.subadditive[OF positive_mu_G additive_mu_G]) auto also have "\ \ (\ i\{1..n}. 2 powr -real i * ?a)" proof (rule sum_mono) fix i assume "i \ {1..n}" hence "i \ n" by simp have "\G (Z i - Z' i) = \G (prod_emb I (\_. borel) (J i) (B i - K i))" unfolding Z'_def Z_def by simp also have "\ = P (J i) (B i - K i)" using J K_sets by (subst mu_G_spec) auto also have "\ = P (J i) (B i) - P (J i) (K i)" using K_sets J \K _ \ B _\ by (simp add: emeasure_Diff) also have "\ = P (J i) (B i) - P' i (K' i)" unfolding K_def P'_def by (auto simp: mapmeasure_PiF borel_eq_PiF_borel[symmetric] compact_imp_closed[OF \compact (K' _)\] space_PiM PiE_def) also have "\ \ ennreal (2 powr - real i) * ?a" using K'(1)[of i] . finally show "\G (Z i - Z' i) \ (2 powr - real i) * ?a" . qed also have "\ = ennreal ((\ i\{1..n}. (2 powr -enn2real i)) * enn2real ?a)" using r by (simp add: sum_distrib_right ennreal_mult[symmetric]) also have "\ < ennreal (1 * enn2real ?a)" proof (intro ennreal_lessI mult_strict_right_mono) have "(\i = 1..n. 2 powr - real i) = (\i = 1.. < 1" by (subst geometric_sum) auto finally show "(\i = 1..n. 2 powr - enn2real i) < 1" by simp qed (auto simp: r enn2real_positive_iff) also have "\ = ?a" by (auto simp: r) also have "\ \ \G (Z n)" using J by (auto intro: INF_lower simp: Z_def mu_G_spec) finally have "\G (Z n) - \G (Y n) < \G (Z n)" . hence R: "\G (Z n) < \G (Z n) + \G (Y n)" using \\G (Y n) \ \\ by (auto simp: zero_less_iff_neq_zero) then have "\G (Y n) > 0" by simp thus "Y n \ {}" using positive_mu_G by (auto simp add: positive_def) qed hence "\n\{1..}. \y. y \ Y n" by auto then obtain y where y: "\n. n \ 1 \ y n \ Y n" unfolding bchoice_iff by force { fix t and n m::nat assume "1 \ n" "n \ m" hence "1 \ m" by simp from Y_mono[OF \m \ n\] y[OF \1 \ m\] have "y m \ Y n" by auto also have "\ \ Z' n" using Y_Z'[OF \1 \ n\] . finally have "fm n (restrict (y m) (J n)) \ K' n" unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff) moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)" using J by (simp add: fm_def) ultimately have "fm n (y m) \ K' n" by simp } note fm_in_K' = this interpret finmap_seqs_into_compact "\n. K' (Suc n)" "\k. fm (Suc k) (y (Suc k))" borel proof fix n show "compact (K' n)" by fact next fix n from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \ Y (Suc n)" by auto also have "\ \ Z' (Suc n)" using Y_Z' by auto finally have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \ K' (Suc n)" unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff) thus "K' (Suc n) \ {}" by auto fix k assume "k \ K' (Suc n)" with K'[of "Suc n"] sets.sets_into_space have "k \ fm (Suc n) ` B (Suc n)" by auto then obtain b where "k = fm (Suc n) b" by auto thus "domain k = domain (fm (Suc n) (y (Suc n)))" by (simp_all add: fm_def) next fix t and n m::nat assume "n \ m" hence "Suc n \ Suc m" by simp assume "t \ domain (fm (Suc n) (y (Suc n)))" then obtain j where j: "t = Utn j" "j \ J (Suc n)" by auto hence "j \ J (Suc m)" using J_mono[OF \Suc n \ Suc m\] by auto have img: "fm (Suc n) (y (Suc m)) \ K' (Suc n)" using \n \ m\ by (intro fm_in_K') simp_all show "(fm (Suc m) (y (Suc m)))\<^sub>F t \ (\k. (k)\<^sub>F t) ` K' (Suc n)" apply (rule image_eqI[OF _ img]) using \j \ J (Suc n)\ \j \ J (Suc m)\ unfolding j by (subst proj_fm, auto)+ qed have "\t. \z. (\i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) \ z" using diagonal_tendsto .. then obtain z where z: "\t. (\i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) \ z t" unfolding choice_iff by blast { fix n :: nat assume "n \ 1" have "\i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)" by simp moreover { fix t assume t: "t \ domain (finmap_of (Utn ` J n) z)" hence "t \ Utn ` J n" by simp then obtain j where j: "t = Utn j" "j \ J n" by auto have "(\i. (fm n (y (Suc (diagseq i))))\<^sub>F t) \ z t" apply (subst (2) tendsto_iff, subst eventually_sequentially) proof safe fix e :: real assume "0 < e" { fix i and x :: "'i \ 'a" assume i: "i \ n" assume "t \ domain (fm n x)" hence "t \ domain (fm i x)" using J_mono[OF \i \ n\] by auto with i have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t" using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn]) } note index_shift = this have I: "\i. i \ n \ Suc (diagseq i) \ n" apply (rule le_SucI) apply (rule order_trans) apply simp apply (rule seq_suble[OF subseq_diagseq]) done from z have "\N. \i\N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" unfolding tendsto_iff eventually_sequentially using \0 < e\ by auto then obtain N where N: "\i. i \ N \ dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto show "\N. \na\N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e " proof (rule exI[where x="max N n"], safe) fix na assume "max N n \ na" hence "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) = dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t by (subst index_shift[OF I]) auto also have "\ < e" using \max N n \ na\ by (intro N) simp finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" . qed qed hence "(\i. (fm n (y (Suc (diagseq i))))\<^sub>F t) \ (finmap_of (Utn ` J n) z)\<^sub>F t" by (simp add: tendsto_intros) } ultimately have "(\i. fm n (y (Suc (diagseq i)))) \ finmap_of (Utn ` J n) z" by (rule tendsto_finmap) hence "((\i. fm n (y (Suc (diagseq i)))) o (\i. i + n)) \ finmap_of (Utn ` J n) z" by (rule LIMSEQ_subseq_LIMSEQ) (simp add: strict_mono_def) moreover have "(\i. ((\i. fm n (y (Suc (diagseq i)))) o (\i. i + n)) i \ K' n)" apply (auto simp add: o_def intro!: fm_in_K' \1 \ n\ le_SucI) apply (rule le_trans) apply (rule le_add2) using seq_suble[OF subseq_diagseq] apply auto done moreover from \compact (K' n)\ have "closed (K' n)" by (rule compact_imp_closed) ultimately have "finmap_of (Utn ` J n) z \ K' n" unfolding closed_sequential_limits by blast also have "finmap_of (Utn ` J n) z = fm n (\i. z (Utn i))" unfolding finmap_eq_iff proof clarsimp fix i assume i: "i \ J n" hence "from_nat_into (\n. J n) (Utn i) = i" unfolding Utn_def by (subst from_nat_into_to_nat_on[OF countable_UN_J]) auto with i show "z (Utn i) = (fm n (\i. z (Utn i)))\<^sub>F (Utn i)" by (simp add: finmap_eq_iff fm_def compose_def) qed finally have "fm n (\i. z (Utn i)) \ K' n" . moreover let ?J = "\n. J n" have "(?J \ J n) = J n" by auto ultimately have "restrict (\i. z (Utn i)) (?J \ J n) \ K n" unfolding K_def by (auto simp: space_P space_PiM) hence "restrict (\i. z (Utn i)) ?J \ Z' n" unfolding Z'_def using J by (auto simp: prod_emb_def PiE_def extensional_def) also have "\ \ Z n" using Z' by simp finally have "restrict (\i. z (Utn i)) ?J \ Z n" . } note in_Z = this hence "(\i\{1..}. Z i) \ {}" by auto thus "(\i. Z i) \ {}" using INT_decseq_offset[OF antimonoI[OF Z_mono]] by simp qed fact+ lemma measure_lim_emb: "J \ I \ finite J \ X \ sets (\\<^sub>M i\J. borel) \ measure lim (emb I J X) = measure (P J) X" unfolding measure_def by (subst emeasure_lim_emb) auto end hide_const (open) PiF hide_const (open) Pi\<^sub>F hide_const (open) Pi' hide_const (open) finmap_of hide_const (open) proj hide_const (open) domain hide_const (open) basis_finmap sublocale polish_projective \ P: prob_space lim proof have *: "emb I {} {\x. undefined} = space (\\<^sub>M i\I. borel)" by (auto simp: prod_emb_def space_PiM) interpret prob_space "P {}" using prob_space_P by simp show "emeasure lim (space lim) = 1" using emeasure_lim_emb[of "{}" "{\x. undefined}"] emeasure_space_1 by (simp add: * PiM_empty space_P) qed locale polish_product_prob_space = product_prob_space "\_. borel::('a::polish_space) measure" I for I::"'i set" sublocale polish_product_prob_space \ P: polish_projective I "\J. PiM J (\_. borel::('a) measure)" .. lemma (in polish_product_prob_space) limP_eq_PiM: "lim = PiM I (\_. borel)" by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_lim_emb) end diff --git a/src/HOL/Probability/Stopping_Time.thy b/src/HOL/Probability/Stopping_Time.thy --- a/src/HOL/Probability/Stopping_Time.thy +++ b/src/HOL/Probability/Stopping_Time.thy @@ -1,262 +1,263 @@ (* Author: Johannes Hölzl *) section \Stopping times\ theory Stopping_Time imports "HOL-Analysis.Analysis" begin subsection \Stopping Time\ text \This is also called strong stopping time. Then stopping time is T with alternative is \T x < t\ measurable.\ definition stopping_time :: "('t::linorder \ 'a measure) \ ('a \ 't) \ bool" where "stopping_time F T = (\t. Measurable.pred (F t) (\x. T x \ t))" lemma stopping_time_cong: "(\t x. x \ space (F t) \ T x = S x) \ stopping_time F T = stopping_time F S" unfolding stopping_time_def by (intro arg_cong[where f=All] ext measurable_cong) simp lemma stopping_timeD: "stopping_time F T \ Measurable.pred (F t) (\x. T x \ t)" by (auto simp: stopping_time_def) lemma stopping_timeD2: "stopping_time F T \ Measurable.pred (F t) (\x. t < T x)" unfolding not_le[symmetric] by (auto intro: stopping_timeD Measurable.pred_intros_logic) lemma stopping_timeI[intro?]: "(\t. Measurable.pred (F t) (\x. T x \ t)) \ stopping_time F T" by (auto simp: stopping_time_def) lemma measurable_stopping_time: fixes T :: "'a \ 't::{linorder_topology, second_countable_topology}" assumes T: "stopping_time F T" and M: "\t. sets (F t) \ sets M" "\t. space (F t) = space M" shows "T \ M \\<^sub>M borel" proof (rule borel_measurableI_le) show "{x \ space M. T x \ t} \ sets M" for t using stopping_timeD[OF T] M by (auto simp: Measurable.pred_def) qed lemma stopping_time_const: "stopping_time F (\x. c)" by (auto simp: stopping_time_def) lemma stopping_time_min: "stopping_time F T \ stopping_time F S \ stopping_time F (\x. min (T x) (S x))" by (auto simp: stopping_time_def min_le_iff_disj intro!: pred_intros_logic) lemma stopping_time_max: "stopping_time F T \ stopping_time F S \ stopping_time F (\x. max (T x) (S x))" by (auto simp: stopping_time_def intro!: pred_intros_logic) section \Filtration\ locale filtration = fixes \ :: "'a set" and F :: "'t::{linorder_topology, second_countable_topology} \ 'a measure" assumes space_F: "\i. space (F i) = \" assumes sets_F_mono: "\i j. i \ j \ sets (F i) \ sets (F j)" begin subsection \$\sigma$-algebra of a Stopping Time\ definition pre_sigma :: "('a \ 't) \ 'a measure" where "pre_sigma T = sigma \ {A. \t. {\\A. T \ \ t} \ sets (F t)}" lemma space_pre_sigma: "space (pre_sigma T) = \" unfolding pre_sigma_def using sets.space_closed[of "F _"] by (intro space_measure_of) (auto simp: space_F) lemma measure_pre_sigma[simp]: "emeasure (pre_sigma T) = (\_. 0)" by (simp add: pre_sigma_def emeasure_sigma) lemma sigma_algebra_pre_sigma: assumes T: "stopping_time F T" shows "sigma_algebra \ {A. \t. {\\A. T \ \ t} \ sets (F t)}" unfolding sigma_algebra_iff2 proof (intro sigma_algebra_iff2[THEN iffD2] conjI ballI allI impI CollectI) show "{A. \t. {\ \ A. T \ \ t} \ sets (F t)} \ Pow \" using sets.space_closed[of "F _"] by (auto simp: space_F) next fix A t assume "A \ {A. \t. {\ \ A. T \ \ t} \ sets (F t)}" then have "{\ \ space (F t). T \ \ t} - {\ \ A. T \ \ t} \ sets (F t)" using T stopping_timeD[measurable] by auto also have "{\ \ space (F t). T \ \ t} - {\ \ A. T \ \ t} = {\ \ \ - A. T \ \ t}" by (auto simp: space_F) finally show "{\ \ \ - A. T \ \ t} \ sets (F t)" . next fix AA :: "nat \ 'a set" and t assume "range AA \ {A. \t. {\ \ A. T \ \ t} \ sets (F t)}" then have "(\i. {\ \ AA i. T \ \ t}) \ sets (F t)" for t by auto also have "(\i. {\ \ AA i. T \ \ t}) = {\ \ \(AA ` UNIV). T \ \ t}" by auto finally show "{\ \ \(AA ` UNIV). T \ \ t} \ sets (F t)" . qed auto lemma sets_pre_sigma: "stopping_time F T \ sets (pre_sigma T) = {A. \t. {\\A. T \ \ t} \ sets (F t)}" unfolding pre_sigma_def by (rule sigma_algebra.sets_measure_of_eq[OF sigma_algebra_pre_sigma]) lemma sets_pre_sigmaI: "stopping_time F T \ (\t. {\\A. T \ \ t} \ sets (F t)) \ A \ sets (pre_sigma T)" unfolding sets_pre_sigma by auto lemma pred_pre_sigmaI: assumes T: "stopping_time F T" shows "(\t. Measurable.pred (F t) (\\. P \ \ T \ \ t)) \ Measurable.pred (pre_sigma T) P" unfolding pred_def space_F space_pre_sigma by (intro sets_pre_sigmaI[OF T]) simp lemma sets_pre_sigmaD: "stopping_time F T \ A \ sets (pre_sigma T) \ {\\A. T \ \ t} \ sets (F t)" unfolding sets_pre_sigma by auto lemma stopping_time_le_const: "stopping_time F T \ s \ t \ Measurable.pred (F t) (\\. T \ \ s)" using stopping_timeD[of F T] sets_F_mono[of _ t] by (auto simp: pred_def space_F) lemma measurable_stopping_time_pre_sigma: assumes T: "stopping_time F T" shows "T \ pre_sigma T \\<^sub>M borel" proof (intro borel_measurableI_le sets_pre_sigmaI[OF T]) fix t t' have "{\\space (F (min t' t)). T \ \ min t' t} \ sets (F (min t' t))" using T unfolding pred_def[symmetric] by (rule stopping_timeD) also have "\ \ sets (F t)" by (rule sets_F_mono) simp finally show "{\ \ {x \ space (pre_sigma T). T x \ t'}. T \ \ t} \ sets (F t)" by (simp add: space_pre_sigma space_F) qed lemma mono_pre_sigma: assumes T: "stopping_time F T" and S: "stopping_time F S" and le: "\\. \ \ \ \ T \ \ S \" shows "sets (pre_sigma T) \ sets (pre_sigma S)" unfolding sets_pre_sigma[OF S] sets_pre_sigma[OF T] proof safe interpret sigma_algebra \ "{A. \t. {\\A. T \ \ t} \ sets (F t)}" using T by (rule sigma_algebra_pre_sigma) fix A t assume A: "\t. {\\A. T \ \ t} \ sets (F t)" then have "A \ \" using sets_into_space by auto from A have "{\\A. T \ \ t} \ {\\space (F t). S \ \ t} \ sets (F t)" using stopping_timeD[OF S] by (auto simp: pred_def) also have "{\\A. T \ \ t} \ {\\space (F t). S \ \ t} = {\\A. S \ \ t}" using \A \ \\ sets_into_space[of A] le by (auto simp: space_F intro: order_trans) finally show "{\\A. S \ \ t} \ sets (F t)" by auto qed lemma stopping_time_less_const: assumes T: "stopping_time F T" shows "Measurable.pred (F t) (\\. T \ < t)" proof - - guess D :: "'t set" by (rule countable_dense_setE) - note D = this + obtain D :: "'t set" + where D: "countable D" "\X. open X \ X \ {} \ \d\D. d \ X" + using countable_dense_setE by auto show ?thesis proof cases assume *: "\t't''. t' < t'' \ t'' < t" { fix t' assume "t' < t" with * have "{t' <..< t} \ {}" by fastforce with D(2)[OF _ this] have "\d\D. t'< d \ d < t" by auto } note ** = this show ?thesis proof (rule measurable_cong[THEN iffD2]) show "T \ < t \ (\r\{r\D. r < t}. T \ \ r)" for \ by (auto dest: ** intro: less_imp_le) show "Measurable.pred (F t) (\w. \r\{r \ D. r < t}. T w \ r)" by (intro measurable_pred_countable stopping_time_le_const[OF T] countable_Collect D) auto qed next assume "\ (\t't''. t' < t'' \ t'' < t)" then obtain t' where t': "t' < t" "\t''. t'' < t \ t'' \ t'" by (auto simp: not_less[symmetric]) show ?thesis proof (rule measurable_cong[THEN iffD2]) show "T \ < t \ T \ \ t'" for \ using t' by auto show "Measurable.pred (F t) (\w. T w \ t')" using \t' by (intro stopping_time_le_const[OF T]) auto qed qed qed lemma stopping_time_eq_const: "stopping_time F T \ Measurable.pred (F t) (\\. T \ = t)" unfolding eq_iff using stopping_time_less_const[of T t] by (intro pred_intros_logic stopping_time_le_const) (auto simp: not_less[symmetric] ) lemma stopping_time_less: assumes T: "stopping_time F T" and S: "stopping_time F S" shows "Measurable.pred (pre_sigma T) (\\. T \ < S \)" proof (rule pred_pre_sigmaI[OF T]) fix t obtain D :: "'t set" where [simp]: "countable D" and semidense_D: "\x y. x < y \ (\b\D. x \ b \ b < y)" using countable_separating_set_linorder2 by auto show "Measurable.pred (F t) (\\. T \ < S \ \ T \ \ t)" proof (rule measurable_cong[THEN iffD2]) let ?f = "\\. if T \ = t then \ S \ \ t else \s\{s\D. s \ t}. T \ \ s \ \ (S \ \ s)" { fix \ assume "T \ \ t" "T \ \ t" "T \ < S \" then have "T \ < min t (S \)" by auto then obtain r where "r \ D" "T \ \ r" "r < min t (S \)" by (metis semidense_D) then have "\s\{s\D. s \ t}. T \ \ s \ s < S \" by auto } then show "(T \ < S \ \ T \ \ t) = ?f \" for \ by (auto simp: not_le) show "Measurable.pred (F t) ?f" by (intro pred_intros_logic measurable_If measurable_pred_countable countable_Collect stopping_time_le_const predE stopping_time_eq_const T S) auto qed qed end lemma stopping_time_SUP_enat: fixes T :: "nat \ ('a \ enat)" shows "(\i. stopping_time F (T i)) \ stopping_time F (SUP i. T i)" unfolding stopping_time_def SUP_apply SUP_le_iff by (auto intro!: pred_intros_countable) lemma less_eSuc_iff: "a < eSuc b \ (a \ b \ a \ \)" by (cases a) auto lemma stopping_time_Inf_enat: fixes F :: "enat \ 'a measure" assumes F: "filtration \ F" assumes P: "\i. Measurable.pred (F i) (P i)" shows "stopping_time F (\\. Inf {i. P i \})" proof (rule stopping_timeI, cases) fix t :: enat assume "t = \" then show "Measurable.pred (F t) (\\. Inf {i. P i \} \ t)" by auto next fix t :: enat assume "t \ \" moreover { fix i \ assume "Inf {i. P i \} \ t" with \t \ \\ have "(\i\t. P i \)" unfolding Inf_le_iff by (cases t) (auto elim!: allE[of _ "eSuc t"] simp: less_eSuc_iff) } ultimately have *: "\\. Inf {i. P i \} \ t \ (\i\{..t}. P i \)" by (auto intro!: Inf_lower2) show "Measurable.pred (F t) (\\. Inf {i. P i \} \ t)" unfolding * using filtration.sets_F_mono[OF F, of _ t] P by (intro pred_intros_countable_bounded) (auto simp: pred_def filtration.space_F[OF F]) qed lemma stopping_time_Inf_nat: fixes F :: "nat \ 'a measure" assumes F: "filtration \ F" assumes P: "\i. Measurable.pred (F i) (P i)" and wf: "\i \. \ \ \ \ \n. P n \" shows "stopping_time F (\\. Inf {i. P i \})" unfolding stopping_time_def proof (intro allI, subst measurable_cong) fix t \ assume "\ \ space (F t)" then have "\ \ \" using filtration.space_F[OF F] by auto from wf[OF this] have "((LEAST n. P n \) \ t) = (\i\t. P i \)" by (rule LeastI2_wellorder_ex) auto then show "(Inf {i. P i \} \ t) = (\i\{..t}. P i \)" by (simp add: Inf_nat_def Bex_def) next fix t from P show "Measurable.pred (F t) (\w. \i\{..t}. P i w)" using filtration.sets_F_mono[OF F, of _ t] by (intro pred_intros_countable_bounded) (auto simp: pred_def filtration.space_F[OF F]) qed end