diff --git a/thys/Actuarial_Mathematics/Preliminaries.thy b/thys/Actuarial_Mathematics/Preliminaries.thy --- a/thys/Actuarial_Mathematics/Preliminaries.thy +++ b/thys/Actuarial_Mathematics/Preliminaries.thy @@ -1,2572 +1,2571 @@ theory Preliminaries imports "HOL-Library.Rewrite" "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Extended_Real" "HOL-Probability.Probability" begin notation powr (infixr ".^" 80) section \Preliminary Lemmas\ lemma Collect_conj_eq2: "{x \ A. P x \ Q x} = {x \ A. P x} \ {x \ A. Q x}" by blast lemma vimage_compl_atMost: fixes f :: "'a \ 'b::linorder" shows "-(f -` {..y}) = f -` {y<..}" by fastforce lemma Ico_nat_disjoint: "disjoint_family (\n::nat. {real n ..< real n + 1})" proof - { fix m n :: nat assume "{real m ..< real m + 1} \ {real n ..< real n + 1} \ {}" then obtain x::real where "x \ {real m ..< real m + 1} \ {real n ..< real n + 1}" by force hence "m = n" by simp } thus ?thesis unfolding disjoint_family_on_def by blast qed lemma Ico_nat_union: "(\n::nat. {real n ..< real n + 1}) = {0..}" proof show "(\n::nat. {real n ..< real n + 1}) \ {0..}" by auto next show "{0..} \ (\n::nat. {real n ..< real n + 1})" proof fix x::real assume "x \ {0..}" hence "nat \x\ \ x \ x < nat \x\ + 1" by (metis add_le_same_cancel1 atLeast_iff le_nat_floor less_one linorder_not_le of_nat_floor) thus "x \ (\n::nat. {real n ..< real n + 1})" unfolding atLeastLessThan_def by force qed qed lemma Ico_nat_union_finite: "(\(n::nat)(n::nat) {0.. (\(n::nat): "x \ {0..x\ < m" using of_nat_floor by fastforce moreover with \ have "nat \x\ \ x \ x < nat \x\ + 1" by (metis Nat.add_0_right atLeastLessThan_iff le_nat_floor less_one linorder_not_le nat_add_left_cancel_le of_nat_floor) ultimately show "x \ (\(n::nat) 0" defines "A \ \i::nat. {i*m ..< (i+1)*m}" shows "\i j. i \ j \ A i \ A j = {}" and "(\i j \ A i \ A j = {}" proof (erule contrapos_np) assume "A i \ A j \ {}" then obtain k where "k \ A i \ A j" by blast hence "i*m < (j+1)*m \ j*m < (i+1)*m" unfolding A_def by force hence "i < j+1 \ j < i+1" using mult_less_cancel2 by blast thus "i = j" by force qed } thus "\i j. i \ j \ A i \ A j = {}" by blast next show "(\ii {..< n*m}" proof fix x::nat assume "x \ (\i n" by linarith hence "x < n*m" by (meson less_le_trans mult_le_cancel2 i_x) thus "x \ {..< n*m}" using diff_mult_distrib mult_1 i_n by auto qed next show "{..< n*m} \ (\i {..< n*m}" hence "?i < n" by (simp add: less_mult_imp_div_less) moreover have "?i*m \ x \ x < (?i+1)*m" using assms div_times_less_eq_dividend dividend_less_div_times by auto ultimately show "x \ (\i b" for a b :: real unfolding frontier_def using that by force lemma(in field) divide_mult_cancel[simp]: fixes a b assumes "b \ 0" shows "a / b * b = a" by (simp add: assms) lemma inverse_powr: "(1/a).^b = a.^-b" if "a > 0" for a b :: real by (smt that powr_divide powr_minus_divide powr_one_eq_one) lemma powr_eq_one_iff_gen[simp]: "a.^x = 1 \ x = 0" if "a > 0" "a \ 1" for a x :: real by (metis powr_eq_0_iff powr_inj powr_zero_eq_one that) lemma powr_less_cancel2: "0 < a \ 0 < x \ 0 < y \ x.^a < y.^a \ x < y" for a x y ::real proof - assume a_pos: "0 < a" and x_pos: "0 < x" and y_pos: "0 < y" show "x.^a < y.^a \ x < y" proof (erule contrapos_pp) assume "\ x < y" hence "x \ y" by fastforce hence "x.^a \ y.^a" proof (cases "x = y") case True thus ?thesis by simp next case False hence "x.^a > y.^a" using \x \ y\ powr_less_mono2 a_pos y_pos by auto thus ?thesis by auto qed thus "\ x.^a < y.^a" by fastforce qed qed lemma geometric_increasing_sum_aux: "(1-r)^2 * (\kk 1" for n::nat and r::real by (subst geometric_increasing_sum_aux[THEN sym], simp add: that) lemma Reals_UNIV[simp]: "\ = {x::real. True}" unfolding Reals_def by auto lemma Lim_cong: assumes "\\<^sub>F x in F. f x = g x" shows "Lim F f = Lim F g" unfolding t2_space_class.Lim_def using tendsto_cong assms by fastforce lemma antimono_onI: "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ antimono_on A f" by (rule monotone_onI) lemma antimono_onD: "\antimono_on A f; r \ A; s \ A; r \ s\ \ f r \ f s" by (rule monotone_onD) lemma antimono_imp_mono_on: "antimono f \ antimono_on A f" by (simp add: antimonoD antimono_onI) lemma antimono_on_subset: "antimono_on A f \ B \ A \ antimono_on B f" by (rule monotone_on_subset) lemma mono_on_antimono_on: fixes f :: "'a::order \ 'b::ordered_ab_group_add" shows "mono_on A f \ antimono_on A (\r. - f r)" by (simp add: monotone_on_def) corollary mono_antimono: fixes f :: "'a::order \ 'b::ordered_ab_group_add" shows "mono f \ antimono (\r. - f r)" by (rule mono_on_antimono_on) lemma mono_on_at_top_le: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_mono: "mono_on {a..} f" and f_to_l: "(f \ l) at_top" shows "\x. x \ {a..} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "f b - l" have lim_top: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_top" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_top" using fb_l by (intro lim_top; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_top_linorder by metis let ?n = "max b N" have "f ?n < f b" using fn_in by force moreover have "f ?n \ f b" using f_mono b_a by (simp add: le_max_iff_disj mono_on_def) ultimately have False by simp } thus "\x\{a..}. f x \ l" apply - by (rule notnotD, rewrite Set.ball_simps) auto qed corollary mono_at_top_le: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "mono f" and "(f \ b) at_top" shows "\x. f x \ b" using mono_on_at_top_le assms by (metis atLeast_iff mono_imp_mono_on nle_le) lemma mono_on_at_bot_ge: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_mono: "mono_on {..a} f" and f_to_l: "(f \ l) at_bot" shows "\x. x \ {..a} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "l - f b" have lim_bot: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_bot" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_bot" using fb_l by (intro lim_bot; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_bot_linorder by metis let ?n = "min b N" have "f ?n > f b" using fn_in by force moreover have "f ?n \ f b" using f_mono b_a by (simp add: min.coboundedI1 mono_onD) ultimately have False by simp } thus "\x\{..a}. f x \ l" apply - by (rule notnotD, rewrite Set.ball_simps) auto qed corollary mono_at_bot_ge: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "mono f" and "(f \ b) at_bot" shows "\x. f x \ b" using mono_on_at_bot_ge assms by (metis atMost_iff mono_imp_mono_on nle_le) lemma antimono_on_at_top_ge: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_antimono: "antimono_on {a..} f" and f_to_l: "(f \ l) at_top" shows "\x. x \ {a..} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "l - f b" have lim_top: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_top" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_top" using fb_l by (intro lim_top; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_top_linorder by metis let ?n = "max b N" have "f ?n > f b" using fn_in by force moreover have "f ?n \ f b" using f_antimono b_a by (simp add: antimono_onD le_max_iff_disj) ultimately have False by simp } thus "\x\{a..}. f x \ l" apply - by (rule notnotD, rewrite Set.ball_simps) auto qed corollary antimono_at_top_le: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "antimono f" and "(f \ b) at_top" shows "\x. f x \ b" using antimono_on_at_top_ge assms antimono_imp_mono_on by blast lemma antimono_on_at_bot_ge: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_antimono: "antimono_on {..a} f" and f_to_l: "(f \ l) at_bot" shows "\x. x \ {..a} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "f b - l" have lim_bot: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_bot" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_bot" using fb_l by (intro lim_bot; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_bot_linorder by metis let ?n = "min b N" have "f ?n < f b" using fn_in by force moreover have "f ?n \ f b" using f_antimono b_a by (simp add: min.coboundedI1 antimono_onD) ultimately have False by simp } thus "\x\{..a}. f x \ l" apply - by (rule notnotD, rewrite Set.ball_simps) auto qed corollary antimono_at_bot_ge: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "antimono f" and "(f \ b) at_bot" shows "\x. f x \ b" using antimono_on_at_bot_ge assms antimono_imp_mono_on by blast lemma continuous_cdivide: fixes c::"'a::real_normed_field" assumes "c \ 0" "continuous F f" shows "continuous F (\x. f x / c)" using assms continuous_mult_right by (rewrite field_class.field_divide_inverse) auto lemma continuous_mult_left_iff: fixes c::"'a::real_normed_field" assumes "c \ 0" shows "continuous F f \ continuous F (\x. c * f x)" using continuous_mult_left continuous_cdivide assms by force lemma continuous_mult_right_iff: fixes c::"'a::real_normed_field" assumes "c \ 0" shows "continuous F f \ continuous F (\x. f x * c)" using continuous_mult_right continuous_cdivide assms by force lemma continuous_cdivide_iff: fixes c::"'a::real_normed_field" assumes "c \ 0" shows "continuous F f \ continuous F (\x. f x / c)" proof assume "continuous F f" thus "continuous F (\x. f x / c)" by (intro continuous_cdivide) (simp add: assms) next assume "continuous F (\x. f x / c)" hence "continuous F (\x. f x / c * c)" using continuous_mult_right by fastforce thus "continuous F f" using assms by force qed lemma continuous_cong: assumes "eventually (\x. f x = g x) F" "f (Lim F (\x. x)) = g (Lim F (\x. x))" shows "continuous F f \ continuous F g" unfolding continuous_def using assms filterlim_cong by force lemma continuous_at_within_cong: assumes "f x = g x" "eventually (\x. f x = g x) (at x within s)" shows "continuous (at x within s) f \ continuous (at x within s) g" proof (cases \x \ closure (s - {x})\) case True thus ?thesis using assms apply (intro continuous_cong, simp) by (rewrite Lim_ident_at, simp add: at_within_eq_bot_iff)+ simp next case False hence "at x within s = bot" using not_trivial_limit_within by blast thus ?thesis by simp qed lemma DERIV_cmult_iff: assumes "c \ 0" shows "(f has_field_derivative D) (at x within s) \ ((\x. c * f x) has_field_derivative c * D) (at x within s)" proof assume "(f has_field_derivative D) (at x within s)" thus "((\x. c * f x) has_field_derivative c * D) (at x within s)" using DERIV_cmult by force next assume "((\x. c * f x) has_field_derivative c * D) (at x within s)" hence "((\x. c * f x / c) has_field_derivative c * D / c) (at x within s)" using DERIV_cdivide assms by blast thus "(f has_field_derivative D) (at x within s)" by (simp add: assms field_simps) qed lemma DERIV_cmult_right_iff: assumes "c \ 0" shows "(f has_field_derivative D) (at x within s) \ ((\x. f x * c) has_field_derivative D * c) (at x within s)" by (rewrite DERIV_cmult_iff[of c], simp_all add: assms mult_ac) lemma DERIV_cdivide_iff: assumes "c \ 0" shows "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" apply (rewrite field_class.field_divide_inverse)+ using DERIV_cmult_right_iff assms inverse_nonzero_iff_nonzero by blast lemma DERIV_ln_divide_chain: fixes f :: "real \ real" assumes "f x > 0" and "(f has_real_derivative D) (at x within s)" shows "((\x. ln (f x)) has_real_derivative (D / f x)) (at x within s)" proof - have "DERIV ln (f x) :> 1 / f x" using assms by (intro DERIV_ln_divide) blast thus ?thesis using DERIV_chain2 assms by fastforce qed lemma inverse_fun_has_integral_ln: fixes f :: "real \ real" and f' :: "real \ real" assumes "a \ b" and "\x. x\{a..b} \ f x > 0" and "continuous_on {a..b} f" and "\x. x\{a<.. (f has_real_derivative f' x) (at x)" shows "((\x. f' x / f x) has_integral (ln (f b) - ln (f a))) {a..b}" proof - have "continuous_on {a..b} (\x. ln (f x))" using assms by (intro continuous_intros; force) moreover have "\x. x\{a<.. ((\x. ln (f x)) has_vector_derivative f' x / f x) (at x)" apply (rewrite has_real_derivative_iff_has_vector_derivative[THEN sym]) using assms by (intro DERIV_ln_divide_chain; simp) ultimately show ?thesis using assms by (intro fundamental_theorem_of_calculus_interior; simp) qed lemma DERIV_fun_powr2: fixes a::real assumes a_pos: "a > 0" and f: "DERIV f x :> r" shows "DERIV (\x. a.^(f x)) x :> a.^(f x) * r * ln a" proof - let ?g = "(\x. a)" have g: "DERIV ?g x :> 0" by simp have pos: "?g x > 0" by (simp add: a_pos) show ?thesis using DERIV_powr[OF g pos f] a_pos by (auto simp add: field_simps) qed lemma has_real_derivative_powr2: assumes a_pos: "a > 0" shows "((\x. a.^x) has_real_derivative a.^x * ln a) (at x)" proof - let ?f = "(\x. x::real)" have f: "DERIV ?f x :> 1" by simp thus ?thesis using DERIV_fun_powr2[OF a_pos f] by simp qed lemma has_integral_powr2_from_0: fixes a c :: real assumes a_pos: "a > 0" and a_neq_1: "a \ 1" and c_nneg: "c \ 0" shows "((\x. a.^x) has_integral ((a.^c - 1) / (ln a))) {0..c}" proof - have "((\x. a.^x) has_integral ((a.^c)/(ln a) - (a.^0)/(ln a))) {0..c}" proof (rule fundamental_theorem_of_calculus[OF c_nneg]) fix x::real assume "x \ {0..c}" show "((\y. a.^y / ln a) has_vector_derivative a.^x) (at x within {0..c})" apply (insert has_real_derivative_powr2[OF a_pos, of x]) apply (drule DERIV_cdivide[where c = "ln a"], simp add: assms) apply (rule has_vector_derivative_within_subset[where S=UNIV and T="{0..c}"], auto) by (rule iffD1[OF has_real_derivative_iff_has_vector_derivative]) qed thus ?thesis using assms powr_zero_eq_one by (simp add: field_simps) qed lemma integrable_on_powr2_from_0: fixes a c :: real assumes a_pos: "a > 0" and a_neq_1: "a \ 1" and c_nneg: "c \ 0" shows "(\x. a.^x) integrable_on {0..c}" using has_integral_powr2_from_0[OF assms] unfolding integrable_on_def by blast lemma integrable_on_powr2_from_0_general: fixes a c :: real assumes a_pos: "a > 0" and c_nneg: "c \ 0" shows "(\x. a.^x) integrable_on {0..c}" proof (cases "a = 1") case True thus ?thesis using has_integral_const_real by auto next case False thus ?thesis using has_integral_powr2_from_0 False assms by auto qed lemma has_integral_null_interval: fixes a b :: real and f::"real \ real" assumes "a \ b" shows "(f has_integral 0) {a..b}" using assms content_real_eq_0 by blast lemma has_integral_interval_reverse: fixes f :: "real \ real" and a b :: real assumes "a \ b" and "continuous_on {a..b} f" shows "((\x. f (a+b-x)) has_integral (integral {a..b} f)) {a..b}" proof - let ?g = "\x. a + b - x" let ?g' = "\x. -1" have g_C0: "continuous_on {a..b} ?g" using continuous_on_op_minus by simp have Dg_g': "\x. x\{a..b} \ (?g has_field_derivative ?g' x) (at x within {a..b})" by (auto intro!: derivative_eq_intros) show ?thesis using has_integral_substitution_general [of "{}" a b ?g a b f, simplified, OF assms g_C0 Dg_g', simplified] apply (simp add: has_integral_null_interval[OF assms(1), THEN integral_unique]) by (simp add: has_integral_neg_iff) qed section \Additional Lemmas for the "Analysis" Library\ lemma continuous_within_shift: fixes a x :: "'a :: {topological_ab_group_add, t2_space}" and s :: "'a set" and f :: "'a \ 'b::topological_space" shows "continuous (at x within s) (\x. f (x+a)) \ continuous (at (x+a) within plus a ` s) f" proof assume "continuous (at x within s) (\x. f (x+a))" moreover have "continuous (at (x+a) within plus a ` s) (plus (-a))" by (simp add: continuous_at_imp_continuous_at_within) moreover have "plus (-a) ` plus a ` s = s" by force ultimately show "continuous (at (x+a) within plus a ` s) f" using continuous_within_compose unfolding comp_def by force next assume "continuous (at (x+a) within plus a ` s) f" moreover have "continuous (at x within s) (plus a)" by (simp add: continuous_at_imp_continuous_at_within) ultimately show "continuous (at x within s) (\x. f (x+a))" using continuous_within_compose unfolding comp_def by (force simp add: add.commute) qed lemma isCont_shift: fixes a x :: "'a :: {topological_ab_group_add, t2_space}" and f :: "'a \ 'b::topological_space" shows "isCont (\x. f (x+a)) x \ isCont f (x+a)" using continuous_within_shift by force lemma differentiable_eq_field_differentiable_real: fixes f :: "real \ real" shows "f differentiable F \ f field_differentiable F" unfolding field_differentiable_def differentiable_def has_real_derivative using has_real_derivative_iff by presburger lemma differentiable_on_eq_field_differentiable_real: fixes f :: "real \ real" shows "f differentiable_on s \ (\x\s. f field_differentiable (at x within s))" unfolding differentiable_on_def using differentiable_eq_field_differentiable_real by simp lemma differentiable_on_cong : assumes "\x. x\s \ f x = g x" and "f differentiable_on s" shows "g differentiable_on s" proof - { fix x assume "x\s" hence "f differentiable at x within s" using assms unfolding differentiable_on_def by simp from this obtain D where "(f has_derivative D) (at x within s)" unfolding differentiable_def by blast hence "(g has_derivative D) (at x within s)" using has_derivative_transform assms \x\s\ by metis hence "g differentiable at x within s" unfolding differentiable_def by blast } hence "\x\s. g differentiable at x within s" by simp thus ?thesis unfolding differentiable_on_def by simp qed lemma C1_differentiable_imp_deriv_continuous_on: "f C1_differentiable_on S \ continuous_on S (deriv f)" using C1_differentiable_on_eq field_derivative_eq_vector_derivative by auto lemma deriv_shift: assumes "f field_differentiable at (x+a)" shows "deriv f (x+a) = deriv (\s. f (x+s)) a" proof - have "(f has_field_derivative deriv f (x+a)) (at (x+a))" using DERIV_deriv_iff_field_differentiable assms by force hence "((\s. f (x+s)) has_field_derivative deriv f (x+a)) (at a)" using DERIV_at_within_shift has_field_derivative_at_within by blast moreover have "((\s. f (x+s)) has_field_derivative deriv (\s. f (x+s)) a) (at a)" using DERIV_imp_deriv calculation by fastforce ultimately show ?thesis using DERIV_unique by force qed lemma piecewise_differentiable_on_cong: assumes "f piecewise_differentiable_on i" and "\x. x \ i \ f x = g x" shows "g piecewise_differentiable_on i" proof - have "continuous_on i g" using continuous_on_cong_simp assms piecewise_differentiable_on_imp_continuous_on by force moreover have "\S. finite S \ (\x \ i - S. g differentiable (at x within i))" proof - from assms piecewise_differentiable_on_def obtain S where fin: "finite S" and "\x \ i - S. f differentiable (at x within i)" by metis hence "\x. x \ i - S \ f differentiable (at x within i)" by simp hence "\x. x \ i - S \ g differentiable (at x within i)" using has_derivative_transform assms by (metis DiffD1 differentiable_def) with fin show ?thesis by blast qed ultimately show ?thesis unfolding piecewise_differentiable_on_def by simp qed lemma differentiable_piecewise: assumes "continuous_on i f" and "f differentiable_on i" shows "f piecewise_differentiable_on i" unfolding piecewise_differentiable_on_def using assms differentiable_onD by auto lemma piecewise_differentiable_scaleR: assumes "f piecewise_differentiable_on S" shows "(\x. a *\<^sub>R f x) piecewise_differentiable_on S" proof - from assms obtain T where fin: "finite T" "\x. x \ S - T \ f differentiable at x within S" unfolding piecewise_differentiable_on_def by blast hence "\x. x \ S - T \ (\x. a *\<^sub>R f x) differentiable at x within S" using differentiable_scaleR by simp moreover have "continuous_on S (\x. a *\<^sub>R f x)" using assms continuous_on_scaleR continuous_on_const piecewise_differentiable_on_def by blast ultimately show "(\x. a *\<^sub>R f x) piecewise_differentiable_on S" using fin piecewise_differentiable_on_def by blast qed lemma differentiable_on_piecewise_compose: assumes "f piecewise_differentiable_on S" and "g differentiable_on f ` S" shows "g \ f piecewise_differentiable_on S" proof - from assms obtain T where fin: "finite T" "\x. x \ S - T \ f differentiable at x within S" unfolding piecewise_differentiable_on_def by blast hence "\x. x \ S - T \ g \ f differentiable at x within S" by (meson DiffD1 assms differentiable_chain_within differentiable_onD image_eqI) hence "\T. finite T \ (\x\S-T. g \ f differentiable at x within S)" using fin by blast moreover have "continuous_on S (g \ f)" using assms continuous_on_compose differentiable_imp_continuous_on unfolding piecewise_differentiable_on_def by blast ultimately show ?thesis unfolding piecewise_differentiable_on_def by force qed lemma MVT_order_free: fixes r h :: real defines "I \ {r..r+h} \ {r+h..r}" assumes "continuous_on I f" and "f differentiable_on interior I" obtains t where "t \ {0<..<1}" and "f (r+h) - f r = h * deriv f (r + t*h)" proof - consider (h_pos) "h > 0" | (h_0) "h = 0" | (h_neg) "h < 0" by force thus ?thesis proof cases case h_pos hence "r < r+h" "I = {r..r+h}" unfolding I_def by simp_all moreover hence "interior I = {r<..x. \r < x; x < r+h\ \ f differentiable (at x)" using assms by (simp add: differentiable_on_eq_differentiable_at) ultimately obtain z where "r < z \ z < r+h \ f (r+h) - f r = h * deriv f z" using MVT assms by (smt (verit) DERIV_imp_deriv) moreover hence "(z-r) / h \ {0<..<1}" by simp moreover have "z = r + (z-r)/h * h" using h_pos by simp ultimately show ?thesis using that by presburger next case h_0 have "1/2 \ {0::real<..<1}" by simp moreover have "f (r+h) - f r = 0" using h_0 by simp moreover have "h * deriv f (r + (1/2)*h) = 0" using h_0 by simp ultimately show ?thesis using that by presburger next case h_neg hence "r+h < r" "I = {r+h..r}" unfolding I_def by simp_all moreover hence "interior I = {r+h<..x. \r+h < x; x < r\ \ f differentiable (at x)" using assms by (simp add: differentiable_on_eq_differentiable_at) ultimately obtain z where "r+h < z \ z < r \ f r - f (r+h) = -h * deriv f z" using MVT assms by (smt (verit) DERIV_imp_deriv) moreover hence "(z-r) / h \ {0<..<1}" by (simp add: neg_less_divide_eq) moreover have "z = r + (z-r)/h * h" using h_neg by simp ultimately show ?thesis using that mult_minus_left by fastforce qed qed lemma integral_combine2: fixes f :: "real \ 'a::banach" assumes "a \ c" "c \ b" and "f integrable_on {a..c}" "f integrable_on {c..b}" shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" apply (rule integral_unique[THEN sym]) apply (rule has_integral_combine[of a c b], simp_all add: assms) using has_integral_integral assms by auto lemma FTC_real_deriv_has_integral: fixes F :: "real \ real" assumes "a \ b" and "F piecewise_differentiable_on {a<..x. x \ {a<.. F differentiable at x within {a<..x. x \ {a<.. (F has_real_derivative deriv F x) (at x)" proof - fix x assume x_in: "x \ {a<..x. x \ T - S \ g x = f x" shows "f integrable_on T \ g integrable_on T" using integrable_spike assms by force lemma set_borel_measurable_lborel: "set_borel_measurable lborel A f \ set_borel_measurable borel A f" unfolding set_borel_measurable_def by simp lemma restrict_space_whole[simp]: "restrict_space M (space M) = M" unfolding restrict_space_def by (simp add: measure_of_of_measure) lemma deriv_measurable_real: fixes f :: "real \ real" assumes "f differentiable_on S" "open S" "f \ borel_measurable borel" shows "set_borel_measurable borel S (deriv f)" proof - have "\x. x \ S \ deriv f x = lim (\i. (f (x + 1 / Suc i) - f x) / (1 / Suc i))" proof - fix x assume x_S: "x \ S" hence "f field_differentiable (at x within S)" using differentiable_on_eq_field_differentiable_real assms by simp hence "(f has_field_derivative deriv f x) (at x)" using assms DERIV_deriv_iff_field_differentiable x_S at_within_open by force hence "(\h. (f (x+h) - f x) / h) \0\ deriv f x" using DERIV_def by auto hence "\d. (\i. d i \ UNIV-{0::real}) \ d \ 0 \ ((\h. (f (x+h) - f x) / h) \ d) \ deriv f x" using tendsto_at_iff_sequentially by blast moreover have "\i. 1 / Suc i \ UNIV - {0::real}" by simp moreover have "(\i. 1 / Suc i) \ 0" using LIMSEQ_Suc lim_const_over_n by blast ultimately have "((\h. (f (x + h) - f x) / h) \ (\i. 1 / Suc i)) \ deriv f x" by auto thus "deriv f x = lim (\i. (f (x + 1 / Suc i) - f x) / (1 / Suc i))" unfolding comp_def by (simp add: limI) qed moreover have "(\x. indicator S x * lim (\i. (f (x + 1 / Suc i) - f x) / (1 / Suc i))) \ borel_measurable borel" using assms by (measurable, simp, measurable) ultimately show ?thesis unfolding set_borel_measurable_def measurable_cong by simp (smt (verit) indicator_simps(2) measurable_cong mult_eq_0_iff) qed lemma piecewise_differentiable_on_deriv_measurable_real: fixes f :: "real \ real" assumes "f piecewise_differentiable_on S" "open S" "f \ borel_measurable borel" shows "set_borel_measurable borel S (deriv f)" proof - from assms obtain T where fin: "finite T" and diff: "\x. x \ S - T \ f differentiable at x within S" unfolding piecewise_differentiable_on_def by blast with assms have "open (S - T)" using finite_imp_closed by blast moreover hence "f differentiable_on (S - T)" unfolding differentiable_on_def using assms by (metis Diff_iff at_within_open diff) ultimately have "set_borel_measurable borel (S - T) (deriv f)" by (intro deriv_measurable_real; simp add: assms) thus ?thesis unfolding set_borel_measurable_def apply simp apply (rule measurable_discrete_difference [where X=T and f="\x. indicat_real (S - T) x * deriv f x"], simp_all) using fin uncountable_infinite apply blast by (simp add: indicator_diff) qed lemma borel_measurable_antimono: fixes f :: "real \ real" shows "antimono f \ f \ borel_measurable borel" using borel_measurable_mono by (smt (verit, del_insts) borel_measurable_uminus_eq monotone_on_def) lemma set_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 \) \ set_borel_measurable M \ f" using assms borel_measurable_restrict_space_iff set_borel_measurable_def by blast lemma set_integrable_restrict_space_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "A \ sets M" shows "set_integrable M A f \ integrable (restrict_space M A) f" unfolding set_integrable_def using assms by (rewrite integrable_restrict_space; simp) lemma set_lebesgue_integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "A \ sets M" shows "set_lebesgue_integral M A f = integral\<^sup>L (restrict_space M A) f" unfolding set_lebesgue_integral_def using assms integral_restrict_space by (metis (mono_tags) sets.Int_space_eq2) lemma distr_borel_lborel: "distr M borel f = distr M lborel f" by (metis distr_cong sets_lborel) lemma AE_translation: assumes "AE x in lborel. P x" shows "AE x in lborel. P (a+x)" proof - from assms obtain N where P: "\x. x \ space lborel - N \ P x" and null: "N \ null_sets lborel" using AE_E3 by blast hence "{y. a+y \ N} \ null_sets lborel" using null_sets_translation[of N "-a", simplified] by (simp add: add.commute) moreover have "\y. y \ space lborel - {y. a+y \ N} \ P (a+y)" using P by force ultimately show "AE y in lborel. P (a+y)" by (smt (verit, del_insts) Diff_iff eventually_ae_filter mem_Collect_eq subsetI) qed lemma set_AE_translation: assumes "AE x\S in lborel. P x" shows "AE x \ plus (-a) ` S in lborel. P (a+x)" proof - have "AE x in lborel. a+x \ S \ P (a+x)" using assms by (rule AE_translation) moreover have "\x. a+x \ S \ x \ plus (-a) ` S" by force ultimately show ?thesis by simp qed lemma AE_scale_measure_iff: assumes "r > 0" shows "(AE x in (scale_measure r M). P x) \ (AE x in M. P x)" unfolding ae_filter_def null_sets_def apply (rewrite space_scale_measure, simp) using assms by (smt Collect_cong not_gr_zero) lemma nn_set_integral_cong2: assumes "AE x\A in M. f x = g x" shows "(\\<^sup>+x\A. f x \M) = (\\<^sup>+x\A. g x \M)" proof - { fix x assume "x \ space M" have "(x \ A \ f x = g x) \ f x * indicator A x = g x * indicator A x" by force } hence "AE x in M. (x \ A \ f x = g x) \ f x * indicator A x = g x * indicator A x" by (rule AE_I2) hence "AE x in M. f x * indicator A x = g x * indicator A x" using assms AE_mp by auto thus ?thesis by (rule nn_integral_cong_AE) qed lemma set_lebesgue_integral_cong_AE2: assumes [measurable]: "A \ sets M" "set_borel_measurable M A f" "set_borel_measurable M A g" assumes "AE x \ A in M. f x = g x" shows "LINT x:A|M. f x = LINT x:A|M. g x" proof - let ?fA = "\x. indicator A x *\<^sub>R f x" and ?gA = "\x. indicator A x *\<^sub>R g x" have "?fA \ borel_measurable M" "?gA \ borel_measurable M" using assms unfolding set_borel_measurable_def by simp_all moreover have "AE x \ A in M. ?fA x = ?gA x" using assms by simp ultimately have "LINT x:A|M. ?fA x = LINT x:A|M. ?gA x" by (intro set_lebesgue_integral_cong_AE; simp) moreover have "LINT x:A|M. f x = LINT x:A|M. ?fA x" "LINT x:A|M. g x = LINT x:A|M. ?gA x" unfolding set_lebesgue_integral_def by (metis indicator_scaleR_eq_if)+ ultimately show ?thesis by simp qed proposition nn_integral_disjoint_family_on_finite: assumes [measurable]: "f \ borel_measurable M" "\(n::nat). n \ S \ B n \ sets M" and "disjoint_family_on B S" "finite S" shows "(\\<^sup>+x \ (\n\S. B n). f x \M) = (\n\S. (\\<^sup>+x \ B n. f x \M))" proof - let ?A = "\n::nat. if n \ S then B n else {}" have "\n::nat. ?A n \ sets M" by simp moreover have "disjoint_family ?A" unfolding disjoint_family_on_def proof - { fix m n :: nat assume "m \ n" hence "(if m \ S then B m else {}) \ (if n \ S then B n else {}) = {}" apply simp using assms unfolding disjoint_family_on_def by blast } thus "\m::nat\UNIV. \n::nat\UNIV. m \ n \ (if m \ S then B m else {}) \ (if n \ S then B n else {}) = {}" by blast qed ultimately have "set_nn_integral M (\ (range ?A)) f = (\n. set_nn_integral M (?A n) f)" by (rewrite nn_integral_disjoint_family; simp) moreover have "set_nn_integral M (\ (range ?A)) f = (\\<^sup>+x \ (\n\S. B n). f x \M)" proof - have "\ (range ?A) = (\n\S. B n)" by simp thus ?thesis by simp qed moreover have "(\n. set_nn_integral M (?A n) f) = (\n\S. set_nn_integral M (B n) f)" by (rewrite suminf_finite[of S]; simp add: assms) ultimately show ?thesis by simp qed lemma nn_integral_distr_set: assumes "T \ measurable M M'" and "f \ borel_measurable (distr M M' T)" and "A \ sets M'" and "\x. x \ space M \ T x \ A" shows "integral\<^sup>N (distr M M' T) f = set_nn_integral (distr M M' T) A f" proof - have "integral\<^sup>N (distr M M' T) f = (\\<^sup>+x\(space M'). f x \(distr M M' T))" by (rewrite nn_set_integral_space[THEN sym], simp) also have "\ = (\\<^sup>+x\A. f x \(distr M M' T))" proof - have [simp]: "sym_diff (space M') A = space M' - A" using assms by (metis Diff_mono sets.sets_into_space sup.orderE) show ?thesis apply (rule nn_integral_null_delta; simp add: assms) unfolding null_sets_def using assms apply (simp, rewrite emeasure_distr; simp) unfolding vimage_def using emeasure_empty by (smt (z3) Collect_empty_eq Diff_iff Int_def mem_Collect_eq) qed finally show ?thesis . qed (* Analogue for "measure_eqI_lessThan" in the "Lebesgue_Measure" Theory *) lemma measure_eqI_Ioc: fixes M N :: "real measure" assumes sets: "sets M = sets borel" "sets N = borel" assumes fin: "\a b. a \ b \ emeasure M {a<..b} < \" assumes eq: "\a b. a \ b \ emeasure M {a<..b} = emeasure N {a<..b}" shows "M = N" proof (rule measure_eqI_generator_eq_countable) let ?Ioc = "\(a::real,b::real). {a<..b}" let ?E = "range ?Ioc" show "Int_stable ?E" using Int_stable_def Int_greaterThanAtMost by force show "?E \ Pow UNIV" "sets M = sigma_sets UNIV ?E" "sets N = sigma_sets UNIV ?E" unfolding sets by (auto simp add: borel_sigma_sets_Ioc) show "\I. I \ ?E \ emeasure M I = emeasure N I" proof - fix I assume "I \ ?E" then obtain a b where "I = {a<..b}" by auto thus "emeasure M I = emeasure N I" by (smt (verit, best) eq greaterThanAtMost_empty) qed show "?Ioc ` (Rats \ Rats) \ ?E" "(\i\(Rats\Rats). ?Ioc i) = UNIV" using Rats_no_bot_less Rats_no_top_le by auto show "countable (?Ioc ` (Rats \ Rats))" using countable_rat by blast show "\I. I \ ?Ioc ` (Rats \ Rats) \ emeasure M I \ \" proof - fix I assume "I \ ?Ioc ` (Rats \ Rats)" then obtain a b where "(a,b) \ (Rats \ Rats)" "I = {a<..b}" by blast thus "emeasure M I \ \" by (smt (verit, best) Ioc_inj fin order.strict_implies_not_eq) qed qed lemma (in finite_measure) distributed_measure: assumes "distributed M N X f" and "\x. x \ space N \ f x \ 0" and "A \ sets N" shows "measure M (X -` A \ space M) = (\x. indicator A x * f x \N)" proof - have [simp]: "(\x. indicat_real A x * f x) \ borel_measurable N" using assms apply (measurable; simp?) using distributed_real_measurable assms by force have "emeasure M (X -` A \ space M) = (\\<^sup>+x\A. ennreal (f x) \N)" by (rule distributed_emeasure; simp add: assms) moreover have "enn2real (\\<^sup>+x\A. ennreal (f x) \N) = \x. indicator A x * f x \N" apply (rewrite enn2real_nn_integral_eq_integral [where f="\x. ennreal (indicator A x * f x)", THEN sym]; (simp add: assms)?) using distributed_emeasure assms - apply (smt (verit) emeasure_finite indicator_mult_ennreal mult.commute + by (smt (verit) emeasure_finite indicator_mult_ennreal mult.commute nn_integral_cong top.not_eq_extremum) - by (rewrite nn_integral_set_ennreal, rewrite mult.commute) simp ultimately show ?thesis using measure_def by metis qed (* Set Integral Version of the Lebesgue's Dominated Convergence Theorem *) lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" assumes "A \ sets M" "set_borel_measurable M A f" "\i. set_borel_measurable M A (s i)" "set_integrable M A w" assumes lim: "AE x\A in M. (\i. s i x) \ f x" assumes bound: "\i::nat. AE x\A in M. norm (s i x) \ w x" shows set_integrable_dominated_convergence: "set_integrable M A f" and set_integrable_dominated_convergence2: "\i. set_integrable M A (s i)" and set_integral_dominated_convergence: "(\i. set_lebesgue_integral M A (s i)) \ set_lebesgue_integral M A f" proof - have "(\x. indicator A x *\<^sub>R f x) \ borel_measurable M" and "\i. (\x. indicator A x *\<^sub>R s i x) \ borel_measurable M" and "integrable M (\x. indicator A x *\<^sub>R w x)" using assms unfolding set_borel_measurable_def set_integrable_def by simp_all moreover have "AE x in M. (\i. indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R f x" proof - obtain N where N_null: "N \ null_sets M" and si_f: "\x. x \ space M - N \ x \ A \ (\i. s i x) \ f x" using lim AE_E3 by (smt (verit)) hence "\x. x \ space M - N \ (\i. indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R f x" proof - fix x assume asm: "x \ space M - N" thus "(\i. indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R f x" proof (cases \x \ A\) case True with si_f asm show ?thesis by simp next case False thus ?thesis by simp qed qed thus ?thesis by (smt (verit) AE_I' DiffI N_null mem_Collect_eq subsetI) qed moreover have "\i. AE x in M. norm (indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R w x" proof - fix i from bound obtain N where N_null: "N \ null_sets M" and "\x. x \ space M - N \ x \ A \ norm (s i x) \ w x" using AE_E3 by (smt (verit)) hence "\x. x \ space M - N \ norm (indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R w x" by (simp add: indicator_scaleR_eq_if) with N_null show "AE x in M. norm (indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R w x" by (smt (verit) DiffI eventually_ae_filter mem_Collect_eq subsetI) qed ultimately show "set_integrable M A f" "\i. set_integrable M A (s i)" "(\i. set_lebesgue_integral M A (s i)) \ set_lebesgue_integral M A f" unfolding set_integrable_def set_lebesgue_integral_def by (rule integrable_dominated_convergence, rule integrable_dominated_convergence2, rule integral_dominated_convergence) qed lemma absolutely_integrable_on_iff_set_integrable: fixes f :: "'a::euclidean_space \ real" assumes "f \ borel_measurable lborel" and "S \ sets lborel" shows "set_integrable lborel S f \ f absolutely_integrable_on S" unfolding set_integrable_def apply (simp, rewrite integrable_completion[THEN sym]) apply measurable using assms by simp_all corollary integrable_on_iff_set_integrable_nonneg: fixes f :: "'a::euclidean_space \ real" assumes "\x. x \ S \ f x \ 0" "f \ borel_measurable lborel" and "S \ sets lborel" shows "set_integrable lborel S f \ f integrable_on S" using absolutely_integrable_on_iff_set_integrable assms by (metis absolutely_integrable_on_iff_nonneg) lemma integrable_on_iff_set_integrable_nonneg_AE: fixes f :: "'a::euclidean_space \ real" assumes "AE x\S in lborel. f x \ 0" "f \ borel_measurable lborel" and "S \ sets lborel" shows "set_integrable lborel S f \ f integrable_on S" proof - from assms obtain N where nonneg: "\x. x \ S - N \ f x \ 0" and null: "N \ null_sets lborel" by (smt (verit, ccfv_threshold) AE_E3 Diff_iff UNIV_I space_borel space_lborel) let ?g = "\x. if x \ N then 0 else f x" have [simp]: "negligible N" using null negligible_iff_null_sets null_sets_completionI by blast have "N \ sets lborel" using null by auto hence [simp]: "?g \ borel_measurable borel" using assms by force have "set_integrable lborel S f \ set_integrable lborel S ?g" proof - have "AE x\S in lborel. f x = ?g x" by (rule AE_I'[of N], simp_all add: null, blast) thus ?thesis using assms by (intro set_integrable_cong_AE[of f _ ?g S]; simp) qed also have "\ \ ?g integrable_on S" using assms by (intro integrable_on_iff_set_integrable_nonneg; simp add: nonneg) also have "\ \ f integrable_on S" by (rule integrable_spike_cong[of N]; simp) finally show ?thesis . qed lemma set_borel_integral_eq_integral_nonneg: fixes f :: "'a::euclidean_space \ real" assumes "\x. x \ S \ f x \ 0" "f \ borel_measurable borel" "S \ sets borel" shows "LINT x : S | lborel. f x = integral S f" \ \Note that 0 = 0 holds when the integral diverges.\ proof (cases \set_integrable lborel S f\) case True thus ?thesis using set_borel_integral_eq_integral by force next case False hence "LINT x : S | lborel. f x = 0" unfolding set_lebesgue_integral_def set_integrable_def by (rewrite not_integrable_integral_eq; simp) moreover have "integral S f = 0" apply (rule not_integrable_integral) using False assms by (rewrite in asm integrable_on_iff_set_integrable_nonneg; simp) ultimately show ?thesis .. qed lemma set_borel_integral_eq_integral_nonneg_AE: fixes f :: "'a::euclidean_space \ real" assumes "AE x\S in lborel. f x \ 0" "f \ borel_measurable borel" "S \ sets borel" shows "LINT x : S | lborel. f x = integral S f" \ \Note that 0 = 0 holds when the integral diverges.\ proof (cases \set_integrable lborel S f\) case True thus ?thesis using set_borel_integral_eq_integral by force next case False hence "LINT x : S | lborel. f x = 0" unfolding set_lebesgue_integral_def set_integrable_def by (rewrite not_integrable_integral_eq; simp) moreover have "integral S f = 0" apply (rule not_integrable_integral) using False assms by (rewrite in asm integrable_on_iff_set_integrable_nonneg_AE; simp) ultimately show ?thesis .. qed subsection \Interchange of Differentiation and Lebesgue Integration\ definition measurize :: "'a measure \ 'b measure \ ('a \ 'b) \ 'a \ 'b" where "measurize M N f = (SOME g. g \ M \\<^sub>M N \ (\S\(null_sets M). {x \ space M. f x \ g x} \ S))" \ \The term "measurize" is what I coined in this formalization.\ lemma fixes f g assumes "g \ M \\<^sub>M N" "S \ null_sets M" "{x \ space M. f x \ g x} \ S" shows measurizeI: "AE x in M. f x = measurize M N f x" and measurizeI2: "AE x in M. g x = measurize M N f x" and measurize_measurable: "measurize M N f \ measurable M N" proof - let ?G = "\g. g \ M \\<^sub>M N" and ?S = "\g. \S\null_sets M. {x \ space M. f x \ g x} \ S" show "AE x in M. f x = measurize M N f x" unfolding measurize_def apply (rule someI2[of "\g. ?G g \ ?S g" g]) using assms apply blast using AE_I' by auto moreover have "AE x in M. g x = f x" using assms by (smt (verit, best) AE_I' Collect_cong) ultimately show "AE x in M. g x = measurize M N f x" by force show "measurize M N f \ measurable M N" unfolding measurize_def apply (rule conjE[of "?G g" "?S g"]) using assms apply auto[1] using someI_ex[of "\g. ?G g \ ?S g"] by auto qed corollary measurable_measurize_AE: fixes f assumes "f \ M \\<^sub>M N" shows "AE x in M. f x = measurize M N f x" by (rule measurizeI[where g=f and S="{}"]; simp add: assms) definition borel_measurize :: "'a measure \ ('a \ 'b::topological_space) \ 'a \ 'b" where "borel_measurize M f = measurize M borel f" lemma fixes f g assumes "g \ borel_measurable M" "S \ null_sets M" "{x \ space M. f x \ g x} \ S" shows borel_measurizeI: "AE x in M. f x = borel_measurize M f x" and borel_measurizeI2: "AE x in M. g x = borel_measurize M f x" and borel_measurize_measurable: "borel_measurize M f \ borel_measurable M" unfolding borel_measurize_def using assms apply - using measurizeI apply blast using measurizeI2 apply blast using measurize_measurable by blast corollary borel_measurable_measurize_AE: fixes f assumes "f \ borel_measurable M" shows "AE x in M. f x = borel_measurize M f x" using assms measurable_measurize_AE unfolding borel_measurize_def by auto definition set_borel_measurize :: "'a measure \ 'a set \ ('a \ 'b::topological_space) \ 'a \ 'b" where "set_borel_measurize M A f = borel_measurize (restrict_space M A) f" lemma fixes f g :: "'a \ 'b::real_normed_vector" and A assumes "A \ sets M" "set_borel_measurable M A g" "S \ null_sets M" "{x \ A. f x \ g x} \ S" shows set_borel_measurizeI: "AE x\A in M. f x = set_borel_measurize M A f x" and set_borel_measurizeI2: "AE x\A in M. g x = set_borel_measurize M A f x" and set_borel_measurize_measurable: "set_borel_measurable M A (set_borel_measurize M A f)" proof - have "g \ borel_measurable (restrict_space M A)" using assms by (rewrite set_borel_measurable_restrict_space_iff; simp) moreover have "S \ A \ null_sets (restrict_space M A)" using assms null_sets_restrict_space by (metis Int_lower2 null_set_Int2) moreover have "{x \ space (restrict_space M A). f x \ g x} \ S \ A" using assms by (rewrite space_restrict_space2; simp) ultimately show "AE x\A in M. f x = set_borel_measurize M A f x" and "AE x\A in M. g x = set_borel_measurize M A f x" and "set_borel_measurable M A (set_borel_measurize M A f)" unfolding set_borel_measurize_def using assms apply - apply (rewrite AE_restrict_space_iff[THEN sym], simp) apply (rule borel_measurizeI[of g _ "S \ A"]; simp) apply (rewrite AE_restrict_space_iff[THEN sym], simp) apply (rule borel_measurizeI2[of g _ "S \ A"]; simp) apply (rewrite set_borel_measurable_restrict_space_iff[THEN sym], simp) by (rule borel_measurize_measurable[of g _ "S \ A"]; simp) qed corollary set_borel_measurable_measurize_AE: fixes f::"'a \ 'b::real_normed_vector" and A assumes "set_borel_measurable M A f" "A \ sets M" shows "AE x\A in M. f x = set_borel_measurize M A f x" using set_borel_measurable_restrict_space_iff borel_measurable_measurize_AE AE_restrict_space_iff unfolding set_borel_measurize_def by (smt (verit) AE_cong sets.Int_space_eq2 assms) proposition interchange_deriv_LINT_general: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" assumes f_integ: "\r. r\{a<.. integrable M (f r)" and f_diff: "AE x in M. (\r. f r x) differentiable_on {a<..r\{a<..deriv (\r. f r x) r\ \ g x" "integrable M g" shows "\r. r\{a<.. ((\r. \x. f r x \M) has_real_derivative \x. borel_measurize M (\x. deriv (\r. f r x) r) x \M) (at r)" proof - text \Preparation\ have f_msr: "\r. r\{a<.. f r \ borel_measurable M" using f_integ by auto from f_diff obtain N1 where N1_null: "N1 \ null_sets M" and "\x. x \ space M - N1 \ (\s. f s x) differentiable_on {a<..x. x \ space M - N1 \ (\s. f s x) differentiable_on {a<.. null_sets M" and "\x. x \ space M - N2 \ \r\{a<..deriv (\s. f s x) r\ \ g x" by (smt (verit) AE_E3) hence Df_boundN2:"\x. x \ space M - N2 \ \r\{a<..deriv (\s. f s x) r\ \ g x" by (meson Diff_iff sets.sets_into_space subset_eq) define N where "N \ N1 \ N2" let ?CN = "space M - N" have N_null: "N \ null_sets M" and N_msr: "N \ sets M" unfolding N_def using N1_null N2_null by auto have f_diffCN: "\x. x\?CN \ (\s. f s x) differentiable_on {a<.. 'a \ real" where "Df r x \ indicator ({a<..?CN) (r,x) * deriv (\s. f s x) r" for r x have Df_boundCN: "\x. x\?CN \ \r\{a<..Df r x\ \ g x" unfolding Df_def N_def using Df_boundN2 by simp text \Main Part of the Proof\ fix r assume r_ab: "r\{a<.. 0" and ball_ab: "ball r e \ {a<..d::nat\real. \\i. d i \ UNIV-{0}; d \ 0\ \ ((\h. ((\x. f (r+h) x \M) - \x. f r x \M) / h) \ d) \ \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" proof - fix d::"nat\real" assume d_neq0: "\i. d i \ UNIV-{0}" and d_to0: "d \ 0" then obtain m where "\i\m. \d i - 0\ < e" using LIMSEQ_def e_pos dist_real_def by metis hence rd_ab: "\n. r + d (n+m) \ {a<..n. (\x. (f (r + d (n+m)) x - f r x) / d (n+m)) \ borel_measurable M" using r_ab by (measurable; (intro f_msr)?; simp) hence limf_msr: "(\x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))) \ borel_measurable M" by measurable moreover have limf_Df: "\x. x\?CN \ (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \ Df r x" proof - fix x assume x_CN: "x\?CN" hence "(\s. f s x) field_differentiable (at r)" using f_diffCN r_ab by (metis at_within_open differentiable_on_eq_field_differentiable_real open_greaterThanLessThan) hence "((\h. (f (r+h) x - f r x) / h) \ Df r x) (at 0)" apply (rewrite in asm DERIV_deriv_iff_field_differentiable[THEN sym]) unfolding Df_def using r_ab x_CN by (simp add: DERIV_def) hence "(\i. (f (r + d i) x - f r x) / d i) \ Df r x" apply (rewrite in asm tendsto_at_iff_sequentially) apply (rule allE'[where x=d], simp) unfolding comp_def using d_neq0 d_to0 by simp thus "(\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \ Df r x" by (rule LIMSEQ_ignore_initial_segment[where k=m]) qed ultimately have Df_eq: "\x. Df r x = indicator ?CN x * lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))" proof - fix x show "Df r x = indicator ?CN x * lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))" proof (cases \x\?CN\) case True hence "lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) = Df r x" by (intro limI, rule limf_Df) thus ?thesis using True by simp next case False thus ?thesis unfolding Df_def by simp qed qed hence Df_msr: "Df r \ borel_measurable M" apply (rewrite in "\x. \" Df_eq) apply (measurable; (rule limf_msr)?) using N_null unfolding null_sets_def by force have "((\h. ((\x. f (r+h) x \M) - \x. f r x \M) / h) \ d) \ \x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M" proof - have "(\n. \x. (f (r + d (n+m)) x - f r x) / d (n+m) \M) \ \x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M" proof - \ \by Lebesgue's Dominated Convergence Theorem\ have "AE x in M. (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \ lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))" using limf_Df Df_eq N_null by (smt (verit) DiffI AE_I' limI mem_Collect_eq subset_eq) moreover have "\n. AE x in M. norm ((f (r + d (n+m)) x - f r x) / d (n+m)) \ g x" proof - fix n { fix x assume x_CN: "x\?CN" let ?I = "{r..(r + d (n+m))} \ {(r + d (n+m))..r}" have f_diffI: "(\s. f s x) differentiable_on ?I" apply (rule differentiable_on_subset[where t="{a<..s. f s x)" "(\s. f s x) differentiable_on interior ?I" apply - using differentiable_imp_continuous_on apply blast by (metis differentiable_on_subset interior_subset) then obtain t where t_01: "t\{0<..<1}" and f_MVT: "f (r + d (n+m)) x - f r x = d (n+m) * deriv (\s. f s x) (r + t * (d (n+m)))" by (rule MVT_order_free) hence "0 < t" "t < 1" by simp_all hence rtd_ab: "r + t * (d (n+m)) \ {a<..s. f s x) (r + t * (d (n+m))) = d (n+m) * Df (r + t * (d (n+m))) x" proof - have "r + t * (d (n+m)) \ {a<..Df (r + t * (d (n+m))) x\ \ g x" using Df_boundCN x_CN rtd_ab by simp ultimately have "\(f (r + d (n+m)) x - f r x) / d (n+m)\ \ g x" by simp } thus "AE x in M. norm ((f (r + d (n+m)) x - f r x) / d (n+m)) \ g x" unfolding real_norm_def using AE_I' N_null by (smt (verit, ccfv_threshold) Diff_iff mem_Collect_eq subsetI) qed ultimately show "((\n. \x. (f (r + d (n+m)) x - f r x) / d (n+m) \M) \ \x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M)" using limf_msr fd_msr Df_bound by (intro integral_dominated_convergence[where w=g], simp_all) qed moreover have "\n. ((\x. f (r + d (n+m)) x \M) - \x. f r x \M) / d (n+m) = \x. (f (r + d (n+m)) x - f r x) / d (n+m) \M" using d_neq0 apply simp by (rewrite Bochner_Integration.integral_diff; (rule f_integ | simp); (rule rd_ab | rule r_ab)) ultimately show ?thesis unfolding comp_def using d_neq0 apply - by (rule LIMSEQ_offset[where k=m]) simp qed moreover have "(\x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M) = \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" proof - have "(\x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M) = \x. Df r x \M" proof - have "AE x in M. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) = Df r x" proof - { fix x assume x_CN: "x\?CN" hence "lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) = Df r x" by (simp add: Df_eq) } thus ?thesis using AE_I' N_null by (smt (verit, del_insts) DiffI mem_Collect_eq subsetI) qed thus ?thesis using limf_msr Df_msr by (intro integral_cong_AE; simp) qed also have "\ = \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" proof - have "AE x in M. Df r x = borel_measurize M (\y. deriv (\s. f s y) r) x" and "borel_measurize M (\y. deriv (\s. f s y) r) \ borel_measurable M" proof - have "{x \ space M. deriv (\s. f s x) r \ Df r x} \ N" proof - { fix x assume "x\?CN" hence "deriv (\s. f s x) r = Df r x" unfolding Df_def using r_ab by simp } thus ?thesis by blast qed thus "AE x in M. Df r x = borel_measurize M (\y. deriv (\s. f s y) r) x" and "borel_measurize M (\y. deriv (\s. f s y) r) \ borel_measurable M" using Df_msr N_null apply - apply (rule borel_measurizeI2[where S=N]; simp) by (rule borel_measurize_measurable[where g="Df r"]; simp) qed thus ?thesis using Df_msr by (intro integral_cong_AE; simp) qed finally show ?thesis . qed ultimately show "((\h. ((\x. f (r+h) x \M) - \x. f r x \M) / h) \ d) \ \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" using tendsto_cong_limit by simp qed thus "((\s. \x. f s x \M) has_real_derivative \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M) (at r)" by (rewrite DERIV_def, rewrite tendsto_at_iff_sequentially) simp qed proposition interchange_deriv_LINT: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" assumes "\r. r\{a<.. integrable M (f r)" and "AE x in M. (\r. f r x) differentiable_on {a<..r. r\{a<.. (\x. (deriv (\r. f r x) r)) \ borel_measurable M" and "AE x in M. \r\{a<..deriv (\r. f r x) r\ \ g x" "integrable M g" shows "\r. r\{a<.. ((\r. \x. f r x \M) has_real_derivative \x. deriv (\r. f r x) r \M) (at r)" proof - fix r assume r_ab: "r\{a<..x. deriv (\s. f s x) r) \ borel_measurable M" using assms by simp have "((\s. \x. f s x \M) has_real_derivative \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M) (at r)" using assms r_ab by (intro interchange_deriv_LINT_general; simp) moreover have "(\x. borel_measurize M (\y. deriv (\s. f s y) r) x \M) = \x. deriv (\s. f s x) r \M" apply (rule integral_cong_AE) apply (rule borel_measurize_measurable[where g="\y. deriv (\s. f s y) r" and S="{}"], simp_all add: Df_msr) using borel_measurable_measurize_AE Df_msr by (smt (verit) AE_cong) ultimately show "((\r. \x. f r x \M) has_real_derivative \x. deriv (\r. f r x) r \M) (at r)" by simp qed proposition interchange_deriv_LINT_set_general: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" and A :: "'a set" assumes A_msr: "A \ sets M" and f_integ: "\r. r\{a<.. set_integrable M A (f r)" and f_diff: "AE x\A in M. (\r. f r x) differentiable_on {a<..A in M. \r\{a<..deriv (\r. f r x) r\ \ g x" "set_integrable M A g" shows "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative \x\A. set_borel_measurize M A (\x. deriv (\r. f r x) r) x \M) (at r)" proof - let ?M_A = "restrict_space M A" have "\r. r\{a<.. integrable ?M_A (f r)" using A_msr f_integ set_integrable_restrict_space_iff by auto moreover have "AE x in ?M_A. (\r. f r x) differentiable_on {a<..r\{a<..deriv (\r. f r x) r\ \ g x" and "integrable ?M_A g" using A_msr Df_bound set_integrable_restrict_space_iff apply - by (simp add: AE_restrict_space_iff, auto) ultimately have "\r. r\{a<.. ((\r. integral\<^sup>L ?M_A (f r)) has_real_derivative integral\<^sup>L ?M_A (borel_measurize ?M_A (\x. deriv (\r. f r x) r))) (at r)" by (rule interchange_deriv_LINT_general[where M="restrict_space M A"]) auto thus "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative \x\A. set_borel_measurize M A (\x. deriv (\r. f r x) r) x \M) (at r)" unfolding set_borel_measurize_def using assms by (rewrite set_lebesgue_integral_restrict_space, simp)+ qed proposition interchange_deriv_LINT_set: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" and A :: "'a set" assumes "A \ sets M" and "\r. r\{a<.. set_integrable M A (f r)" and "AE x\A in M. (\r. f r x) differentiable_on {a<..r. r\{a<.. set_borel_measurable M A (\x. (deriv (\r. f r x) r))" and "AE x\A in M. \r\{a<..deriv (\r. f r x) r\ \ g x" "set_integrable M A g" shows "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative \x\A. deriv (\r. f r x) r \M) (at r)" proof - fix r assume r_ab: "r\{a<..x. deriv (\s. f s x) r)" using assms by simp have "((\s. \x\A. f s x \M) has_real_derivative \x\A. set_borel_measurize M A (\y. deriv (\s. f s y) r) x \M) (at r)" using assms r_ab by (intro interchange_deriv_LINT_set_general; simp) moreover have "(\x\A. set_borel_measurize M A (\y. deriv (\s. f s y) r) x \M) = \x\A. deriv (\s. f s x) r \M" apply (rule set_lebesgue_integral_cong_AE2, simp add: assms) apply (rule set_borel_measurize_measurable[where g="\y. deriv (\s. f s y) r" and S="{}"], simp_all add: Df_msr assms) using set_borel_measurable_measurize_AE Df_msr assms by (smt (verit) AE_cong) ultimately show "((\r. \x\A. f r x \M) has_real_derivative \x\A. deriv (\r. f r x) r \M) (at r)" by simp qed section \Additional Lemmas for the "Probability" Library\ lemma (in finite_borel_measure) fixes F :: "real \ real" assumes nondecF : "\ x y. x \ y \ F x \ F y" and right_cont_F : "\a. continuous (at_right a) F" and lim_F_at_bot : "(F \ 0) at_bot" and lim_F_at_top : "(F \ m) at_top" and m : "0 \ m" shows emeasure_interval_measure_Ioi: "emeasure (interval_measure F) {x<..} = m - F x" and measure_interval_measure_Ioi: "measure (interval_measure F) {x<..} = m - F x" proof - interpret F_FM: finite_measure "interval_measure F" using finite_borel_measure.axioms(1) finite_borel_measure_interval_measure lim_F_at_bot lim_F_at_top m nondecF right_cont_F by blast have "UNIV = {..x} \ {x<..}" by auto moreover have "{..x} \ {x<..} = {}" by auto ultimately have "emeasure (interval_measure F) UNIV = emeasure (interval_measure F) {..x} + emeasure (interval_measure F) {x<..}" by (simp add: plus_emeasure) moreover have "emeasure (interval_measure F) UNIV = m" using assms interval_measure_UNIV by presburger ultimately show \: "emeasure (interval_measure F) {x<..} = m - F x" using assms emeasure_interval_measure_Iic by (metis ennreal_add_diff_cancel_left ennreal_minus measure_interval_measure_Iic measure_nonneg top_neq_ennreal) hence "ennreal (measure (interval_measure F) {x<..}) = m - F x" using emeasure_eq_measure by (metis emeasure_eq_ennreal_measure top_neq_ennreal) moreover have "\x. F x \ m" using lim_F_at_top nondecF by (intro mono_at_top_le[where f=F]; simp add: mono_def) ultimately show "measure (interval_measure F) {x<..} = m - F x" using ennreal_inj F_FM.emeasure_eq_measure by force qed lemma (in prob_space) cond_prob_nonneg[simp]: "cond_prob M P Q \ 0" by (auto simp: cond_prob_def) lemma (in prob_space) cond_prob_whole_1: "cond_prob M P P = 1" if "prob {\ \ space M. P \} \ 0" unfolding cond_prob_def using that by simp lemma (in prob_space) cond_prob_0_null: "cond_prob M P Q = 0" if "prob {\ \ space M. Q \} = 0" unfolding cond_prob_def using that by simp lemma (in prob_space) cond_prob_AE_prob: assumes "{\ \ space M. P \} \ events" "{\ \ space M. Q \} \ events" and "AE \ in M. Q \" shows "cond_prob M P Q = prob {\ \ space M. P \}" proof - let ?setP = "{\ \ space M. P \}" let ?setQ = "{\ \ space M. Q \}" have [simp]: "prob ?setQ = 1" using assms prob_Collect_eq_1 by simp hence "cond_prob M P Q = prob (?setP \ ?setQ)" unfolding cond_prob_def by (simp add: Collect_conj_eq2) also have "\ = prob ?setP" proof (rule antisym) show "prob (?setP \ ?setQ) \ prob ?setP" using assms finite_measure_mono inf_sup_ord(1) by blast next show "prob ?setP \ prob (?setP \ ?setQ)" proof - have "prob (?setP \ ?setQ) = prob ?setP + prob ?setQ - prob (?setP \ ?setQ)" using assms by (smt (verit) finite_measure_Diff' finite_measure_Union' sup_commute) also have "\ = prob ?setP + (1 - prob (?setP \ ?setQ))" by simp also have "\ \ prob ?setP" by simp finally show ?thesis . qed qed finally show ?thesis . qed subsection \More Properties of cdf's\ context finite_borel_measure begin lemma cdf_diff_eq2: assumes "x \ y" shows "cdf M y - cdf M x = measure M {x<..y}" proof (cases \x = y\) case True thus ?thesis by force next case False hence "x < y" using assms by simp thus ?thesis by (rule cdf_diff_eq) qed end context prob_space begin lemma cdf_distr_measurable [measurable]: assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel X) \ borel_measurable borel" proof (rule borel_measurable_mono) show "mono (cdf (distr M borel X))" unfolding mono_def using finite_borel_measure.cdf_nondecreasing by (simp add: real_distribution.finite_borel_measure_M) qed lemma cdf_distr_P: assumes "random_variable borel X" shows "cdf (distr M borel X) x = \

(\ in M. X \ \ x)" unfolding cdf_def apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob], force) lemma cdf_continuous_distr_P_lt: assumes "random_variable borel X" "isCont (cdf (distr M borel X)) x" shows "cdf (distr M borel X) x = \

(\ in M. X \ < x)" proof - have "\

(\ in M. X \ < x) = measure (distr M borel X) {.. = measure (distr M borel X) ({.. {x})" apply (rewrite finite_measure.measure_zero_union, simp_all add: assms finite_measure_distr) using finite_borel_measure.isCont_cdf real_distribution.finite_borel_measure_M assms by blast also have "\ = measure (distr M borel X) {..x}" by (metis ivl_disj_un_singleton(2)) also have "\ = cdf (distr M borel X) x" unfolding cdf_def by simp finally show ?thesis by simp qed lemma cdf_distr_diff_P: assumes "x \ y" and "random_variable borel X" shows "cdf (distr M borel X) y - cdf (distr M borel X) x = \

(\ in M. x < X \ \ X \ \ y)" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp have "cdf (distr M borel X) y - cdf (distr M borel X) x = measure (distr M borel X) {x<..y}" by (rewrite distrX_FBM.cdf_diff_eq2; simp add: assms) also have "\ = \

(\ in M. x < X \ \ X \ \ y)" apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob], force) finally show ?thesis . qed lemma cdf_distr_max: fixes c::real assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel (\x. max (X x) c)) u = cdf (distr M borel X) u * indicator {c..} u" proof (cases \c \ u\) case True thus ?thesis unfolding cdf_def apply (rewrite measure_distr; simp?)+ by (smt (verit) Collect_cong atMost_iff vimage_def) next case False thus ?thesis unfolding cdf_def apply (rewrite measure_distr; simp?)+ by (smt (verit, best) Int_emptyI atMost_iff measure_empty vimage_eq) qed lemma cdf_distr_min: fixes c::real assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel (\x. min (X x) c)) u = 1 - (1 - cdf (distr M borel X) u) * indicator {..c \ u\) case True thus ?thesis unfolding cdf_def using real_distribution.finite_borel_measure_M real_distribution_distr apply (rewrite measure_distr; simp?) by (smt (verit, del_insts) Int_absorb1 atMost_iff prob_space subset_vimage_iff) next case False thus ?thesis unfolding cdf_def using real_distribution.finite_borel_measure_M real_distribution_distr apply (rewrite measure_distr; simp?)+ using prob_space_axioms assms by (smt (verit) Collect_cong Int_def atMost_iff prob_space prob_space.cdf_distr_P vimage_eq) qed lemma cdf_distr_floor_P: fixes X :: "'a \ real" assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel (\x. \X x\)) u = \

(x in M. X x < \u\ + 1)" unfolding cdf_def apply (rewrite measure_distr; simp?) apply (rule arg_cong[where f=prob]) unfolding vimage_def using floor_le_iff le_floor_iff by blast lemma expectation_nonneg_tail: assumes [measurable]: "random_variable borel X" and X_nonneg: "\x. x \ space M \ X x \ 0" defines "F u \ cdf (distr M borel X) u" shows "(\\<^sup>+x. ennreal (X x) \M) = \\<^sup>+u\{0..}. ennreal (1 - F u) \lborel" proof - let ?distrX = "distr M borel X" have "(\\<^sup>+x. ennreal (X x) \M) = (\\<^sup>+u. ennreal u \?distrX)" by (rewrite nn_integral_distr; simp) also have "\ = (\\<^sup>+u\{0..}. ennreal u \?distrX)" by (rule nn_integral_distr_set; simp add: X_nonneg) also have "\ = \\<^sup>+u\{0..}. (\\<^sup>+v\{0..}. indicator {..lborel) \?distrX" proof - have "\u::real. u\{0..} \ ennreal u = \\<^sup>+v\{0..}. indicator {..lborel" apply (rewrite indicator_inter_arith[THEN sym]) apply (rewrite nn_integral_indicator, measurable, simp) by (metis atLeastLessThan_def diff_zero emeasure_lborel_Ico inf.commute) thus ?thesis by (metis (no_types, lifting) indicator_eq_0_iff mult_eq_0_iff) qed also have "\ = \\<^sup>+v\{0..}. (\\<^sup>+u\{0..}. indicator {..?distrX) \lborel" proof - have "\\<^sup>+u\{0..}. (\\<^sup>+v\{0..}. indicator {..lborel) \?distrX = \\<^sup>+u. (\\<^sup>+v. indicator {..lborel) \?distrX" by (rewrite nn_integral_multc; simp) also have "\ = \\<^sup>+v. (\\<^sup>+u. indicator {..?distrX) \lborel" apply (rewrite pair_sigma_finite.Fubini'; simp?) using pair_sigma_finite.intro assms prob_space_distr prob_space_imp_sigma_finite sigma_finite_lborel apply blast by measurable auto also have "\ = \\<^sup>+v\{0..}. (\\<^sup>+u\{0..}. indicator {..?distrX) \lborel" apply (rewrite nn_integral_multc[THEN sym]; measurable; simp?) apply (rule nn_integral_cong)+ using mult.assoc mult.commute by metis finally show ?thesis by simp qed also have "\ = \\<^sup>+v\{0..}. (\\<^sup>+u. indicator {v<..} u \?distrX) \lborel" apply (rule nn_integral_cong) apply (rewrite nn_integral_multc[THEN sym], measurable; (simp del: nn_integral_indicator)?)+ apply (rule nn_integral_cong) using lessThan_iff greaterThan_iff atLeast_iff indicator_simps by (smt (verit, del_insts) mult_1 mult_eq_0_iff) also have "\ = (\\<^sup>+v\{0..}. ennreal (1 - F v) \lborel)" apply (rule nn_integral_cong, simp) apply (rewrite emeasure_distr; simp?) apply (rewrite vimage_compl_atMost[THEN sym]) unfolding F_def cdf_def apply (rewrite measure_distr; simp?) apply (rewrite prob_compl[THEN sym], simp) by (metis (no_types, lifting) Diff_Compl Diff_Diff_Int Int_commute emeasure_eq_measure) finally show ?thesis . qed lemma expectation_nonpos_tail: assumes [measurable]: "random_variable borel X" and X_nonpos: "\x. x \ space M \ X x \ 0" defines "F u \ cdf (distr M borel X) u" shows "(\\<^sup>+x. ennreal (- X x) \M) = \\<^sup>+u\{..0}. ennreal (F u) \lborel" proof - let ?distrX = "distr M borel X" have "(\\<^sup>+x. ennreal (- X x) \M) = (\\<^sup>+u. ennreal (-u) \?distrX)" by (rewrite nn_integral_distr; simp) also have "\ = \\<^sup>+u\{..0}. ennreal (-u) \?distrX" proof - have [simp]: "{..0::real} \ {0<..} = UNIV" by force have "(\\<^sup>+u. ennreal (-u) \?distrX) = (\\<^sup>+u\{..0}. ennreal (-u) \?distrX) + (\\<^sup>+u\{0<..}. ennreal (-u) \?distrX)" by (rewrite nn_integral_disjoint_pair[THEN sym], simp_all, force) also have "\ = (\\<^sup>+u\{..0}. ennreal (-u) \?distrX)" apply (rewrite nn_integral_zero'[of "\u. ennreal (-u) * indicator {0<..} u"]; simp?) unfolding indicator_def using always_eventually ennreal_lt_0 by force finally show ?thesis . qed also have "\ = \\<^sup>+u\{..0}. (\\<^sup>+v\{..0}. indicator {u..} v \lborel) \?distrX" proof - have "\u::real. u\{..0} \ ennreal (-u) = \\<^sup>+v\{..0}. indicator {u..} v \lborel" apply (rewrite indicator_inter_arith[THEN sym]) apply (rewrite nn_integral_indicator, measurable, simp) by (metis emeasure_lborel_Icc atLeastAtMost_def diff_0) thus ?thesis by (metis (no_types, lifting) indicator_eq_0_iff mult_eq_0_iff) qed also have "\ = \\<^sup>+v\{..0}. (\\<^sup>+u\{..0}. indicator {u..} v \?distrX) \lborel" proof - have "\\<^sup>+u\{..0}. (\\<^sup>+v\{..0}. indicator {u..} v \lborel) \?distrX = \\<^sup>+u. (\\<^sup>+v. indicator {u..} v * indicator {..0} v * indicator {..0} u \lborel) \?distrX" by (rewrite nn_integral_multc; simp) also have "\ = \\<^sup>+v. (\\<^sup>+u. indicator {u..} v * indicator {..0} v * indicator {..0} u \?distrX) \lborel" apply (rewrite pair_sigma_finite.Fubini'; simp?) using pair_sigma_finite.intro assms prob_space_distr prob_space_imp_sigma_finite sigma_finite_lborel apply blast by measurable auto also have "\ = \\<^sup>+v\{..0}. (\\<^sup>+u\{..0}. indicator {u..} v \?distrX) \lborel" apply (rewrite nn_integral_multc[THEN sym]; measurable; simp?) apply (rule nn_integral_cong)+ using mult.assoc mult.commute by metis finally show ?thesis by simp qed also have "\ = \\<^sup>+v\{..0}. (\\<^sup>+u. indicator {..v} u \?distrX) \lborel" apply (rule nn_integral_cong) apply (rewrite nn_integral_multc[THEN sym], measurable; (simp del: nn_integral_indicator)?)+ apply (rule nn_integral_cong) using atMost_iff atLeast_iff indicator_simps by (smt (verit, del_insts) mult_1 mult_eq_0_iff) also have "\ = \\<^sup>+v\{..0}. ennreal (F v) \lborel" apply (rule nn_integral_cong, simp) apply (rewrite emeasure_distr; simp?) unfolding F_def cdf_def by (rewrite measure_distr; simp add: emeasure_eq_measure) finally show ?thesis . qed proposition expectation_tail: assumes [measurable]: "integrable M X" defines "F u \ cdf (distr M borel X) u" shows "expectation X = (LBINT u:{0..}. 1 - F u) - (LBINT u:{..0}. F u)" proof - have "expectation X = expectation (\x. max (X x) 0) + expectation (\x. min (X x) 0)" using integrable_max integrable_min apply (rewrite Bochner_Integration.integral_add[THEN sym], measurable) by (rule Bochner_Integration.integral_cong; simp) also have "\ = expectation (\x. max (X x) 0) - expectation (\x. - min (X x) 0)" by force also have "\ = (LBINT u:{0..}. 1 - F u) - (LBINT u:{..0}. F u)" proof - have "expectation (\x. max (X x) 0) = LBINT u:{0..}. 1 - F u" proof - have "expectation (\x. max (X x) 0) = enn2real (\\<^sup>+x. ennreal (max (X x) 0) \M)" by (rule integral_eq_nn_integral; simp) also have "\ = enn2real (\\<^sup>+u\{0..}. ennreal (1 - F u) \lborel)" apply (rewrite expectation_nonneg_tail; simp?) apply (rewrite cdf_distr_max, simp) unfolding F_def by (metis (opaque_lifting) indicator_simps mult.commute mult_1 mult_eq_0_iff) also have "\ = enn2real (\\<^sup>+u. ennreal ((1 - F u) * indicator {0..} u) \lborel)" by (simp add: indicator_mult_ennreal mult.commute) also have "\ = LBINT u:{0..}. 1 - F u" apply (rewrite integral_eq_nn_integral[THEN sym], simp add: F_def) unfolding F_def using real_distribution.cdf_bounded_prob apply force unfolding set_lebesgue_integral_def by (rule Bochner_Integration.integral_cong; simp) finally show ?thesis . qed moreover have "expectation (\x. - min (X x) 0) = LBINT u:{..0}. F u" proof - have "expectation (\x. - min (X x) 0) = enn2real (\\<^sup>+x. ennreal (- min (X x) 0) \M)" by (rule integral_eq_nn_integral; simp) also have "\ = enn2real (\\<^sup>+u\{..0}. ennreal (F u) \lborel)" proof - let ?distrminX = "distr M borel (\x. min (X x) 0)" have [simp]: "sym_diff {..0} {..<0} = {0::real}" by force have "enn2real (\\<^sup>+x. ennreal (- min (X x) 0) \M) = enn2real (\\<^sup>+u\{..0}. ennreal (cdf ?distrminX u) \lborel)" by (rewrite expectation_nonpos_tail; simp) also have "\ = enn2real (\\<^sup>+u\{..<0}. ennreal (cdf ?distrminX u) \lborel)" by (rewrite nn_integral_null_delta, auto) also have "\ = enn2real (\\<^sup>+u\{..<0}. ennreal (F u) \lborel)" apply (rewrite cdf_distr_min, simp) apply (rule arg_cong[where f=enn2real], rule nn_integral_cong) unfolding F_def by (smt (verit) indicator_simps mult_cancel_left1 mult_eq_0_iff) also have "\ = enn2real (\\<^sup>+u\{..0}. ennreal (F u) \lborel)" by (rewrite nn_integral_null_delta, auto simp add: sup_commute) finally show ?thesis . qed also have "\ = enn2real (\\<^sup>+u. ennreal (F u * indicator {..0} u) \lborel)" using mult.commute indicator_mult_ennreal by metis also have "\ = LBINT u:{..0}. F u" apply (rewrite integral_eq_nn_integral[THEN sym], simp add: F_def) unfolding F_def using finite_borel_measure.cdf_nonneg real_distribution.finite_borel_measure_M apply simp unfolding set_lebesgue_integral_def by (rule Bochner_Integration.integral_cong; simp) finally show ?thesis . qed ultimately show ?thesis by simp qed finally show ?thesis . qed proposition distributed_deriv_cdf: assumes [measurable]: "random_variable borel X" defines "F u \ cdf (distr M borel X) u" assumes "finite S" "\x. x \ S \ (F has_real_derivative f x) (at x)" and "continuous_on UNIV F" "f \ borel_measurable lborel" shows "distributed M lborel X f" proof - have FBM: "finite_borel_measure (distr M borel X)" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp then interpret distrX_FBM: finite_borel_measure "distr M borel X" . have FBMl: "finite_borel_measure (distr M lborel X)" using FBM distr_borel_lborel by smt then interpret distrlX_FBM: finite_borel_measure "distr M lborel X" . have [simp]: "(\x. ennreal (f x)) \ borel_measurable borel" using assms by simp moreover have "distr M lborel X = density lborel f" proof - have "\a b. a \ b \ emeasure (distr M lborel X) {a<..b} < \" using distrlX_FBM.emeasure_real less_top_ennreal by blast moreover have "\a b. a \ b \ emeasure (distr M lborel X) {a<..b} = emeasure (density lborel f) {a<..b}" proof - fix a b :: real assume "a \ b" hence [simp]: "sym_diff {a<..b} {a..b} = {a}" by force have "emeasure (density lborel f) {a<..b} = \\<^sup>+x\{a<..b}. ennreal (f x) \lborel" by (rewrite emeasure_density; simp) also have "\ = \\<^sup>+x\{a..b}. ennreal (f x) \lborel" by (rewrite nn_integral_null_delta, auto) also have "\ = \\<^sup>+x. ennreal (indicat_real {a..b} x * f x) \lborel" by (metis indicator_mult_ennreal mult.commute) also have "\ = ennreal (F b - F a)" proof - define g where "g x = (if x \ S then 0 else f x)" for x :: real have [simp]: "\x. g x \ 0" unfolding g_def apply (split if_split, auto) apply (rule mono_on_imp_deriv_nonneg[of UNIV F], auto) unfolding F_def mono_on_def using distrX_FBM.cdf_nondecreasing apply blast using assms unfolding F_def by force have "(\\<^sup>+x. ennreal (indicat_real {a..b} x * f x) \lborel) = \\<^sup>+x. ennreal (indicat_real {a..b} x * g x) \lborel" apply (rule nn_integral_cong_AE) apply (rule AE_mp[where P= "\x. x \ S"]) using assms finite_imp_null_set_lborel AE_not_in apply blast unfolding g_def by simp also have "\ = ennreal (F b - F a)" apply (rewrite nn_integral_has_integral_lebesgue, simp) apply (rule fundamental_theorem_of_calculus_strong[of S], auto simp: \a \ b\ g_def assms) using has_real_derivative_iff_has_vector_derivative assms apply presburger using assms continuous_on_subset subsetI by fastforce finally show ?thesis . qed also have "\ = emeasure (distr M lborel X) {a <.. b}" apply (rewrite distrlX_FBM.emeasure_Ioc, simp add: \a \ b\) unfolding F_def cdf_def apply (rewrite ennreal_minus[THEN sym], simp)+ by (metis distr_borel_lborel) finally show "emeasure (distr M lborel X) {a<..b} = emeasure (density lborel f) {a<..b}" by simp qed ultimately show ?thesis by (intro measure_eqI_Ioc; simp) qed ultimately show ?thesis unfolding distributed_def by auto qed end text \ Define the conditional probability space. This is obtained by rescaling the restricted space of a probability space. \ subsection \Conditional Probability Space\ lemma restrict_prob_space: assumes "measure_space \ A \" "a \ A" and "0 < \ a" "\ a < \" shows "prob_space (scale_measure (1 / \ a) (restrict_space (measure_of \ A \) a))" proof let ?M = "measure_of \ A \" let ?Ma = "restrict_space ?M a" let ?rMa = "scale_measure (1 / \ a) ?Ma" have "emeasure ?rMa (space ?rMa) = 1 / \ a * emeasure ?Ma (space ?rMa)" by simp also have "\ = 1 / \ a * emeasure ?M (space ?rMa)" using assms apply (rewrite emeasure_restrict_space) apply (simp add: measure_space_def sigma_algebra.sets_measure_of_eq) by (simp add: space_restrict_space space_scale_measure) simp also have "\ = 1 / \ a * emeasure ?M (space ?Ma)" by (rewrite space_scale_measure) simp also have "\ = 1 / \ a * emeasure ?M a" using assms apply (rewrite space_restrict_space2) by (simp add: measure_space_closed)+ also have "\ = 1" using assms measure_space_def apply (rewrite emeasure_measure_of_sigma, blast+) by (simp add: ennreal_divide_times) finally show "emeasure ?rMa (space ?rMa) = 1" . qed definition cond_prob_space :: "'a measure \ 'a set \ 'a measure" (infix "\" 200) where "M\A \ scale_measure (1 / emeasure M A) (restrict_space M A)" context prob_space begin lemma cond_prob_space_whole[simp]: "M \ space M = M" unfolding cond_prob_space_def using emeasure_space_1 by simp lemma cond_prob_space_correct: assumes "A \ events" "prob A > 0" shows "prob_space (M\A)" unfolding cond_prob_space_def apply (subst(2) measure_of_of_measure[of M, THEN sym]) using assms by (intro restrict_prob_space; (simp add: measure_space)?; simp_all add: emeasure_eq_measure) lemma space_cond_prob_space: assumes "A \ events" shows "space (M\A) = A" unfolding cond_prob_space_def using assms by (simp add: space_scale_measure) lemma sets_cond_prob_space: "sets (M\A) = (\) A ` events" unfolding cond_prob_space_def by (metis sets_restrict_space sets_scale_measure) lemma measure_cond_prob_space: assumes "A \ events" "B \ events" and "prob A > 0" shows "measure (M\A) (B \ A) = prob (B \ A) / prob A" proof - have "1 / emeasure M A = ennreal (1 / prob A)" using assms by (smt (verit) divide_ennreal emeasure_eq_measure ennreal_1) hence "measure (M\A) (B \ A) = (1 / prob A) * measure (restrict_space M A) (B \ A)" unfolding cond_prob_space_def using measure_scale_measure by force also have "\ = (1 / prob A) * prob (B \ A)" using measure_restrict_space assms by (metis inf.cobounded2 sets.Int_space_eq2) also have "\ = prob (B \ A) / prob A" by simp finally show ?thesis . qed corollary measure_cond_prob_space_subset: assumes "A \ events" "B \ events" "B \ A" and "prob A > 0" shows "measure (M\A) B = prob B / prob A" proof - have "B = B \ A" using assms by auto moreover have "measure (M\A) (B \ A) = prob (B \ A) / prob A" using assms measure_cond_prob_space by simp ultimately show ?thesis by auto qed lemma cond_cond_prob_space: assumes "A \ events" "B \ events" "B \ A" "prob B > 0" shows "(M\A)\B = M\B" proof (rule measure_eqI) have pA_pos[simp]: "prob A > 0" using assms by (smt (verit, ccfv_SIG) finite_measure_mono) interpret MA_PS: prob_space "M\A" using cond_prob_space_correct assms by simp interpret MB_PS: prob_space "M\B" using cond_prob_space_correct assms by simp have "1 / emeasure M A = ennreal (1 / prob A)" using pA_pos by (smt (verit, ccfv_SIG) divide_ennreal emeasure_eq_measure ennreal_1) hence [simp]: "0 < MA_PS.prob B" using assms pA_pos by (metis divide_eq_0_iff measure_cond_prob_space_subset zero_less_measure_iff) have [simp]: "B \ MA_PS.events" using assms by (rewrite sets_cond_prob_space, unfold image_def) blast have [simp]: "finite_measure ((M\A)\B)" by (rule prob_space.finite_measure, rule prob_space.cond_prob_space_correct, simp_all add: MA_PS.prob_space_axioms) show sets_MAB: "sets ((M\A)\B) = sets (M\B)" apply (rewrite prob_space.sets_cond_prob_space) using MA_PS.prob_space_axioms apply presburger apply (rewrite sets_cond_prob_space, unfold image_def)+ using assms by blast show "\C. C \ sets ((M\A)\B) \ emeasure ((M\A)\B) C = emeasure (M\B) C" proof - fix C assume "C \ sets ((M\A)\B)" hence "C \ sets (M\B)" using sets_MAB by simp from this obtain D where "D \ events" "C = B \ D" by (rewrite in asm sets_cond_prob_space, auto) hence [simp]: "C \ events" and [simp]: "C \ B" and [simp]: "C \ A" using assms by auto hence [simp]: "C \ MA_PS.events" using assms by (rewrite sets_cond_prob_space, unfold image_def) blast show "emeasure ((M\A)\B) C = emeasure (M\B) C" apply (rewrite finite_measure.emeasure_eq_measure, simp)+ apply (rewrite ennreal_inj, simp_all) apply (rewrite prob_space.measure_cond_prob_space_subset, simp_all add: assms prob_space_axioms MA_PS.prob_space_axioms)+ using pA_pos by fastforce qed qed lemma cond_prob_space_prob: assumes[measurable]: "Measurable.pred M P" "Measurable.pred M Q" and "\

(x in M. Q x) > 0" shows "measure (M \ {x \ space M. Q x}) {x \ space M. P x \ Q x} = \

(x in M. P x \ Q x)" proof - let ?SetP = "{x \ space M. P x}" let ?SetQ = "{x \ space M. Q x}" have "measure (M\?SetQ) {x \ space M. P x \ Q x} = measure (M\?SetQ) (?SetP \ ?SetQ)" by (smt (verit, ccfv_SIG) Collect_cong Int_def mem_Collect_eq) also have "\ = prob (?SetP \ ?SetQ) / prob ?SetQ" using assms by (rewrite measure_cond_prob_space; simp?) also have "\ = \

(x in M. P x \ Q x)" unfolding cond_prob_def assms by (smt (verit) Collect_cong Int_def mem_Collect_eq) finally show ?thesis . qed lemma cond_prob_space_cond_prob: assumes [measurable]: "Measurable.pred M P" "Measurable.pred M Q" and "\

(x in M. Q x) > 0" shows "\

(x in M. P x \ Q x) = \

(x in (M \ {x \ space M. Q x}). P x)" proof - let ?setQ = "{x \ space M. Q x}" have "\

(x in M. P x \ Q x) = measure (M\?setQ) {x \ space M. P x \ Q x}" using cond_prob_space_prob assms by simp also have "\ = \

(x in (M\?setQ). P x)" proof - have "{x \ space M. P x \ Q x} = {x \ space (M\?setQ). P x}" using space_cond_prob_space assms by force thus ?thesis by simp qed finally show ?thesis . qed lemma cond_prob_neg: assumes[measurable]: "Measurable.pred M P" "Measurable.pred M Q" and "\

(x in M. Q x) > 0" shows "\

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

(x in M. P x \ Q x)" proof - let ?setP = "{x \ space M. P x}" let ?setQ = "{x \ space M. Q x}" interpret setQ_PS: prob_space "M\?setQ" using cond_prob_space_correct assms by simp have [simp]: "{x \ space (M\?setQ). P x} \ setQ_PS.events" proof - have "{x \ space (M\?setQ). P x} = ?setQ \ ?setP" using space_cond_prob_space by force thus ?thesis using sets_cond_prob_space by simp qed have "\

(x in M. \ P x \ Q x) = \

(x in M\?setQ. \ P x)" by (rewrite cond_prob_space_cond_prob; simp add: assms) also have "\ = 1 - \

(x in M\?setQ. P x)" by (rewrite setQ_PS.prob_neg, simp_all add: assms) also have "\ = 1 - \

(x in M. P x \ Q x)" by (rewrite cond_prob_space_cond_prob; simp add: assms) finally show ?thesis . qed lemma random_variable_cond_prob_space: assumes "A \ events" "prob A > 0" and [measurable]: "random_variable borel X" shows "X \ borel_measurable (M\A)" proof (rule borel_measurableI) fix S :: "'b set" assume [measurable]: "open S" show "X -` S \ space (M \ A) \ sets (M \ A)" apply (rewrite space_cond_prob_space, simp add: assms) apply (rewrite sets_cond_prob_space, simp add: image_def) apply (rule bexI[of _ "X -` S \ space M"]; measurable) using sets.Int_space_eq2 Int_commute assms by auto qed lemma AE_cond_prob_space_iff: assumes "A \ events" "prob A > 0" shows "(AE x in M\A. P x) \ (AE x in M. x \ A \ P x)" proof - have [simp]: "1 / emeasure M A > 0" using assms divide_ennreal emeasure_eq_measure ennreal_1 by (smt (verit) divide_pos_pos ennreal_eq_0_iff not_gr_zero) show ?thesis unfolding cond_prob_space_def apply (rewrite AE_scale_measure_iff, simp) by (rewrite AE_restrict_space_iff; simp add: assms) qed lemma integral_cond_prob_space_nn: assumes "A \ events" "prob A > 0" and [measurable]: "random_variable borel X" and nonneg: "AE x in M. x \ A \ 0 \ X x" shows "integral\<^sup>L (M\A) X = expectation (\x. indicator A x * X x) / prob A" proof - have [simp]: "X \ borel_measurable (M\A)" by (rule random_variable_cond_prob_space; (simp add: assms)) have [simp]: "AE x in (M\A). 0 \ X x" by (rewrite AE_cond_prob_space_iff; simp add: assms) have [simp]: "random_variable borel (\x. indicat_real A x * X x)" using borel_measurable_indicator assms by force have [simp]: "AE x in M. 0 \ indicat_real A x * X x" using nonneg by fastforce have "integral\<^sup>L (M\A) X = enn2real (\\<^sup>+ x. ennreal (X x) \(M\A))" by (rewrite integral_eq_nn_integral; simp) also have "\ = enn2real (1 / prob A * \\<^sup>+ x. ennreal (X x) \(restrict_space M A))" unfolding cond_prob_space_def apply (rewrite nn_integral_scale_measure, simp add: measurable_restrict_space1) using divide_ennreal emeasure_eq_measure ennreal_1 assms by smt also have "\ = enn2real (1 / prob A * (\\<^sup>+ x. ennreal (indicator A x * X x) \M))" apply (rewrite nn_integral_restrict_space, simp add: assms) by (metis indicator_mult_ennreal mult.commute) also have "\ = integral\<^sup>L M (\x. indicator A x * X x) / prob A" apply (rewrite integral_eq_nn_integral; simp?) by (simp add: divide_nonneg_pos enn2real_mult) finally show ?thesis by simp qed end text \ Define the complementary cumulative distribution function, also known as tail distribution. The theory presented below is a slight modification of the subsection "Properties of cdf's" in the theory "@{text Distribution_Functions}". \ subsection \Complementary Cumulative Distribution Function\ definition ccdf :: "real measure \ real \ real" where "ccdf M \ \x. measure M {x<..}" \ \complementary cumulative distribution function (tail distribution)\ lemma ccdf_def2: "ccdf M x = measure M {x<..}" by (simp add: ccdf_def) context finite_borel_measure begin lemma add_cdf_ccdf: "cdf M x + ccdf M x = measure M (space M)" proof - have "{..x} \ {x<..} = UNIV" by auto moreover have "{..x} \ {x<..} = {}" by auto ultimately have "emeasure M {..x} + emeasure M {x<..} = emeasure M UNIV" using plus_emeasure M_is_borel atMost_borel greaterThan_borel by metis hence "measure M {..x} + measure M {x<..} = measure M UNIV" using finite_emeasure_space emeasure_eq_measure ennreal_inj by (smt (verit, ccfv_SIG) ennreal_plus measure_nonneg) thus ?thesis unfolding cdf_def ccdf_def using borel_UNIV by simp qed lemma ccdf_cdf: "ccdf M = (\x. measure M (space M) - cdf M x)" by (rule ext) (smt add_cdf_ccdf) lemma cdf_ccdf: "cdf M = (\x. measure M (space M) - ccdf M x)" by (rule ext) (smt add_cdf_ccdf) lemma isCont_cdf_ccdf: "isCont (cdf M) x \ isCont (ccdf M) x" proof show "isCont (cdf M) x \ isCont (ccdf M) x" by (rewrite ccdf_cdf) auto next show "isCont (ccdf M) x \ isCont (cdf M) x" by (rewrite cdf_ccdf) auto qed lemma isCont_ccdf: "isCont (ccdf M) x \ measure M {x} = 0" using isCont_cdf_ccdf isCont_cdf by simp lemma continuous_cdf_ccdf: shows "continuous F (cdf M) \ continuous F (ccdf M)" (is "?LHS \ ?RHS") proof assume ?LHS thus ?RHS using continuous_diff continuous_const by (rewrite ccdf_cdf) blast next assume ?RHS thus ?LHS using continuous_diff continuous_const by (rewrite cdf_ccdf) blast qed lemma has_real_derivative_cdf_ccdf: "(cdf M has_real_derivative D) F \ (ccdf M has_real_derivative -D) F" proof assume "(cdf M has_real_derivative D) F" thus "(ccdf M has_real_derivative -D) F" using ccdf_cdf DERIV_const Deriv.field_differentiable_diff by fastforce next assume "(ccdf M has_real_derivative -D) F" thus "(cdf M has_real_derivative D) F" using cdf_ccdf DERIV_const Deriv.field_differentiable_diff by fastforce qed lemma differentiable_cdf_ccdf: "(cdf M differentiable F) \ (ccdf M differentiable F)" unfolding differentiable_def apply (rewrite has_real_derivative_iff[THEN sym])+ apply (rewrite has_real_derivative_cdf_ccdf) by (metis verit_minus_simplify(4)) lemma deriv_cdf_ccdf: assumes "cdf M differentiable at x" shows "deriv (cdf M) x = - deriv (ccdf M) x" using has_real_derivative_cdf_ccdf differentiable_cdf_ccdf assms by (simp add: DERIV_deriv_iff_real_differentiable DERIV_imp_deriv) lemma ccdf_diff_eq2: assumes "x \ y" shows "ccdf M x - ccdf M y = measure M {x<..y}" proof - have "ccdf M x - ccdf M y = cdf M y - cdf M x" using add_cdf_ccdf by (smt (verit)) also have "\ = measure M {x<..y}" using cdf_diff_eq2 assms by simp finally show ?thesis . qed lemma ccdf_nonincreasing: "x \ y \ ccdf M x \ ccdf M y" using add_cdf_ccdf cdf_nondecreasing by smt lemma ccdf_nonneg: "ccdf M x \ 0" using add_cdf_ccdf cdf_bounded by smt lemma ccdf_bounded: "ccdf M x \ measure M (space M)" using add_cdf_ccdf cdf_nonneg by smt lemma ccdf_lim_at_top: "(ccdf M \ 0) at_top" proof - have "((\x. measure M (space M) - cdf M x) \ measure M (space M) - measure M (space M)) at_top" apply (intro tendsto_intros) by (rule cdf_lim_at_top) thus ?thesis by (rewrite ccdf_cdf) simp qed lemma ccdf_lim_at_bot: "(ccdf M \ measure M (space M)) at_bot" proof - have "((\x. measure M (space M) - cdf M x) \ measure M (space M) - 0) at_bot" apply (intro tendsto_intros) by (rule cdf_lim_at_bot) thus ?thesis by (rewrite ccdf_cdf) simp qed lemma ccdf_is_right_cont: "continuous (at_right a) (ccdf M)" proof - have "continuous (at_right a) (\x. measure M (space M) - cdf M x)" apply (intro continuous_intros) by (rule cdf_is_right_cont) thus ?thesis by (rewrite ccdf_cdf) simp qed end context prob_space begin lemma ccdf_distr_measurable [measurable]: assumes [measurable]: "random_variable borel X" shows "ccdf (distr M borel X) \ borel_measurable borel" using real_distribution.finite_borel_measure_M by (rewrite finite_borel_measure.ccdf_cdf; simp) lemma ccdf_distr_P: assumes "random_variable borel X" shows "ccdf (distr M borel X) x = \

(\ in M. X \ > x)" unfolding ccdf_def apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob]) force lemma ccdf_continuous_distr_P_ge: assumes "random_variable borel X" "isCont (ccdf (distr M borel X)) x" shows "ccdf (distr M borel X) x = \

(\ in M. X \ \ x)" proof - have "ccdf (distr M borel X) x = measure (distr M borel X) {x<..}" unfolding ccdf_def by simp also have "\ = measure (distr M borel X) ({x<..} \ {x})" apply (rewrite finite_measure.measure_zero_union, simp_all add: assms finite_measure_distr) using finite_borel_measure.isCont_ccdf real_distribution.finite_borel_measure_M assms by blast also have "\ = measure (distr M borel X) {x..}" by (metis Un_commute ivl_disj_un_singleton(1)) also have "\ = \

(\ in M. X \ \ x)" apply (rewrite measure_distr, simp_all add: assms) unfolding vimage_def by simp (smt (verit) Collect_cong Int_def mem_Collect_eq) finally show ?thesis . qed lemma ccdf_distr_diff_P: assumes "x \ y" and "random_variable borel X" shows "ccdf (distr M borel X) x - ccdf (distr M borel X) y = \

(\ in M. x < X \ \ X \ \ y)" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp have "ccdf (distr M borel X) x - ccdf (distr M borel X) y = measure (distr M borel X) {x<..y}" by (rewrite distrX_FBM.ccdf_diff_eq2; simp add: assms) also have "\ = \

(\ in M. x < X \ \ X \ \ y)" apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob], force) finally show ?thesis . qed end context real_distribution begin lemma ccdf_bounded_prob: "\x. ccdf M x \ 1" by (subst prob_space[THEN sym], rule ccdf_bounded) lemma ccdf_lim_at_bot_prob: "(ccdf M \ 1) at_bot" by (subst prob_space[THEN sym], rule ccdf_lim_at_bot) end text \Introduce the hazard rate. This notion will be used to define the force of mortality.\ subsection \Hazard Rate\ context prob_space begin definition hazard_rate :: "('a \ real) \ real \ real" where "hazard_rate X t \ Lim (at_right 0) (\dt. \

(x in M. t < X x \ X x \ t + dt \ X x > t) / dt)" \ \Here, X is supposed to be a random variable.\ lemma hazard_rate_0_ccdf_0: assumes [measurable]: "random_variable borel X" and "ccdf (distr M borel X) t = 0" shows "hazard_rate X t = 0" \ \Note that division by 0 is calculated as 0 in Isabelle/HOL.\ proof - have "\dt. \

(x in M. t < X x \ X x \ t + dt \ X x > t) = 0" unfolding cond_prob_def using ccdf_distr_P assms by simp hence "hazard_rate X t = Lim (at_right 0) (\dt::real. 0)" unfolding hazard_rate_def by (rewrite Lim_cong; simp) also have "\ = 0" by (rule tendsto_Lim; simp) finally show ?thesis . qed lemma hazard_rate_deriv_cdf: assumes [measurable]: "random_variable borel X" and "(cdf (distr M borel X)) differentiable at t" shows "hazard_rate X t = deriv (cdf (distr M borel X)) t / ccdf (distr M borel X) t" proof (cases \ccdf (distr M borel X) t = 0\) case True with hazard_rate_0_ccdf_0 show ?thesis by simp next case False let ?F = "cdf (distr M borel X)" have "\\<^sub>F dt in at_right 0. \

(x in M. t < X x \ X x \ t + dt \ X x > t) / dt = (?F (t + dt) - ?F t) / dt / ccdf (distr M borel X) t" apply (rule eventually_at_rightI[where b=1]; simp) unfolding cond_prob_def apply (rewrite cdf_distr_diff_P; simp) apply (rewrite ccdf_distr_P[THEN sym], simp) by (smt (verit) Collect_cong mult.commute) moreover have "((\dt. (?F (t + dt) - ?F t) / dt / ccdf (distr M borel X) t) \ deriv ?F t / ccdf (distr M borel X) t) (at_right 0)" apply (rule tendsto_intros, simp_all add: False) apply (rule Lim_at_imp_Lim_at_within) using DERIV_deriv_iff_real_differentiable assms DERIV_def by blast ultimately show ?thesis unfolding hazard_rate_def using tendsto_cong by (intro tendsto_Lim; force) qed lemma deriv_cdf_hazard_rate: assumes [measurable]: "random_variable borel X" and "(cdf (distr M borel X)) differentiable at t" shows "deriv (cdf (distr M borel X)) t = ccdf (distr M borel X) t * hazard_rate X t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp show ?thesis proof (cases \ccdf (distr M borel X) t = 0\) case True hence "cdf (distr M borel X) t = 1" using distrX_FBM.cdf_ccdf by simp (metis assms(1) distrX_FBM.borel_UNIV prob_space.prob_space prob_space_distr) moreover obtain D where "(cdf (distr M borel X) has_real_derivative D) (at t)" using assms real_differentiable_def by atomize_elim blast ultimately have "(cdf (distr M borel X) has_real_derivative 0) (at t)" using assms by (smt (verit) DERIV_local_max real_distribution.cdf_bounded_prob real_distribution_distr) thus ?thesis using True by (simp add: DERIV_imp_deriv) next case False thus ?thesis using hazard_rate_deriv_cdf assms by simp qed qed lemma hazard_rate_deriv_ccdf: assumes [measurable]: "random_variable borel X" and "(ccdf (distr M borel X)) differentiable at t" shows "hazard_rate X t = - deriv (ccdf (distr M borel X)) t / ccdf (distr M borel X) t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp show ?thesis using hazard_rate_deriv_cdf distrX_FBM.deriv_cdf_ccdf assms distrX_FBM.differentiable_cdf_ccdf by presburger qed lemma deriv_ccdf_hazard_rate: assumes [measurable]: "random_variable borel X" and "(ccdf (distr M borel X)) differentiable at t" shows "deriv (ccdf (distr M borel X)) t = - ccdf (distr M borel X) t * hazard_rate X t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp show ?thesis using deriv_cdf_hazard_rate distrX_FBM.deriv_cdf_ccdf assms distrX_FBM.differentiable_cdf_ccdf by simp qed lemma hazard_rate_deriv_ln_ccdf: assumes [measurable]: "random_variable borel X" and "(ccdf (distr M borel X)) differentiable at t" and "ccdf (distr M borel X) t \ 0" shows "hazard_rate X t = - deriv (\t. ln (ccdf (distr M borel X) t)) t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp let ?srvl = "ccdf (distr M borel X)" have "?srvl t > 0" using distrX_FBM.ccdf_nonneg assms by (smt (verit)) moreover have "(?srvl has_real_derivative (deriv ?srvl t)) (at t)" using DERIV_deriv_iff_real_differentiable assms by blast ultimately have "((\t. ln (?srvl t)) has_real_derivative 1 / ?srvl t * deriv ?srvl t) (at t)" by (rule derivative_intros) hence "deriv (\t. ln (?srvl t)) t = deriv ?srvl t / ?srvl t" by (simp add: DERIV_imp_deriv) also have "\ = - hazard_rate X t" using hazard_rate_deriv_ccdf assms by simp finally show ?thesis by simp qed lemma hazard_rate_has_integral: assumes [measurable]: "random_variable borel X" and "t \ u" and "(ccdf (distr M borel X)) piecewise_differentiable_on {t<..s. s \ {t..u} \ ccdf (distr M borel X) s \ 0" shows "(hazard_rate X has_integral ln (ccdf (distr M borel X) t / ccdf (distr M borel X) u)) {t..u}" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp let ?srvl = "ccdf (distr M borel X)" have [simp]: "\s. t \ s \ s \ u \ ?srvl s > 0" using distrX_FBM.ccdf_nonneg assms by (smt (verit) atLeastAtMost_iff) have "(deriv (\s. - ln (?srvl s)) has_integral - ln (?srvl u) - - ln (?srvl t)) {t..u}" proof - have "continuous_on {t..u} (\s. - ln (?srvl s))" by (rule continuous_intros, rule continuous_on_ln, auto simp add: assms) moreover hence "(\s. - ln (?srvl s)) piecewise_differentiable_on {t<.. {0<..}" proof - { fix s assume "s \ {t<.. 0" using assms by simp moreover have "?srvl s \ 0" using distrX_FBM.ccdf_nonneg by simp ultimately have "?srvl s > 0" by simp } thus ?thesis by auto qed hence "(\r. - ln r) \ ?srvl piecewise_differentiable_on {t<..s. - ln (?srvl s)) has_integral ln (?srvl t / ?srvl u)) {t..u}" by simp (rewrite ln_div; force simp: assms) thus "((hazard_rate X) has_integral ln (?srvl t / ?srvl u)) {t..u}" proof - from assms obtain S0 where finS0: "finite S0" and diffS0: "\s. s \ {t<.. ?srvl differentiable at s within {t<..s. s \ {t..u} - S \ ?srvl differentiable at s" proof (atomize_elim) let ?S = "S0 \ {t, u}" have "finite ?S" using finS0 by simp moreover have "\s. s \ {t..u} - ?S \ ccdf (distr M borel X) differentiable at s" proof - { fix s assume s_in: "s \ {t..u} - ?S" hence "?srvl differentiable at s within {t<..S. finite S \ (\s. s \ {t..u} - S \ ccdf (distr M borel X) differentiable at s)" by blast qed thus ?thesis apply (rewrite has_integral_spike_finite_eq [of S _ "deriv (\s. - ln (?srvl s))"], simp) apply (rewrite hazard_rate_deriv_ln_ccdf, simp_all add: assms) apply (rewrite deriv_minus, simp_all) apply (rewrite in asm differentiable_eq_field_differentiable_real) apply (rewrite comp_def[THEN sym], rule field_differentiable_compose[of "?srvl"], simp_all) unfolding field_differentiable_def apply (rule exI, rule DERIV_ln, simp) using ln by simp qed qed corollary hazard_rate_integrable: assumes [measurable]: "random_variable borel X" and "t \ u" and "(ccdf (distr M borel X)) piecewise_differentiable_on {t<..s. s \ {t..u} \ ccdf (distr M borel X) s \ 0" shows "hazard_rate X integrable_on {t..u}" using has_integral_integrable hazard_rate_has_integral assms by blast lemma ccdf_exp_cumulative_hazard: assumes [measurable]: "random_variable borel X" and "t \ u" and "(ccdf (distr M borel X)) piecewise_differentiable_on {t<..s. s \ {t..u} \ ccdf (distr M borel X) s \ 0" shows "ccdf (distr M borel X) u / ccdf (distr M borel X) t = exp (- integral {t..u} (hazard_rate X))" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp let ?srvl = "ccdf (distr M borel X)" have [simp]: "\s. t \ s \ s \ u \ ?srvl s > 0" using distrX_FBM.ccdf_nonneg assms by (smt (verit) atLeastAtMost_iff) have "integral {t..u} (hazard_rate X) = ln (?srvl t / ?srvl u)" using hazard_rate_has_integral has_integral_integrable_integral assms by auto also have "\ = - ln (?srvl u / ?srvl t)" using ln_div assms by simp finally have "- integral {t..u} (hazard_rate X) = ln (?srvl u / ?srvl t)" by simp thus ?thesis using assms by simp qed lemma hazard_rate_density_ccdf: assumes "distributed M lborel X f" and "\s. f s \ 0" "t < u" "continuous_on {t..u} f" shows "hazard_rate X t = f t / ccdf (distr M borel X) t" proof (cases \ccdf (distr M borel X) t = 0\) case True thus ?thesis apply (rewrite hazard_rate_0_ccdf_0, simp_all) using assms(1) distributed_measurable by force next case False have [simp]: "t \ u" using assms by simp have [measurable]: "random_variable borel X" using assms distributed_measurable measurable_lborel1 by blast have [measurable]: "f \ borel_measurable lborel" using assms distributed_real_measurable by metis have [simp]: "integrable lborel f" proof - have "prob (X -` UNIV \ space M) = LINT x|lborel. indicat_real UNIV x * f x" by (rule distributed_measure; simp add: assms) thus ?thesis using prob_space not_integrable_integral_eq by fastforce qed have "((\dt. (LBINT s:{t..t+dt}. f s) / dt) \ f t) (at_right 0)" proof - have "\dt. (\\<^sup>+ x. ennreal (indicat_real {t..t+dt} x * f x) \lborel) < \" proof - fix dt :: real have "(\\<^sup>+ x. ennreal (indicat_real {t..t+dt} x * f x) \lborel) = set_nn_integral lborel {t..t+dt} f" by (metis indicator_mult_ennreal mult.commute) moreover have "emeasure M (X -` {t..t+dt} \ space M) = set_nn_integral lborel {t..t+dt} f" by (rule distributed_emeasure; simp add: assms) moreover have "emeasure M (X -` {t..t+dt} \ space M) < \" using emeasure_eq_measure ennreal_less_top infinity_ennreal_def by presburger ultimately show "(\\<^sup>+ x. ennreal (indicat_real {t..t+dt} x * f x) \lborel) < \" by simp qed hence "\dt. LBINT s:{t..t+dt}. f s = integral {t..t+dt} f" apply (intro set_borel_integral_eq_integral) unfolding set_integrable_def - apply (rule integrableI_real_bounded; simp) + apply (rule integrableI_nonneg; simp) by (rule AE_I2, simp add: assms) moreover have "((\dt. (integral {t..t+dt} f) / dt) \ f t) (at_right 0)" proof - have "((\x. integral {t..x} f) has_real_derivative f t) (at t within {t..u})" by (rule integral_has_real_derivative; simp add: assms) moreover have "(at t within {t..u}) = (at (t+0) within (+)t ` {0..u-t})" by simp ultimately have "((\x. integral {t..x} f) \ (+)t has_real_derivative f t) (at 0 within {0..u-t})" by (metis DERIV_at_within_shift_lemma) hence "((\dt. (integral {t..t+dt} f) / dt) \ f t) (at 0 within {0..u-t})" using has_field_derivative_iff by force thus ?thesis using at_within_Icc_at_right assms by smt qed ultimately show ?thesis by simp qed moreover have "\dt. dt > 0 \ \

(x in M. X x \ {t <.. t+dt}) = LBINT s:{t..t+dt}. f s" proof - fix dt :: real assume "dt > 0" hence [simp]: "sym_diff {t<..t + dt} {t..t + dt} = {t}" by force have "prob (X -` {t<..t+dt} \ space M) = \s. indicator {t<..t+dt} s * f s \lborel" by (rule distributed_measure; simp add: assms) hence "\

(x in M. X x \ {t <.. t+dt}) = LBINT s:{t<..t+dt}. f s" unfolding set_lebesgue_integral_def vimage_def Int_def by simp (smt (verit) Collect_cong) moreover have "LBINT s:{t<..t+dt}. f s = LBINT s:{t..t+dt}. f s" by (rule set_integral_null_delta; force) ultimately show "\

(x in M. X x \ {t <.. t+dt}) = LBINT s:{t..t+dt}. f s" by simp qed ultimately have "((\dt. \

(x in M. t < X x \ X x \ t + dt) / dt) \ f t) (at_right 0)" by simp (smt (verit) Lim_cong_within greaterThan_iff) hence "((\dt. \

(x in M. t < X x \ X x \ t + dt \ X x > t) / dt) \ f t / ccdf (distr M borel X) t) (at_right 0)" unfolding cond_prob_def apply (rewrite ccdf_distr_P[THEN sym]; simp) apply (rewrite mult.commute, rewrite divide_divide_eq_left[THEN sym]) by (rule tendsto_intros; (simp add: False)?) (smt (verit) Collect_cong Lim_cong_within) thus ?thesis unfolding hazard_rate_def by (intro tendsto_Lim; simp) qed end end diff --git a/thys/Actuarial_Mathematics/Survival_Model.thy b/thys/Actuarial_Mathematics/Survival_Model.thy --- a/thys/Actuarial_Mathematics/Survival_Model.thy +++ b/thys/Actuarial_Mathematics/Survival_Model.thy @@ -1,2732 +1,2732 @@ theory Survival_Model imports "HOL-Library.Rewrite" "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Extended_Real" "HOL-Probability.Probability" Preliminaries begin section \Survival Model\ text \ The survival model is built on the probability space @{text "\"}. Additionally, the random variable @{text "X : space \ \ \"} is introduced, which represents the age at death. \ locale prob_space_actuary = MM_PS: prob_space \ for \ \ \Since the letter M may be used as a commutation function, adopt the letter @{text "\"} instead as a notation for the measure space.\ locale survival_model = prob_space_actuary + fixes X :: "'a \ real" assumes X_RV[simp]: "MM_PS.random_variable (borel :: real measure) X" and X_pos_AE[simp]: "AE \ in \. X \ > 0" begin subsection \General Theory of Survival Model\ interpretation distrX_RD: real_distribution "distr \ borel X" using MM_PS.real_distribution_distr by simp lemma X_le_event[simp]: "{\ \ space \. X \ \ x} \ MM_PS.events" by measurable simp lemma X_gt_event[simp]: "{\ \ space \. X \ > x} \ MM_PS.events" by measurable simp lemma X_compl_le_gt: "space \ - {\ \ space \. X \ \ x} = {\ \ space \. X \ > x}" for x::real proof - have "space \ - {\ \ space \. X \ \ x} = space \ - (X -` {..x})" by blast also have "\ = (X -` {x<..}) \ space \" using vimage_compl_atMost by fastforce also have "\ = {\ \ space \. X \ > x}" by blast finally show ?thesis . qed lemma X_compl_gt_le: "space \ - {\ \ space \. X \ > x} = {\ \ space \. X \ \ x}" for x::real using X_compl_le_gt by blast subsubsection \Introduction of Survival Function for X\ text \Note that @{text "ccdf (distr \ borel X)"} is the survival (distributive) function for X.\ lemma ccdfX_0_1: "ccdf (distr \ borel X) 0 = 1" apply (rewrite MM_PS.ccdf_distr_P, simp) using X_pos_AE MM_PS.prob_space using MM_PS.prob_Collect_eq_1 X_gt_event by presburger lemma ccdfX_unborn_1: "ccdf (distr \ borel X) x = 1" if "x \ 0" proof (rule antisym) show "ccdf (distr \ borel X) x \ 1" using MM_PS.ccdf_distr_P by simp show "ccdf (distr \ borel X) x \ 1" proof - have "ccdf (distr \ borel X) x \ ccdf (distr \ borel X) 0" using finite_borel_measure.ccdf_nonincreasing distrX_RD.finite_borel_measure_M that by simp also have "ccdf (distr \ borel X) 0 = 1" using ccdfX_0_1 that by simp finally show ?thesis . qed qed definition death_pt :: ereal ("$\") where "$\ \ Inf (ereal ` {x \ \. ccdf (distr \ borel X) x = 0})" \ \This is my original notation, which is used to develop life insurance mathematics rigorously.\ lemma psi_nonneg: "$\ \ 0" unfolding death_pt_def proof (rule Inf_greatest) fix x'::ereal assume "x' \ ereal ` {x \ \. ccdf (distr \ borel X) x = 0}" then obtain x::real where "x' = ereal x" and "ccdf (distr \ borel X) x = 0" by blast hence "ccdf (distr \ borel X) 0 > ccdf (distr \ borel X) x" using ccdfX_0_1 X_pos_AE by simp hence "x \ 0" using mono_invE finite_borel_measure.ccdf_nonincreasing distrX_RD.finite_borel_measure_M X_RV by (smt(verit)) thus "x' \ 0" using \x' = ereal x\ by simp qed lemma ccdfX_beyond_0: "ccdf (distr \ borel X) x = 0" if "x > $\" for x::real proof - have "ereal ` {y \ \. ccdf (distr \ borel X) y = 0} \ {}" using death_pt_def that by force hence "\y'\(ereal ` {y \ \. ccdf (distr \ borel X) y = 0}). y' < ereal x" using that unfolding death_pt_def by (rule cInf_lessD) then obtain "y'" where "y' \ (ereal ` {y \ \. ccdf (distr \ borel X) y = 0})" and "y' < ereal x" by blast then obtain y::real where "y' = ereal y" and "ccdf (distr \ borel X) y = 0" and "ereal y < ereal x" by blast hence "ccdf (distr \ borel X) y = 0" and "y < x" by simp_all hence "ccdf (distr \ borel X) x \ 0" using finite_borel_measure.ccdf_nonincreasing distrX_RD.finite_borel_measure_M X_RV by (metis order_less_le) thus ?thesis using finite_borel_measure.ccdf_nonneg distrX_RD.finite_borel_measure_M X_RV by smt qed lemma ccdfX_psi_0: "ccdf (distr \ borel X) (real_of_ereal $\) = 0" if "$\ < \" proof - have "\$\\ \ \" using that psi_nonneg by simp then obtain u::real where "$\ = ereal u" using ereal_real' by blast hence "real_of_ereal $\ = u" by simp moreover have "ccdf (distr \ borel X) u = 0" proof - have "\x::real. x \ u \ x \ {u<..} \ ccdf (distr \ borel X) x = 0" by (rule ccdfX_beyond_0, simp add: \$\ = ereal u\) hence "(ccdf (distr \ borel X) \ 0) (at_right u)" apply - by (rule iffD2[OF Lim_cong_within[where ?g="(\x.0)"]], simp_all+) moreover have "(ccdf (distr \ borel X) \ ccdf (distr \ borel X) u) (at_right u)" using finite_borel_measure.ccdf_is_right_cont distrX_RD.finite_borel_measure_M continuous_within X_RV by blast ultimately show ?thesis using tendsto_unique trivial_limit_at_right_real by blast qed ultimately show ?thesis by simp qed lemma ccdfX_0_equiv: "ccdf (distr \ borel X) x = 0 \ x \ $\" for x::real proof assume "ccdf (distr \ borel X) x = 0" thus "ereal x \ $\" unfolding death_pt_def by (simp add: INF_lower) next assume "$\ \ ereal x" hence "$\ = ereal x \ $\ < ereal x" unfolding less_eq_ereal_def by auto thus "ccdf (distr \ borel X) x = 0" proof assume \: "$\ = ereal x" hence "$\ < \" by simp moreover have "real_of_ereal $\ = x" using \ by simp ultimately show "ccdf (distr \ borel X) x = 0" using ccdfX_psi_0 by simp next assume "$\ < ereal x" thus "ccdf (distr \ borel X) x = 0" by (rule ccdfX_beyond_0) qed qed lemma psi_pos[simp]: "$\ > 0" proof (rule not_le_imp_less, rule notI) show "$\ \ (0::ereal) \ False" proof - assume "$\ \ (0::ereal)" hence "ccdf (distr \ borel X) 0 = 0" using ccdfX_0_equiv by (simp add: zero_ereal_def) moreover have "ccdf (distr \ borel X) 0 = 1" using ccdfX_0_1 by simp ultimately show "False" by simp qed qed corollary psi_pos'[simp]: "$\ > ereal 0" using psi_pos zero_ereal_def by presburger subsubsection \Introdution of Future-Lifetime Random Variable T(x)\ definition alive :: "real \ 'a set" where "alive x \ {\ \ space \. X \ > x}" lemma alive_event[simp]: "alive x \ MM_PS.events" for x::real unfolding alive_def by simp lemma X_alivex_measurable[measurable, simp]: "X \ borel_measurable (\ \ alive x)" for x::real unfolding cond_prob_space_def by (measurable; simp add: measurable_restrict_space1) definition futr_life :: "real \ ('a \ real)" ("T") where "T x \ (\\. X \ - x)" \ \Note that @{text "T(x) : space \ \ \"} represents the time until death of a person aged x.\ lemma T0_eq_X[simp]: "T 0 = X" unfolding futr_life_def by simp lemma Tx_measurable[measurable, simp]: "T x \ borel_measurable \" for x::real unfolding futr_life_def by (simp add: borel_measurable_diff) lemma Tx_alivex_measurable[measurable, simp]: "T x \ borel_measurable (\ \ alive x)" for x::real unfolding futr_life_def by (simp add: borel_measurable_diff) lemma alive_T: "alive x = {\ \ space \. T x \ > 0}" for x::real unfolding alive_def futr_life_def by force lemma alivex_Tx_pos[simp]: "0 < T x \" if "\ \ space (\ \ alive x)" for x::real using MM_PS.space_cond_prob_space alive_T that by simp lemma PT0_eq_PX_lborel: "\

(\ in \. T 0 \ \ A \ T 0 \ > 0) = \

(\ in \. X \ \ A)" if "A \ sets lborel" for A :: "real set" apply (rewrite MM_PS.cond_prob_AE_prob, simp_all) using that X_RV measurable_lborel1 predE pred_sets2 by blast subsubsection \Actuarial Notations on the Survival Model\ definition survive :: "real \ real \ real" ("$p'_{_&_}" [0,0] 200) where "$p_{t&x} \ ccdf (distr (\ \ alive x) borel (T x)) t" \ \the probability that a person aged x will survive for t years\ \ \Note that the function @{text "$p_{\&x}"} is the survival function on @{text "(\ \ alive x)"} for the random variable T(x).\ \ \The parameter t is usually nonnegative, but theoretically it can be negative.\ abbreviation survive_1 :: "real \ real" ("$p'__" [101] 200) where "$p_x \ $p_{1&x}" definition die :: "real \ real \ real" ("$q'_{_&_}" [0,0] 200) where "$q_{t&x} \ cdf (distr (\ \ alive x) borel (T x)) t" \ \the probability that a person aged x will die within t years\ \ \Note that the function @{text "$q_{\&x}"} is the cumulative distributive function on @{text "(\ \ alive x)"} for the random variable T(x).\ \ \The parameter t is usually nonnegative, but theoretically it can be negative.\ abbreviation die_1 :: "real \ real" ("$q'__" [101] 200) where "$q_x \ $q_{1&x}" definition die_defer :: "real \ real \ real \ real" ("$q'_{_\_&_}" [0,0,0] 200) where "$q_{f\t&x} = \$q_{f+t&x} - $q_{f&x}\" \ \the probability that a person aged x will die within t years, deferred f years\ \ \The parameters f and t are usually nonnegative, but theoretically they can be negative.\ abbreviation die_defer_1 :: "real \ real \ real" ("$q'_{_\&_}" [0,0] 200) where "$q_{f\&x} \ $q_{f\1&x}" definition life_expect :: "real \ real" ("$e`\'__" [101] 200) where "$e`\_x \ integral\<^sup>L (\ \ alive x) (T x)" \ \complete life expectation\ \ \Note that @{text "$e`\_x"} is calculated as 0 when @{text "nn_integral (\ \ alve x) (T x) = \"}.\ definition temp_life_expect :: "real \ real \real" ("$e`\'_{_:_}" [0,0] 200) where "$e`\_{x:n} \ integral\<^sup>L (\ \ alive x) (\\. min (T x \) n)" \ \temporary complete life expectation\ definition curt_life_expect :: "real \ real" ("$e'__" [101] 200) where "$e_x \ integral\<^sup>L (\ \ alive x) (\\. \T x \\)" \ \curtate life expectation\ \ \Note that @{text "$e_x"} is calculated as 0 when @{text "nn_integral (\ \ alive x) \T x\ = \"}.\ definition temp_curt_life_expect :: "real \ real \ real" ("$e'_{_:_}" [0,0] 200) where "$e_{x:n} \ integral\<^sup>L (\ \ alive x) (\\. \min (T x \) n\)" \ \temporary curtate life expectation\ \ \In the definition n can be a real number, but in practice n is usually a natural number.\ subsubsection \Properties of Survival Function for T(x)\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin lemma PXx_pos[simp]: "\

(\ in \. X \ > x) > 0" proof - have "\

(\ in \. X \ > x) = ccdf (distr \ borel X) x" unfolding alive_def using MM_PS.ccdf_distr_P by simp also have "\ > 0" using ccdfX_0_equiv distrX_RD.ccdf_nonneg x_lt_psi by (smt (verit) linorder_not_le) finally show ?thesis . qed lemma PTx_pos[simp]: "\

(\ in \. T x \ > 0) > 0" apply (rewrite alive_T[THEN sym]) unfolding alive_def by simp interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct, simp_all add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma ccdfTx_cond_prob: "ccdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. T x \ > t \ T x \ > 0)" for t::real apply (rewrite alivex_PS.ccdf_distr_P, simp) unfolding alive_def apply (rewrite MM_PS.cond_prob_space_cond_prob[THEN sym], simp_all add: pred_def) unfolding futr_life_def by simp lemma ccdfTx_0_1: "ccdf (distr (\ \ alive x) borel (T x)) 0 = 1" apply (rewrite ccdfTx_cond_prob) unfolding futr_life_def cond_prob_def by (smt (verit, best) Collect_cong PXx_pos divide_eq_1_iff) lemma ccdfTx_nonpos_1: "ccdf (distr (\ \ alive x) borel (T x)) t = 1" if "t \ 0" for t :: real using antisym ccdfTx_0_1 that by (metis distrTx_RD.ccdf_bounded_prob distrTx_RD.ccdf_nonincreasing) lemma ccdfTx_0_equiv: "ccdf (distr (\ \ alive x) borel (T x)) t = 0 \ x+t \ $\" for t::real proof - have "ccdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. X \ > x+t \ X \ > x) / \

(\ in \. X \ > x)" apply (rewrite ccdfTx_cond_prob) unfolding cond_prob_def futr_life_def by (smt (verit) Collect_cong) hence "ccdf (distr (\ \ alive x) borel (T x)) t = 0 \ \

(\ in \. X \ > x+t \ X \ > x) / \

(\ in \. X \ > x) = 0" by simp also have "\ \ \

(\ in \. X \ > x+t \ X \ > x) = 0" using x_lt_psi PXx_pos by (smt (verit) divide_eq_0_iff) also have "\ \ x+t \ $\" using ccdfX_0_equiv MM_PS.ccdf_distr_P by (smt (verit) Collect_cong X_RV le_ereal_le linorder_not_le x_lt_psi) finally show ?thesis . qed lemma ccdfTx_continuous_on_nonpos[simp]: "continuous_on {..0} (ccdf (distr (\ \ alive x) borel (T x)))" by (metis atMost_iff ccdfTx_nonpos_1 continuous_on_cong continuous_on_const) lemma ccdfTx_differentiable_on_nonpos[simp]: "(ccdf (distr (\ \ alive x) borel (T x))) differentiable_on {..0}" by (rewrite differentiable_on_cong[where f="\_. 1"]; simp add: ccdfTx_nonpos_1) lemma ccdfTx_has_real_derivative_0_at_neg: "(ccdf (distr (\ \ alive x) borel (T x)) has_real_derivative 0) (at t)" if "t < 0" for t::real apply (rewrite has_real_derivative_iff_has_vector_derivative) apply (rule has_vector_derivative_transform_within_open[of "\_. 1" _ _ "{..<0}"]) using ccdfTx_nonpos_1 that by simp_all lemma ccdfTx_integrable_Icc: "set_integrable lborel {a..b} (ccdf (distr (\ \ alive x) borel (T x)))" for a b :: real proof - have "(\\<^sup>+t. ennreal (indicat_real {a..b} t * ccdf (distr (\ \ alive x) borel (T x)) t) \lborel) < \" proof - have "(\\<^sup>+t. ennreal (indicat_real {a..b} t * ccdf (distr (\ \ alive x) borel (T x)) t) \lborel) \ (\\<^sup>+t. ennreal (indicat_real {a..b} t) \lborel)" apply (rule nn_integral_mono) using distrTx_RD.ccdf_bounded by (simp add: distrTx_RD.ccdf_bounded_prob indicator_times_eq_if(1)) also have "\ = nn_integral lborel (indicator {a..b})" by (meson ennreal_indicator) also have "\ = emeasure lborel {a..b}" by (rewrite nn_integral_indicator; simp) also have "\ < \" using emeasure_lborel_Icc_eq ennreal_less_top infinity_ennreal_def by presburger finally show ?thesis . qed thus ?thesis unfolding set_integrable_def apply (intro integrableI_nonneg, simp_all) using distrTx_RD.ccdf_nonneg by (intro always_eventually) auto qed corollary ccdfTx_integrable_on_Icc: "ccdf (distr (\ \ alive x) borel (T x)) integrable_on {a..b}" for a b :: real using set_borel_integral_eq_integral ccdfTx_integrable_Icc by force lemma ccdfTx_PX: "ccdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. X \ > x+t) / \

(\ in \. X \ > x)" if "t \ 0" for t::real apply (rewrite ccdfTx_cond_prob) unfolding cond_prob_def futr_life_def PXx_pos by (smt (verit) Collect_cong that) lemma ccdfTx_ccdfX: "ccdf (distr (\ \ alive x) borel (T x)) t = ccdf (distr \ borel X) (x + t) / ccdf (distr \ borel X) x" if "t \ 0" for t::real using ccdfTx_PX that MM_PS.ccdf_distr_P X_RV by presburger lemma ccdfT0_eq_ccdfX: "ccdf (distr (\ \ alive 0) borel (T 0)) = ccdf (distr \ borel X)" proof fix x show "ccdf (distr (\ \ alive 0) borel (T 0)) x = ccdf (distr \ borel X) x" proof (cases \x \ 0\) case True thus ?thesis using survival_model.ccdfTx_ccdfX[where x=0] ccdfX_0_1 survival_model_axioms by fastforce next case False hence "x \ 0" by simp thus ?thesis apply (rewrite ccdfX_unborn_1, simp) by (rewrite survival_model.ccdfTx_nonpos_1; simp add: survival_model_axioms) qed qed lemma continuous_ccdfX_ccdfTx: "continuous (at (x+t) within {x..}) (ccdf (distr \ borel X)) \ continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" if "t \ 0" for t::real proof - let ?srvl = "ccdf (distr \ borel X)" have [simp]: "\

(\ in \. X \ > x) \ 0" using PXx_pos by force have \: "\u. u \ 0 \ ccdf (distr (\ \ alive x) borel (T x)) u = ?srvl (x + u) / \

(\ in \. X \ > x)" using survive_def MM_PS.ccdf_distr_P ccdfTx_PX that by simp have "continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x))) \ continuous (at t within {0..}) (\u. ?srvl (x+u) / \

(\ in \. x < X \))" proof - have "\\<^sub>F u in at t within {0..}. ccdf (distr (\ \ alive x) borel (T x)) u = ?srvl (x+u) / \

(\ in \. X \ > x)" using \ by (rewrite eventually_at_topological, simp_all) blast thus ?thesis by (intro continuous_at_within_cong, simp_all add: \ that) qed also have "\ \ continuous (at t within {0..}) (\u. ?srvl (x+u))" by (rewrite at "_ = \" continuous_cdivide_iff[of "\

(\ in \. X \ > x)"], simp_all) also have "\ \ continuous (at (x+t) within {x..}) ?srvl" proof let ?subx = "\v. v-x" assume LHS: "continuous (at t within {0..}) (\u. ?srvl (x+u))" hence "continuous (at (?subx (x+t)) within ?subx ` {x..}) (\u. ?srvl (x+u))" proof - have "?subx ` {x..} = {0..}" by (metis (no_types, lifting) add.commute add_uminus_conv_diff diff_self image_add_atLeast image_cong) thus ?thesis using LHS by simp qed moreover have "continuous (at (x+t) within {x..}) ?subx" by (simp add: continuous_diff) ultimately have "continuous (at (x+t) within {x..}) (\u. ?srvl (x + (?subx u)))" using continuous_within_compose2 by force thus "continuous (at (x+t) within {x..}) ?srvl" by simp next assume RHS: "continuous (at (x+t) within {x..}) ?srvl" hence "continuous (at ((plus x) t) within (plus x) ` {0..}) ?srvl" by simp moreover have "continuous (at t within {0..}) (plus x)" by (simp add: continuous_add) ultimately show "continuous (at t within {0..}) (\u. ?srvl (x+u))" using continuous_within_compose2 by force qed finally show ?thesis by simp qed lemma isCont_ccdfX_ccdfTx: "isCont (ccdf (distr \ borel X)) (x+t) \ isCont (ccdf (distr (\ \ alive x) borel (T x))) t" if "t > 0" for t::real proof - have "isCont (ccdf (distr \ borel X)) (x+t) \ continuous (at (x+t) within {x<..}) (ccdf (distr \ borel X))" by (smt (verit) at_within_open greaterThan_iff open_greaterThan that) also have "\ \ continuous (at (x+t) within {x..}) (ccdf (distr \ borel X))" by (meson Ioi_le_Ico calculation continuous_within_subset top_greatest) also have "\ \ continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" using that continuous_ccdfX_ccdfTx by force also have "\ \ continuous (at t within {0<..}) (ccdf (distr (\ \ alive x) borel (T x)))" by (metis Ioi_le_Ico at_within_open continuous_at_imp_continuous_at_within continuous_within_subset greaterThan_iff open_greaterThan that) also have "\ \ isCont (ccdf (distr (\ \ alive x) borel (T x))) t" by (metis at_within_open greaterThan_iff open_greaterThan that) finally show ?thesis . qed lemma has_real_derivative_ccdfX_ccdfTx: "((ccdf (distr \ borel X)) has_real_derivative D) (at (x+t)) \ ((ccdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" if "t > 0" for t D :: real proof - have "((ccdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t) \ ((\t. (ccdf (distr \ borel X)) (x+t) / \

(\ in \. X \ > x)) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" proof - let ?d = "t/2" { fix u::real assume "dist u t < ?d" hence "u > 0" by (smt (verit) dist_real_def dist_triangle_half_r) hence "ccdf (distr (\ \ alive x) borel (T x)) u = ccdf (distr \ borel X) (x + u) / MM_PS.prob {\::'a \ space \. x < X \}" using survive_def MM_PS.ccdf_distr_P ccdfTx_PX that by simp } moreover have "?d > 0" using that by simp ultimately show ?thesis apply - apply (rule DERIV_cong_ev, simp) apply (rewrite eventually_nhds_metric, blast) by simp qed also have "\ \ ((\t. (ccdf (distr \ borel X)) (x+t)) has_real_derivative D) (at t)" using PXx_pos by (rewrite DERIV_cdivide_iff[of "\

(\ in \. X \ > x)", THEN sym]; force) also have "\ \ (ccdf (distr \ borel X) has_real_derivative D) (at (x+t))" by (simp add: DERIV_shift add.commute) finally show ?thesis by simp qed lemma differentiable_ccdfX_ccdfTx: "(ccdf (distr \ borel X)) differentiable at (x+t) \ (ccdf (distr (\ \ alive x) borel (T x))) differentiable at t" if "t > 0" for t::real apply (rewrite differentiable_eq_field_differentiable_real)+ unfolding field_differentiable_def using has_real_derivative_ccdfX_ccdfTx that by (smt (verit, del_insts) PXx_pos nonzero_mult_div_cancel_left) subsubsection \Properties of @{text "$p_{t&x}"}\ lemma p_0_1: "$p_{0&x} = 1" unfolding survive_def using ccdfTx_0_1 by simp lemma p_nonneg[simp]: "$p_{t&x} \ 0" for t::real unfolding survive_def using distrTx_RD.ccdf_nonneg by simp lemma p_0_equiv: "$p_{t&x} = 0 \ x+t \ $\" for t::real unfolding survive_def by (rule ccdfTx_0_equiv) lemma p_PTx: "$p_{t&x} = \

(\ in \. T x \ > t \ T x \ > 0)" for t::real unfolding survive_def using ccdfTx_cond_prob by simp lemma p_PX: "$p_{t&x} = \

(\ in \. X \ > x + t) / \

(\ in \. X \ > x)" if "t \ 0" for t::real unfolding survive_def using ccdfTx_PX that by simp lemma p_mult: "$p_{t+t' & x} = $p_{t&x} * $p_{t' & x+t}" if "t \ 0" "t' \ 0" "x+t < $\" for t t' :: real proof - have "$p_{t+t' & x} = \

(\ in \. X \ > x+t+t') / \

(\ in \. X \ > x)" apply (rewrite p_PX; (simp add: that)?) by (rule disjI2, smt (verit, best) Collect_cong) also have "\ = (\

(\ in \. X \ > x+t+t') / \

(\ in \. X \ > x+t)) * (\

(\ in \. X \ > x+t) / \

(\ in \. X \ > x))" using that survival_model.PXx_pos survival_model_axioms by fastforce also have "\ = $p_{t&x} * $p_{t' & x+t}" apply (rewrite p_PX, simp add: that) by (rewrite survival_model.p_PX, simp_all add: that survival_model_axioms) finally show ?thesis . qed lemma p_PTx_ge_ccdf_isCont: "$p_{t&x} = \

(\ in \. T x \ \ t \ T x \ > 0)" if "isCont (ccdf (distr \ borel X)) (x+t)" "t > 0" for t::real unfolding survive_def using that isCont_ccdfX_ccdfTx apply (rewrite alivex_PS.ccdf_continuous_distr_P_ge, simp_all) by (rewrite MM_PS.cond_prob_space_cond_prob, simp_all add: alive_T) end subsubsection \Properties of Survival Function for X\ lemma ccdfX_continuous_unborn[simp]: "continuous_on {..0} (ccdf (distr \ borel X))" using ccdfTx_continuous_on_nonpos by (metis ccdfT0_eq_ccdfX psi_pos') lemma ccdfX_differentiable_unborn[simp]: "(ccdf (distr \ borel X)) differentiable_on {..0}" using ccdfTx_differentiable_on_nonpos by (metis ccdfT0_eq_ccdfX psi_pos') lemma ccdfX_has_real_derivative_0_unborn: "(ccdf (distr \ borel X) has_real_derivative 0) (at x)" if "x < 0" for x::real using ccdfTx_has_real_derivative_0_at_neg by (metis ccdfT0_eq_ccdfX psi_pos' that) lemma ccdfX_integrable_Icc: "set_integrable lborel {a..b} (ccdf (distr \ borel X))" for a b :: real using ccdfTx_integrable_Icc by (metis ccdfT0_eq_ccdfX psi_pos') corollary ccdfX_integrable_on_Icc: "ccdf (distr \ borel X) integrable_on {a..b}" for a b :: real using set_borel_integral_eq_integral ccdfX_integrable_Icc by force lemma ccdfX_p: "ccdf (distr \ borel X) x = $p_{x&0}" for x::real by (metis ccdfT0_eq_ccdfX survive_def psi_pos') subsubsection \Introduction of Cumulative Distributive Function for X\ lemma cdfX_0_0: "cdf (distr \ borel X) 0 = 0" using ccdfX_0_1 distrX_RD.ccdf_cdf distrX_RD.prob_space by fastforce lemma cdfX_unborn_0: "cdf (distr \ borel X) x = 0" if "x \ 0" using ccdfX_unborn_1 cdfX_0_0 distrX_RD.cdf_ccdf that by fastforce lemma cdfX_beyond_1: "cdf (distr \ borel X) x = 1" if "x > $\" for x::real using ccdfX_beyond_0 distrX_RD.cdf_ccdf that distrX_RD.prob_space by force lemma cdfX_psi_1: "cdf (distr \ borel X) (real_of_ereal $\) = 1" if "$\ < \" using ccdfX_psi_0 distrX_RD.cdf_ccdf distrX_RD.prob_space that by fastforce lemma cdfX_1_equiv: "cdf (distr \ borel X) x = 1 \ x \ $\" for x::real using ccdfX_0_equiv distrX_RD.cdf_ccdf distrX_RD.prob_space by force subsubsection \Properties of Cumulative Distributive Function for T(x)\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct, simp_all add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma cdfTx_cond_prob: "cdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. T x \ \ t \ T x \ > 0)" for t::real apply (rewrite distrTx_RD.cdf_ccdf, rewrite distrTx_RD.prob_space) apply (rewrite ccdfTx_cond_prob, simp) by (rewrite not_less[THEN sym], rewrite MM_PS.cond_prob_neg; simp) lemma cdfTx_0_0: "cdf (distr (\ \ alive x) borel (T x)) 0 = 0" using ccdfTx_0_1 distrTx_RD.cdf_ccdf distrTx_RD.prob_space by force lemma cdfTx_nonpos_0: "cdf (distr (\ \ alive x) borel (T x)) t = 0" if "t \ 0" for t :: real using ccdfTx_nonpos_1 distrTx_RD.cdf_ccdf distrTx_RD.prob_space that by force lemma cdfTx_1_equiv: "cdf (distr (\ \ alive x) borel (T x)) t = 1 \ x+t \ $\" for t::real using ccdfTx_0_equiv distrTx_RD.cdf_ccdf distrTx_RD.prob_space by force lemma cdfTx_continuous_on_nonpos[simp]: "continuous_on {..0} (cdf (distr (\ \ alive x) borel (T x)))" by (rewrite continuous_on_cong[where g="\t. 0"]) (simp_all add: cdfTx_nonpos_0)+ lemma cdfTx_differentiable_on_nonpos[simp]: "(cdf (distr (\ \ alive x) borel (T x))) differentiable_on {..0}" by (rewrite differentiable_on_cong[where f="\t. 0"]; simp add: cdfTx_nonpos_0) lemma cdfTx_has_real_derivative_0_at_neg: "(cdf (distr (\ \ alive x) borel (T x)) has_real_derivative 0) (at t)" if "t < 0" for t::real apply (rewrite has_real_derivative_iff_has_vector_derivative) apply (rule has_vector_derivative_transform_within_open[of "\_. 0" _ _ "{..<0}"]) using cdfTx_nonpos_0 that by simp_all lemma cdfTx_integrable_Icc: "set_integrable lborel {a..b} (cdf (distr (\ \ alive x) borel (T x)))" for a b :: real proof - have "set_integrable lborel {a..b} (\_. 1::real)" unfolding set_integrable_def using emeasure_compact_finite by (simp, intro integrable_real_indicator; force) thus ?thesis apply (rewrite distrTx_RD.cdf_ccdf, rewrite distrTx_RD.prob_space) using ccdfTx_integrable_Icc by (rewrite set_integral_diff; simp) qed corollary cdfTx_integrable_on_Icc: "cdf (distr (\ \ alive x) borel (T x)) integrable_on {a..b}" for a b :: real using cdfTx_integrable_Icc set_borel_integral_eq_integral by force lemma cdfTx_PX: "cdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. x < X \ \ X \ \ x+t) / \

(\ in \. X \ > x)" for t::real apply (rewrite cdfTx_cond_prob) unfolding cond_prob_def futr_life_def PXx_pos by (smt (verit) Collect_cong) lemma cdfT0_eq_cdfX: "cdf (distr (\ \ alive 0) borel (T 0)) = cdf (distr \ borel X)" proof interpret alive0_PS: prob_space "\ \ alive 0" apply (rule MM_PS.cond_prob_space_correct, simp) using PXx_pos alive_def psi_pos' by presburger interpret distrT0_RD: real_distribution "distr (\ \ alive 0) borel (T 0)" by simp show "\x. cdf (distr (\ \ alive 0) borel (T 0)) x = cdf (distr \ borel X) x" using ccdfT0_eq_ccdfX distrX_RD.ccdf_cdf distrT0_RD.ccdf_cdf by (smt (verit, best) distrT0_RD.prob_space distrX_RD.prob_space psi_pos') qed lemma continuous_cdfX_cdfTx: "continuous (at (x+t) within {x..}) (cdf (distr \ borel X)) \ continuous (at t within {0..}) (cdf (distr (\ \ alive x) borel (T x)))" if "t \ 0" for t::real proof - have "continuous (at (x+t) within {x..}) (cdf (distr \ borel X)) \ continuous (at (x+t) within {x..}) (ccdf (distr \ borel X))" by (rule distrX_RD.continuous_cdf_ccdf) also have "\ \ continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" using continuous_ccdfX_ccdfTx that by simp also have "\ \ continuous (at t within {0..}) (cdf (distr (\ \ alive x) borel (T x)))" using distrTx_RD.continuous_cdf_ccdf by simp finally show ?thesis . qed lemma isCont_cdfX_cdfTx: "isCont (cdf (distr \ borel X)) (x+t) \ isCont (cdf (distr (\ \ alive x) borel (T x))) t" if "t > 0" for t::real apply (rewrite distrX_RD.isCont_cdf_ccdf) apply (rewrite isCont_ccdfX_ccdfTx, simp_all add: that) by (rule distrTx_RD.isCont_cdf_ccdf[THEN sym]) lemma has_real_derivative_cdfX_cdfTx: "((cdf (distr \ borel X)) has_real_derivative D) (at (x+t)) \ ((cdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" if "t > 0" for t D :: real proof - have "((cdf (distr \ borel X)) has_real_derivative D) (at (x+t)) \ (ccdf (distr \ borel X) has_real_derivative -D) (at (x+t))" using distrX_RD.has_real_derivative_cdf_ccdf by force also have "\ \ ((ccdf (distr (\ \ alive x) borel (T x))) has_real_derivative (-D / \

(\ in \. X \ > x))) (at t)" using has_real_derivative_ccdfX_ccdfTx that by simp also have "\ \ ((cdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" by (simp add: distrTx_RD.has_real_derivative_cdf_ccdf) finally show ?thesis . qed lemma differentiable_cdfX_cdfTx: "(cdf (distr \ borel X)) differentiable at (x+t) \ (cdf (distr (\ \ alive x) borel (T x))) differentiable at t" if "t > 0" for t::real apply (rewrite differentiable_eq_field_differentiable_real)+ unfolding field_differentiable_def using has_real_derivative_cdfX_cdfTx that by (meson differentiable_ccdfX_ccdfTx distrTx_RD.finite_borel_measure_axioms distrX_RD.finite_borel_measure_axioms finite_borel_measure.differentiable_cdf_ccdf real_differentiable_def x_lt_psi) subsubsection \Properties of @{text "$q_{t&x}"}\ lemma q_0_0: "$q_{0&x} = 0" unfolding die_def using cdfTx_0_0 by simp lemma q_nonneg[simp]: "$q_{t&x} \ 0" for t::real unfolding die_def using distrTx_RD.cdf_nonneg by simp lemma q_1_equiv: "$q_{t&x} = 1 \ x+t \ $\" for t::real unfolding die_def using cdfTx_1_equiv by simp lemma q_PTx: "$q_{t&x} = \

(\ in \. T x \ \ t \ T x \ > 0)" for t::real unfolding die_def using cdfTx_cond_prob by simp lemma q_PX: "$q_{t&x} = \

(\ in \. x < X \ \ X \ \ x + t) / \

(\ in \. X \ > x)" unfolding die_def using cdfTx_PX by simp lemma q_defer_0_q[simp]: "$q_{0\t&x} = $q_{t&x}" for t::real unfolding die_defer_def using q_0_0 by simp lemma q_defer_0_0: "$q_{f\0&x} = 0" for f::real unfolding die_defer_def by simp lemma q_defer_nonneg[simp]: "$q_{f\t&x} \ 0" for f t :: real unfolding die_defer_def by simp lemma q_defer_q: "$q_{f\t&x} = $q_{f+t & x} - $q_{f&x}" if "t \ 0" for f t :: real unfolding die_defer_def die_def using distrTx_RD.cdf_nondecreasing that by simp lemma q_defer_PTx: "$q_{f\t&x} = \

(\ in \. f < T x \ \ T x \ \ f + t \ T x \ > 0)" if "t \ 0" for f t :: real proof - have "$q_{f\t&x} = $q_{f+t & x} - $q_{f&x}" using q_defer_q that by simp also have "\ = \

(\ in \. T x \ \ f + t \ T x \ > 0) - \

(\ in \. T x \ \ f \ T x \ > 0)" using q_PTx by simp also have "\ = \

(\ in (\ \ alive x). T x \ \ f + t) - \

(\ in (\ \ alive x). T x \ \ f)" using MM_PS.cond_prob_space_cond_prob alive_T by simp also have "\ = \

(\ in (\ \ alive x). f < T x \ \ T x \ \ f + t)" proof - have "{\ \ space (\ \ alive x). T x \ \ f + t} - {\ \ space (\ \ alive x). T x \ \ f} = {\ \ space (\ \ alive x). f < T x \ \ T x \ \ f + t}" using that by force hence "alivex_PS.prob ({\ \ space (\ \ alive x). T x \ \ f + t} - {\ \ space (\ \ alive x). T x \ \ f}) = \

(\ in (\ \ alive x). f < T x \ \ T x \ \ f + t)" by simp moreover have "{\ \ space (\ \ alive x). T x \ \ f} \ {\ \ space (\ \ alive x). T x \ \ f + t}" using that by force ultimately show ?thesis by (rewrite alivex_PS.finite_measure_Diff[THEN sym]; simp) qed also have "\ = \

(\ in \. f < T x \ \ T x \ \ f + t \ T x \ > 0)" using MM_PS.cond_prob_space_cond_prob alive_T by simp finally show ?thesis . qed lemma q_defer_PX: "$q_{f\t&x} = \

(\ in \. x + f < X \ \ X \ \ x + f + t) / \

(\ in \. X \ > x)" if "f \ 0" "t \ 0" for f t :: real proof - have "$q_{f\t&x} = \

(\ in \. f < T x \ \ T x \ \ f + t \ T x \ > 0) / \

(\ in \. T x \ > 0)" apply (rewrite q_defer_PTx; (simp add: that)?) unfolding cond_prob_def by simp also have "\ = \

(\ in \. f < T x \ \ T x \ \ f + t) / \

(\ in \. T x \ > 0)" proof - have "\\. \ \ space \ \ f < T x \ \ T x \ \ f + t \ T x \ > 0 \ f < T x \ \ T x \ \ f + t" using that by auto hence "{\ \ space \. f < T x \ \ T x \ \ f + t \ T x \ > 0} = {\ \ space \. f < T x \ \ T x \ \ f + t}" by blast thus ?thesis by simp qed also have "\ = \

(\ in \. x + f < X \ \ X \ \ x + f + t) / \

(\ in \. X \ > x)" unfolding futr_life_def by (smt (verit) Collect_cong) finally show ?thesis . qed lemma q_defer_old_0: "$q_{f\t&x} = 0" if "x+f \ $\" "t \ 0" for f t :: real proof - have "$q_{f\t&x} = $q_{f+t & x} - $q_{f&x}" using q_defer_q that by simp moreover have "$q_{f+t & x} = 1" using q_1_equiv that le_ereal_le by auto moreover have "$q_{f&x} = 1" using q_1_equiv that by simp ultimately show ?thesis by simp qed end subsubsection \Properties of Cumulative Distributive Function for X\ lemma cdfX_continuous_unborn[simp]: "continuous_on {..0} (cdf (distr \ borel X))" using cdfTx_continuous_on_nonpos by (metis cdfT0_eq_cdfX psi_pos') lemma cdfX_differentiable_unborn[simp]: "(cdf (distr \ borel X)) differentiable_on {..0}" using cdfTx_differentiable_on_nonpos by (metis cdfT0_eq_cdfX psi_pos') lemma cdfX_has_real_derivative_0_unborn: "(cdf (distr \ borel X) has_real_derivative 0) (at x)" if "x < 0" for x::real using cdfTx_has_real_derivative_0_at_neg by (metis cdfT0_eq_cdfX psi_pos' that) lemma cdfX_integrable_Icc: "set_integrable lborel {a..b} (cdf (distr \ borel X))" for a b :: real using cdfTx_integrable_Icc by (metis cdfT0_eq_cdfX psi_pos') corollary cdfX_integrable_on_Icc: "cdf (distr \ borel X) integrable_on {a..b}" for a b :: real using cdfX_integrable_Icc set_borel_integral_eq_integral by force lemma cdfX_q: "cdf (distr \ borel X) x = $q_{x&0}" if "x \ 0" for x::real by (metis cdfT0_eq_cdfX die_def psi_pos') subsubsection \Relations between @{text "$p_{t&x}"} and @{text "$q_{t&x}"}\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct, simp_all add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma p_q_1: "$p_{t&x} + $q_{t&x} = 1" for t::real unfolding survive_def die_def using distrTx_RD.add_cdf_ccdf by (smt (verit) distrTx_RD.prob_space x_lt_psi) lemma q_defer_p: "$q_{f\t&x} = $p_{f&x} - $p_{f+t & x}" if "t \ 0" for f t :: real using q_defer_q p_q_1 that x_lt_psi by smt lemma q_defer_p_q_defer: "$p_{f&x} * $q_{f'\t & x+f} = $q_{f+f'\t & x}" if "x+f < $\" "f \ 0" "f' \ 0" "t \ 0" for f f' t :: real proof - have "$p_{f&x} * $q_{f'\t & x+f} = \

(\ in \. X \ > x+f) / \

(\ in \. X \ > x) * \

(\ in \. x+f+f' < X \ \ X \ \ x+f+f'+t) / \

(\ in \. X \ > x+f)" apply (rewrite p_PX, (simp_all add: that)[2]) by (rewrite survival_model.q_defer_PX, simp_all add: that survival_model_axioms) also have "\ = \

(\ in \. x+f+f' < X \ \ X \ \ x+f+f'+t) / \

(\ in \. X \ > x)" using survival_model.PXx_pos[of \ X "x+f"] nonzero_mult_div_cancel_left that by (smt (verit, ccfv_SIG) survival_model_axioms times_divide_eq_left times_divide_eq_right) also have "\ = $q_{f+f'\t & x}" by (rewrite q_defer_PX; simp add: that group_cancel.add1) finally show ?thesis . qed lemma q_defer_pq: "$q_{f\t&x} = $p_{f&x} * $q_{t & x+f}" if "x+f < $\" "t \ 0" "f \ 0" for f t :: real using q_defer_p_q_defer[where f'=0] that by (simp add: survival_model.q_defer_0_q survival_model_axioms) subsubsection \Properties of Life Expectation\ lemma e_nonneg: "$e`\_x \ 0" unfolding life_expect_def by (rule Bochner_Integration.integral_nonneg, simp add: less_eq_real_def) lemma e_P: "$e`\_x = MM_PS.expectation (\\. indicator (alive x) \ * T x \) / \

(\ in \. T x \ > 0)" unfolding life_expect_def by (rewrite MM_PS.integral_cond_prob_space_nn, auto simp add: alive_T) proposition nn_integral_T_p: "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = \\<^sup>+t\{0..}. ennreal ($p_{t&x}) \lborel" apply (rewrite alivex_PS.expectation_nonneg_tail, simp_all add: less_imp_le) apply (rule nn_integral_cong) unfolding survive_def using distrTx_RD.prob_space distrTx_RD.ccdf_cdf by presburger lemma nn_integral_T_pos: "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) > 0" proof - let ?f = "\t. - ccdf (distr (\ \ alive x) borel (T x)) t" have "\t u. t \ u \ ?f t \ ?f u" using distrTx_RD.ccdf_nonincreasing by simp moreover have "continuous (at_right 0) ?f" using distrTx_RD.ccdf_is_right_cont by (intro continuous_intros) ultimately have "\e>0. \d>0. ?f (0 + d) - ?f 0 < e" using continuous_at_right_real_increasing by simp hence "\d>0. ?f (0 + d) - ?f 0 < 1/2" by (smt (verit, del_insts) field_sum_of_halves) from this obtain d where d_pos: "d > 0" and "$p_{d&x} \ 1/2" using p_0_1 unfolding survive_def by auto hence "\t. t\{0..d} \ $p_{t&x} \ 1/2" unfolding survive_def using distrTx_RD.ccdf_nonincreasing by force hence "\\<^sup>+t\{0..d}. ennreal ($p_{t&x}) \lborel \ \\<^sup>+t\{0..d}. ennreal (1/2) \lborel" apply (intro nn_set_integral_mono, simp_all) unfolding survive_def using Tx_alivex_measurable apply force by (rule AE_I2) (smt (verit) ennreal_half ennreal_leI half_bounded_equal) moreover have "\\<^sup>+t\{0..}. ennreal ($p_{t&x}) \lborel \ \\<^sup>+t\{0..d}. ennreal ($p_{t&x}) \lborel" by (rule nn_set_integral_set_mono) simp moreover have "\\<^sup>+t\{0..d}. ennreal (1/2) \lborel > 0" apply (rewrite nn_integral_cmult_indicator, simp_all) using d_pos emeasure_lborel_Icc ennreal_zero_less_mult_iff by fastforce ultimately show ?thesis using nn_integral_T_p by simp qed proposition e_LBINT_p: "$e`\_x = LBINT t:{0..}. $p_{t&x}" \ \Note that 0 = 0 holds when the integral diverges.\ unfolding life_expect_def apply (rewrite integral_eq_nn_integral, simp_all add: less_imp_le) unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral, simp_all) apply (measurable, simp add: survive_def) by (rewrite nn_integral_T_p) (simp add: indicator_mult_ennreal mult.commute) corollary e_integral_p: "$e`\_x = integral {0..} (\t. $p_{t&x})" \ \Note that 0 = 0 holds when the integral diverges.\ proof - have "$e`\_x = LBINT t:{0..}. $p_{t&x}" using e_LBINT_p by simp also have "\ = integral {0..} (\t. $p_{t&x})" apply (rule set_borel_integral_eq_integral_nonneg, simp_all) unfolding survive_def by simp finally show ?thesis . qed lemma e_LBINT_p_Icc: "$e`\_x = LBINT t:{0..n}. $p_{t&x}" if "x+n \ $\" for n::real proof - have [simp]: "{0..n} \ {n<..} = {}" using ivl_disj_int_one(7) by blast have [simp]: "{0..n} \ {n<..} = {0..}" by (smt (verit) ereal_less_le ivl_disj_un_one(7) leD that x_lt_psi) have [simp]: "\t. n < t \ 0 \ t" using that x_lt_psi by (smt (verit) ereal_less_le leD) have [simp]: "\t. n < t \ $\ \ ereal (x+t)" using that by (simp add: le_ereal_le) have gt_n_0: "has_bochner_integral lborel (\t. indicat_real {n<..} t * $p_{t&x}) 0" apply (rewrite has_bochner_integral_cong[where N=lborel and g="\t.0" and y=0], simp_all) using p_0_equiv that x_lt_psi apply (smt (verit, ccfv_SIG) greaterThan_iff indicator_simps le_ereal_le linorder_not_le) by (rule has_bochner_integral_zero) hence gt_n: "set_integrable lborel {n<..} (\t. $p_{t&x})" unfolding set_integrable_def using integrable.simps by auto moreover have le_n: "set_integrable lborel {0..n} (\t. $p_{t&x})" unfolding survive_def by (intro ccdfTx_integrable_Icc) simp ultimately have "set_integrable lborel ({0..n} \ {n<..}) (\t. $p_{t&x})" using set_integrable_Un by force hence "set_integrable lborel {0..} (\t. $p_{t&x})" by force thus ?thesis apply (rewrite e_LBINT_p, simp) apply (rewrite set_integral_Un[of "{0..n}" "{n<..}", simplified], simp_all add: gt_n le_n) unfolding set_lebesgue_integral_def using gt_n_0 has_bochner_integral_integral_eq by fastforce qed lemma e_integral_p_Icc: "$e`\_x = integral {0..n} (\t. $p_{t&x})" if "x+n \ $\" for n::real using that apply (rewrite e_LBINT_p_Icc, simp_all) using ccdfTx_integrable_Icc unfolding survive_def by (rewrite set_borel_integral_eq_integral; simp) lemma temp_e_P: "$e`\_{x:n} = MM_PS.expectation (\\. indicator (alive x) \ * min (T x \) n) / \

(\ in \. T x \ > 0)" if "n \ 0" for n::real unfolding temp_life_expect_def by (rewrite MM_PS.integral_cond_prob_space_nn; simp add: alive_T that) lemma temp_e_LBINT_p: "$e`\_{x:n} = LBINT t:{0..n}. $p_{t&x}" if "n \ 0" for n::real proof - let ?minTxn = "\\. min (T x \) n" let ?F = "cdf (distr (\ \ alive x) borel (T x))" let ?Fn = "cdf (distr (\ \ alive x) borel ?minTxn)" interpret distrTxn_RD: real_distribution "distr (\ \ alive x) borel ?minTxn" by (simp add: that) have [simp]: "\\. \ \ space (\ \ alive x) \ 0 \ T x \" by (smt (verit) alivex_Tx_pos) have "(\\<^sup>+\. ennreal (min (T x \) n) \(\ \ alive x)) = \\<^sup>+t\{0..}. ennreal (1 - ?Fn t) \lborel" by (rewrite alivex_PS.expectation_nonneg_tail; simp add: that) also have "\ = \\<^sup>+t\{0..}. (ennreal (1 - ?F t) * indicator {..lborel" apply (rule nn_integral_cong) by (rewrite alivex_PS.cdf_distr_min; simp add: indicator_mult_ennreal mult.commute) also have "\ = \\<^sup>+t\{0..lborel" apply (rule nn_integral_cong) using nn_integral_set_ennreal by (smt (verit, best) Int_def atLeastLessThan_def ennreal_mult_right_cong indicator_simps mem_Collect_eq mult.commute mult_1) also have "\ = \\<^sup>+t\{0..n}. ennreal (1 - ?F t) \lborel" proof - have "sym_diff {0.. = ennreal (LBINT t:{0..n}. $p_{t&x})" proof - have "set_integrable lborel {0..n} (\t. $p_{t&x})" unfolding survive_def by (intro ccdfTx_integrable_Icc) simp thus ?thesis unfolding set_lebesgue_integral_def unfolding set_integrable_def apply (rewrite nn_integral_eq_integral[THEN sym]; simp) apply (rule nn_integral_cong, simp) unfolding survive_def using distrTx_RD.ccdf_cdf distrTx_RD.prob_space nn_integral_set_ennreal by (simp add: indicator_mult_ennreal mult.commute) qed finally have "(\\<^sup>+\. ennreal (min (T x \) n) \(\ \ alive x)) = ennreal (LBINT t:{0..n}. $p_{t&x})" . thus ?thesis unfolding temp_life_expect_def apply (rewrite integral_eq_nn_integral; simp add: that) apply (rewrite enn2real_ennreal; simp?) unfolding set_lebesgue_integral_def by (simp add: that) qed lemma temp_e_integral_p: "$e`\_{x:n} = integral {0..n} (\t. $p_{t&x})" if "n \ 0" for n::real using that apply (rewrite temp_e_LBINT_p, simp_all) using ccdfTx_integrable_Icc unfolding survive_def by (rewrite set_borel_integral_eq_integral; simp) lemma e_eq_temp: "$e`\_x = $e`\_{x:n}" if "n \ 0" "x+n \ $\" for n::real using that e_LBINT_p_Icc temp_e_LBINT_p by simp lemma curt_e_P: "$e_x = MM_PS.expectation (\\. indicator (alive x) \ * \T x \\) / \

(\ in \. T x \ > 0)" unfolding curt_life_expect_def apply (rewrite MM_PS.integral_cond_prob_space_nn; simp add: alive_T) by (metis (no_types, lifting) Bochner_Integration.integral_cong indicator_simps of_int_0 of_int_1) lemma curt_e_sum_P: "$e_x = (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" if "summable (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" proof - let ?F_flrTx = "cdf (distr (\ \ alive x) borel (\\. \T x \\))" have [simp]: "\\. \ \ space (\ \ alive x) \ 0 \ T x \" by (smt (verit) alivex_Tx_pos) have "integral\<^sup>N (\ \ alive x) (\\. ennreal \T x \\) = \\<^sup>+t\{0..}. ennreal (1 - ?F_flrTx t) \lborel" by (rewrite alivex_PS.expectation_nonneg_tail; simp) also have "\ = \\<^sup>+t\{0::real..}. ennreal \

(\ in \. T x \ \ \t\ + 1 \ T x \ > 0) \lborel" proof - { fix t::real assume "t \ 0" hence "1 - ?F_flrTx t = \

(\ in \. T x \ \ real_of_int \t\ + 1 \ T x \ > 0)" proof - have "1 - ?F_flrTx t = 1 - \

(\ in (\ \ alive x). T x \ < real_of_int \t\ + 1)" by (rewrite alivex_PS.cdf_distr_floor_P; simp) also have "\ = 1 - \

(\ in \. T x \ < real_of_int \t\ + 1 \ T x \ > 0)" using alive_T by (rewrite MM_PS.cond_prob_space_cond_prob; simp) also have "\ = \

(\ in \. T x \ \ real_of_int \t\ + 1 \ T x \ > 0)" by (rewrite not_le[THEN sym], rewrite MM_PS.cond_prob_neg; simp) finally show ?thesis . qed } thus ?thesis apply - by (rule nn_set_integral_cong2, rule AE_I2) simp qed also have "\ = (\k. \\<^sup>+t\{k..(\ in \. T x \ \ \t\ + 1 \ T x \ > 0) \lborel)" apply (rewrite nn_integral_disjoint_family[THEN sym]; simp) apply (rewrite add.commute, rule Ico_nat_disjoint) by (rewrite Ico_nat_union[THEN sym], simp add: add.commute) also have "\ = (\k. \\<^sup>+t\{k..(\ in \. T x \ \ k + 1 \ T x \ > 0) \lborel)" proof - { fix k::nat and t::real assume "real k \ t" and "t < 1 + real k" hence "real_of_int \t\ = real k" by (metis add.commute floor_eq2 of_int_of_nat_eq) hence "\

(\ in \. T x \ \ real_of_int \t\ + 1 \ T x \ > 0) = \

(\ in \. T x \ \ 1 + real k \ T x \ > 0)" by (simp add: add.commute) } thus ?thesis apply - apply (rule suminf_cong, rule nn_set_integral_cong2, rule AE_I2) by (rule impI) simp qed also have "\ = (\k. ennreal \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" by (rewrite nn_integral_cmult_indicator; simp add: add.commute) also have "\ = ennreal (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" by (rewrite suminf_ennreal2; simp add: that) finally have "integral\<^sup>N (\ \ alive x) (\\. ennreal \T x \\) = ennreal (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" . hence "integral\<^sup>L (\ \ alive x) (\\. \T x \\) = (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" apply (rewrite integral_eq_nn_integral; simp) apply (rewrite enn2real_ennreal; simp add: add.commute) apply (rule suminf_nonneg; simp?) by (rewrite add.commute, simp add: that) thus ?thesis unfolding curt_life_expect_def by (simp add: add.commute) qed lemma curt_e_sum_P_finite: "$e_x = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" if "x+n+1 > $\" for n::nat proof - from that have psi_fin: "$\ < \" by force let ?P = "\k::nat. \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" let ?P_fin = "\k::nat. if k\{..k. ?P k = ?P_fin k" proof - fix k show "?P k = ?P_fin k" proof (cases \k \ {..) case True thus ?thesis by simp next case False hence "\ k < n" by simp hence "x + k + 1 > real_of_ereal $\" using that psi_nonneg real_of_ereal_ord_simps(4) by fastforce hence "{\ \ space \. T x \ \ k + 1 \ T x \ > 0} \ {\ \ space \. X \ > real_of_ereal $\}" unfolding futr_life_def using that less_ereal_le of_nat_1 of_nat_add by force hence "\

(\ in \. T x \ \ k + 1 \ T x \ > 0) \ \

(\ in \. X \ > real_of_ereal $\)" by (intro MM_PS.finite_measure_mono, simp_all) also have "\ = 0" using MM_PS.ccdf_distr_P X_RV ccdfX_psi_0 psi_fin by presburger finally have "\

(\ in \. T x \ \ k + 1 \ T x \ > 0) = 0" using measure_le_0_iff by blast hence "?P k = 0" unfolding cond_prob_def by (simp add: add.commute) thus ?thesis by simp qed qed moreover have "?P_fin sums (\k: "?P sums (\kk) thus ?thesis by (simp add: add.commute) qed lemma curt_e_sum_p: "$e_x = (\k. $p_{k+1&x})" if "summable (\k. $p_{k+1&x})" "\k::nat. isCont (\t. $p_{t&x}) (k+1)" proof - have "\k::nat. $p_{k+1&x} = \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" apply (rewrite p_PTx_ge_ccdf_isCont, simp_all) using that(2) isCont_ccdfX_ccdfTx unfolding survive_def by simp thus ?thesis using that p_PTx_ge_ccdf_isCont curt_e_sum_P by presburger qed lemma curt_e_rec: "$e_x = $p_x * (1 + $e_(x+1))" if "summable (\k. $p_{k+1&x})" "\k::nat. isCont (\t. $p_{t&x}) (real k + 1)" "x+1 < $\" proof - have px_neq_0[simp]: "$p_x \ 0" using p_0_equiv that by auto have "(\k. $p_{k+1&x}) sums $e_x" using that apply (rewrite curt_e_sum_p, simp_all add: add.commute) by (rule summable_sums, simp add: that) hence "(\k. $p_x * $p_{k&x+1}) sums $e_x" apply (rewrite sums_cong[where g="\k. $p_{k+1&x}"]; simp?) using p_mult by (smt (verit) of_nat_0_le_iff that(3) x_lt_psi) hence "(\k. $p_{k&x+1}) sums ($e_x / $p_x)" using sums_mult_D that by (smt (verit, best) linorder_not_le p_0_equiv sums_cong x_lt_psi) hence p_e_p: "(\k. $p_{Suc k & x+1}) sums ($e_x / $p_x - $p_{0&x+1})" using sums_split_initial_segment[where n=1] by force moreover have "(\k. $p_{Suc k & x+1}) sums $e_(x+1)" proof - have [simp]: "summable (\k::nat. $p_{real k + 1 & x + 1})" apply (intro sums_summable[where l="$e_x / $p_x - $p_{0&x+1}"]) using p_e_p by (simp add: add.commute) have [simp]: "\k::nat. isCont (\t. $p_{t&x+1}) (real k + 1)" proof - fix k::nat have "isCont (\t. $p_x * $p_{t-1&x+1}) (real k + 2)" proof - let ?S="{real k + 1 <..< real k + 3}" have "open ?S" by simp moreover have "real k + 2 \ ?S" by simp moreover have "\t. t \ ?S \ $p_x * $p_{t-1&x+1} = $p_{t&x}" using p_mult by (smt (verit, del_insts) greaterThanLessThan_iff of_nat_0_le_iff that(3) x_lt_psi) ultimately show ?thesis apply (rewrite isCont_cong[where g="\t. $p_{t&x}"]) apply (rewrite eventually_nhds, blast) using that by (smt (verit) of_nat_1 of_nat_add) qed hence "isCont (\t. $p_x * $p_{t-1&x+1} / $p_x) (real k + 2)" by (intro isCont_divide[where g="\t. $p_x"], auto) hence "isCont ((\t. $p_{t-1&x+1}) \ (\t. t+1)) (real k + 1)" by simp (rule continuous_at_compose, simp_all add: add.commute) thus "isCont (\t. $p_{t&x+1}) (real k + 1)" unfolding comp_def by simp qed show ?thesis apply (rewrite survival_model.curt_e_sum_p; simp add: survival_model_axioms that) using summable_sums by (rewrite add.commute) force qed ultimately have "$e_x / $p_x - $p_{0&x+1} = $e_(x+1)" by (rule sums_unique2) thus ?thesis using p_0_1 that by (smt (verit) px_neq_0 divide_mult_cancel mult.commute mult_cancel_left2 p_mult that(3)) qed lemma curt_e_sum_p_finite: "$e_x = (\kk::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" "x+n+1 > $\" for n::nat proof - have "\k::nat. k < n \ $p_{k+1&x} = \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" apply (rewrite p_PTx_ge_ccdf_isCont, simp_all) using that isCont_ccdfX_ccdfTx unfolding survive_def by (smt (verit) of_nat_0_le_iff x_lt_psi) thus ?thesis using that p_PTx_ge_ccdf_isCont curt_e_sum_P_finite by simp qed lemma temp_curt_e_P: "$e_{x:n} = MM_PS.expectation (\\. indicator (alive x) \ * \min (T x \) n\) / \

(\ in \. T x \ > 0)" if "n \ 0" for n::real unfolding temp_curt_life_expect_def apply (rewrite MM_PS.integral_cond_prob_space_nn; simp add: alive_T that) apply (rule disjI2, rule Bochner_Integration.integral_cong; simp) using indicator_simps of_int_0 of_int_1 by smt lemma temp_curt_e_sum_P: "$e_{x:n} = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" for n::nat proof - let ?F_flrminTx = "cdf (distr (\ \ alive x) borel (\\. \min (T x \) n\))" have [simp]: "\\. \ \ space (\ \ alive x) \ 0 \ T x \" by (smt (verit) alivex_Tx_pos) have "integral\<^sup>N (\ \ alive x) (\\. ennreal \min (T x \) n\) = (\\<^sup>+t\{0..}. ennreal (1 - ?F_flrminTx t) \lborel)" by (rewrite alivex_PS.expectation_nonneg_tail; simp) also have "\ = \\<^sup>+t\{0::real..}. ennreal ((1 - \

(\ in (\ \ alive x). T x \ < \t\ + 1)) * of_bool (\t\ + 1 \ n)) \lborel" proof - have "\t. t \ 0 \ ennreal (1 - ?F_flrminTx t) = ennreal ((1 - \

(\ in (\ \ alive x). T x \ < \t\ + 1)) * of_bool (\t\ + 1 \ n))" proof - fix t::real assume "t \ 0" thus "ennreal (1 - ?F_flrminTx t) = ennreal ((1 - \

(\ in (\ \ alive x). T x \ < \t\ + 1)) * of_bool (\t\ + 1 \ n))" proof (cases \\t\ + 1 \ n\) case True thus ?thesis apply (rewrite alivex_PS.cdf_distr_floor_P; simp) using min_less_iff_disj by (smt (verit, ccfv_SIG) Collect_cong floor_eq floor_less_cancel floor_of_nat of_int_floor_le) next case False thus ?thesis apply (rewrite alivex_PS.cdf_distr_floor_P; simp) using min_less_iff_disj by (smt (verit, ccfv_SIG) Collect_cong MM_PS.space_cond_prob_space alive_T alive_event alivex_PS.prob_space mem_Collect_eq of_int_of_nat_eq of_nat_less_of_int_iff) qed qed thus ?thesis by (intro nn_set_integral_cong2, intro AE_I2) auto qed also have "\ = \\<^sup>+t\{0..(\ in (\ \ alive x). T x \ < \t\ + 1)) \lborel" proof - { fix t::real have "(\t\ + 1 \ n) = (t < n)" by linarith hence "\r::real. ennreal (r * of_bool (\t\ + 1 \ n)) * indicator {0..} t = ennreal r * indicator {0.. = \\<^sup>+t\{0..(\ in (\ \ alive x). T x \ \ \t\ + 1) \lborel" by (rewrite alivex_PS.prob_neg[THEN sym]; simp add: not_less) also have "\ = (\k\<^sup>+t\{k..(\ in (\ \ alive x). T x \ \ \t\ + 1) \lborel)" apply (rewrite Ico_nat_union_finite[of n, THEN sym]) apply (rewrite nn_integral_disjoint_family_on_finite; simp add: add.commute) apply (rule disjoint_family_on_mono[of _ "{0..}"]; simp) by (rewrite add.commute, rule Ico_nat_disjoint) also have "\ = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" proof - { fix k::nat assume "k < n" hence "\\<^sup>+t\{k..<(1 + real k)}. ennreal \

(\ in (\ \ alive x). T x \ \ real_of_int \t\ + 1) \lborel = \

(\ in \. T x \ \ 1 + real k \ T x \ > 0)" (is "?LHS = ?RHS") proof - have "?LHS = \\<^sup>+t\{k..<(1 + real k)}. ennreal \

(\ in (\ \ alive x). T x \ \ k + 1) \lborel" proof - { fix t::real assume "k \ t" "t < 1 + real k" hence "real_of_int \t\ = real k" by (smt (verit) floor_eq2 of_int_of_nat_eq) hence "\

(\ in (\ \ alive x). T x \ \ real_of_int \t\ + 1) = \

(\ in (\ \ alive x). T x \ \ 1 + real k)" by (simp add: add.commute) } thus ?thesis by (intro nn_set_integral_cong2, intro AE_I2) auto qed also have "\ = ennreal \

(\ in (\ \ alive x). T x \ \ k + 1)" by (rewrite nn_integral_cmult_indicator; simp) also have "\ = ?RHS" by (rewrite MM_PS.cond_prob_space_cond_prob, simp_all add: alive_T) finally show ?thesis . qed } thus ?thesis by (intro sum.cong) auto qed also have "\ = ennreal (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" by simp finally have "integral\<^sup>N (\ \ alive x) (\\. ennreal \min (T x \) n\) = ennreal (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" . hence "integral\<^sup>L (\ \ alive x) (\\. \min (T x \) n\) = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" apply (intro nn_integral_eq_integrable[THEN iffD1, THEN conjunct2]; simp) using MM_PS.cond_prob_nonneg by (meson sum_nonneg) thus ?thesis unfolding temp_curt_life_expect_def by simp qed corollary curt_e_eq_temp: "$e_x = $e_{x:n}" if "x+n+1 > $\" for n::nat using curt_e_sum_P_finite temp_curt_e_sum_P that by simp lemma temp_curt_e_sum_p: "$e_{x:n} = (\kk::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" for n::nat proof - have "\k::nat. k < n \ $p_{k+1&x} = \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" apply (rewrite p_PTx_ge_ccdf_isCont, simp_all) using that isCont_ccdfX_ccdfTx unfolding survive_def by (smt (verit) of_nat_0_le_iff x_lt_psi) thus ?thesis apply (rewrite temp_curt_e_sum_P) by (rule sum.cong) simp_all qed lemma temp_curt_e_rec: "$e_{x:n} = $p_x * (1 + $e_{x+1:n-1})" if "\k::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" "x+1 < $\" "n \ 0" for n::nat proof - have [simp]: "$p_x \ 0" using p_0_equiv that by auto have [simp]: "\k. k < n-1 \ isCont (\t. $p_{t&x+1}) (real k + 1)" proof - fix k::nat assume k_n: "k < n-1" have "isCont (\t. $p_x * $p_{t-1&x+1}) (real k + 2)" proof - let ?S="{real k + 1 <..< real k + 3}" have "open ?S" by simp moreover have "real k + 2 \ ?S" by simp moreover have "\t. t \ ?S \ $p_x * $p_{t-1&x+1} = $p_{t&x}" using p_mult by (smt (verit, del_insts) greaterThanLessThan_iff of_nat_0_le_iff that(2) x_lt_psi) ultimately show ?thesis apply (rewrite isCont_cong[where g="\t. $p_{t&x}"]) apply (rewrite eventually_nhds, blast) using that k_n by (smt (verit) less_diff_conv of_nat_1 of_nat_add) qed hence "isCont (\t. $p_x * $p_{t-1&x+1} / $p_x) (real k + 2)" by (intro isCont_divide[where g="\t. $p_x"], auto) hence "isCont ((\t. $p_{t-1&x+1}) \ (\t. t+1)) (real k + 1)" by simp (rule continuous_at_compose, simp_all add: add.commute) thus "isCont (\t. $p_{t&x+1}) (real k + 1)" unfolding comp_def by simp qed have "$p_x * (1 + $e_{x+1:n-1}) = $p_x * (1 + (\k<(n-1). $p_{k+1&x+1}))" by (rewrite survival_model.temp_curt_e_sum_p; simp add: survival_model_axioms that) also have "\ = $p_x + (\k<(n-1). $p_x * $p_{k+1&x+1})" apply (rewrite distrib_left, simp) by (rewrite vector_space_over_itself.scale_sum_right, simp) also have "\ = $p_x + (\k<(n-1). $p_{k+2&x})" apply (rewrite sum.cong, simp_all add: add.commute) using p_mult by (smt (verit) of_nat_0_le_iff that(2) x_lt_psi) also have "\ = (\k = $e_{x:n}" using that by (rewrite temp_curt_e_sum_p; simp) finally show ?thesis by simp qed end end subsection \Piecewise Differentiable Survival Function\ locale smooth_survival_function = survival_model + assumes ccdfX_piecewise_differentiable[simp]: "(ccdf (distr \ borel X)) piecewise_differentiable_on UNIV" begin interpretation distrX_RD: real_distribution "distr \ borel X" using MM_PS.real_distribution_distr by simp subsubsection \Properties of Survival Function for X\ lemma ccdfX_continuous[simp]: "continuous_on UNIV (ccdf (distr \ borel X))" using ccdfX_piecewise_differentiable piecewise_differentiable_on_imp_continuous_on by fastforce corollary ccdfX_borel_measurable[measurable]: "ccdf (distr \ borel X) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma ccdfX_nondifferentiable_finite_set[simp]: "finite {x. \ ccdf (distr \ borel X) differentiable at x}" proof - obtain S where fin: "finite S" and "\x. x \ S \ (ccdf (distr \ borel X)) differentiable at x" using ccdfX_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast hence "{x. \ ccdf (distr \ borel X) differentiable at x} \ S" by blast thus ?thesis using finite_subset fin by blast qed lemma ccdfX_differentiable_open_set: "open {x. ccdf (distr \ borel X) differentiable at x}" using ccdfX_nondifferentiable_finite_set finite_imp_closed by (metis (mono_tags, lifting) Collect_cong open_Collect_neg) lemma ccdfX_differentiable_borel_set[measurable, simp]: "{x. ccdf (distr \ borel X) differentiable at x} \ sets borel" using ccdfX_differentiable_open_set by simp lemma ccdfX_differentiable_AE: "AE x in lborel. (ccdf (distr \ borel X)) differentiable at x" apply (rule AE_I'[of "{x. \ ccdf (distr \ borel X) differentiable at x}"], simp_all) using ccdfX_nondifferentiable_finite_set by (simp add: finite_imp_null_set_lborel) lemma deriv_ccdfX_measurable[measurable]: "deriv (ccdf (distr \ borel X)) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (ccdf (distr \ borel X)))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Properties of Cumulative Distributive Function for X\ lemma cdfX_piecewise_differentiable[simp]: "(cdf (distr \ borel X)) piecewise_differentiable_on UNIV" by (rewrite distrX_RD.cdf_ccdf) (rule piecewise_differentiable_diff; simp) lemma cdfX_continuous[simp]: "continuous_on UNIV (cdf (distr \ borel X))" using cdfX_piecewise_differentiable piecewise_differentiable_on_imp_continuous_on by fastforce corollary cdfX_borel_measurable[measurable]: "cdf (distr \ borel X) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma cdfX_nondifferentiable_finite_set[simp]: "finite {x. \ cdf (distr \ borel X) differentiable at x}" using distrX_RD.differentiable_cdf_ccdf ccdfX_nondifferentiable_finite_set by simp lemma cdfX_differentiable_open_set: "open {x. cdf (distr \ borel X) differentiable at x}" using distrX_RD.differentiable_cdf_ccdf ccdfX_differentiable_open_set by simp lemma cdfX_differentiable_borel_set[measurable, simp]: "{x. cdf (distr \ borel X) differentiable at x} \ sets borel" using cdfX_differentiable_open_set by simp lemma cdfX_differentiable_AE: "AE x in lborel. (cdf (distr \ borel X)) differentiable at x" using ccdfX_differentiable_AE distrX_RD.differentiable_cdf_ccdf AE_iffI by simp lemma deriv_cdfX_measurable[measurable]: "deriv (cdf (distr \ borel X)) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (cdf (distr \ borel X)))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Introduction of Probability Density Functions of X and T(x)\ definition pdfX :: "real \ real" where "pdfX x \ if cdf (distr \ borel X) differentiable at x then deriv (cdf (distr \ borel X)) x else 0" \ \This function is defined to be always nonnegative for future application.\ definition pdfT :: "real \ real \ real" where "pdfT x t \ if cdf (distr (\ \ alive x) borel (T x)) differentiable at t then deriv (cdf (distr (\ \ alive x) borel (T x))) t else 0" \ \This function is defined to be always nonnegative for future application.\ lemma pdfX_measurable[measurable]: "pdfX \ borel_measurable borel" proof - let ?cdfX = "cdf (distr \ borel X)" have "countable {x. \ ?cdfX differentiable at x}" using cdfX_nondifferentiable_finite_set uncountable_infinite by force thus ?thesis unfolding pdfX_def apply - by (rule measurable_discrete_difference [where X="{x. \ ?cdfX differentiable at x}" and f="deriv ?cdfX"]; simp) qed lemma distributed_pdfX: "distributed \ lborel X pdfX" proof - let ?cdfX = "cdf (distr \ borel X)" obtain S where fin: "finite S" and diff: "\t. t \ S \ ?cdfX differentiable at t" using cdfX_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast { fix t::real assume t_notin: "t \ S" have "?cdfX differentiable at t" using diff t_notin by simp hence "(?cdfX has_real_derivative pdfX t) (at t)" unfolding pdfX_def using DERIV_deriv_iff_real_differentiable by auto } thus ?thesis by (intro MM_PS.distributed_deriv_cdf[where S=S]; simp add: fin) qed lemma pdfT0_X: "pdfT 0 = pdfX" unfolding pdfT_def pdfX_def using cdfT0_eq_cdfX psi_pos' by fastforce subsubsection \Properties of Survival Function for T(x)\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct; simp add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma ccdfTx_continuous_on_nonneg[simp]: "continuous_on {0..} (ccdf (distr (\ \ alive x) borel (T x)))" apply (rewrite continuous_on_eq_continuous_within, auto) apply (rewrite continuous_ccdfX_ccdfTx[THEN sym], auto) by (metis UNIV_I ccdfX_continuous continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) lemma ccdfTx_continuous[simp]: "continuous_on UNIV (ccdf (distr (\ \ alive x) borel (T x)))" proof - have [simp]: "{..0::real} \ {0..} = UNIV" by auto have "continuous_on {..0} (ccdf (distr (\ \ alive x) borel (T x)))" by (rule ccdfTx_continuous_on_nonpos) simp moreover have "continuous_on {0..} (ccdf (distr (\ \ alive x) borel (T x)))" by simp ultimately have "continuous_on ({..0} \ {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" by (intro continuous_on_closed_Un) simp_all thus ?thesis by simp qed corollary ccdfTx_borel_measurable[measurable]: "ccdf (distr (\ \ alive x) borel (T x)) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma ccdfTx_nondifferentiable_finite_set[simp]: "finite {t. \ ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}" proof - let ?P = "\t. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t" have "{t. t < 0 \ \ ?P t} = {}" proof (rule equals0I) fix t assume asm: "t \ {t. t < 0 \ \ ?P t}" hence "?P t" using ccdfTx_has_real_derivative_0_at_neg real_differentiable_def x_lt_psi by blast with asm show False by simp qed hence "{t. \ ?P t} \ insert 0 {t. t > 0 \ \ ?P t}" by force moreover have "finite {t. t > 0 \ \ ?P t}" proof - have "{t. \ ccdf (distr \ borel X) differentiable at (x+t)} = plus (-x) ` {s. \ ccdf (distr \ borel X) differentiable at s}" by force hence "finite {t. \ ccdf (distr \ borel X) differentiable at (x+t)}" using ccdfX_nondifferentiable_finite_set by simp thus ?thesis using finite_subset differentiable_ccdfX_ccdfTx by (metis (no_types, lifting) mem_Collect_eq subsetI x_lt_psi) qed ultimately show ?thesis using finite_subset by auto qed lemma ccdfTx_differentiable_open_set: "open {t. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}" using ccdfTx_nondifferentiable_finite_set finite_imp_closed by (metis (mono_tags, lifting) Collect_cong open_Collect_neg) lemma ccdfTx_differentiable_borel_set[measurable, simp]: "{t. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t} \ sets borel" using ccdfTx_differentiable_open_set by simp lemma ccdfTx_differentiable_AE: "AE t in lborel. (ccdf (distr (\ \ alive x) borel (T x))) differentiable at t" apply (rule AE_I'[of "{t. \ ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}"]; simp?) using ccdfTx_nondifferentiable_finite_set by (simp add: finite_imp_null_set_lborel) lemma ccdfTx_piecewise_differentiable[simp]: "(ccdf (distr (\ \ alive x) borel (T x))) piecewise_differentiable_on UNIV" proof - have "\t\UNIV-{t. \ ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t" by auto thus ?thesis unfolding piecewise_differentiable_on_def using ccdfTx_continuous ccdfTx_nondifferentiable_finite_set by blast qed lemma deriv_ccdfTx_measurable[measurable]: "deriv (ccdf (distr (\ \ alive x) borel (T x))) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (ccdf (distr (\ \ alive x) borel (T x))))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Properties of Cumulative Distributive Function for T(x)\ lemma cdfTx_continuous[simp]: "continuous_on UNIV (cdf (distr (\ \ alive x) borel (T x)))" using distrTx_RD.cdf_ccdf ccdfTx_continuous by (simp add: continuous_on_eq_continuous_within) corollary cdfTx_borel_measurable[measurable]: "cdf (distr (\ \ alive x) borel (T x)) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma cdfTx_nondifferentiable_finite_set[simp]: "finite {t. \ cdf (distr (\ \ alive x) borel (T x)) differentiable at t}" using distrTx_RD.differentiable_cdf_ccdf by simp lemma cdfTx_differentiable_open_set: "open {t. cdf (distr (\ \ alive x) borel (T x)) differentiable at t}" using distrTx_RD.differentiable_cdf_ccdf ccdfTx_differentiable_open_set by simp lemma cdfTx_differentiable_borel_set[measurable, simp]: "{t. cdf (distr (\ \ alive x) borel (T x)) differentiable at t} \ sets borel" using cdfTx_differentiable_open_set by simp lemma cdfTx_differentiable_AE: "AE t in lborel. (cdf (distr (\ \ alive x) borel (T x))) differentiable at t" by (rewrite distrTx_RD.differentiable_cdf_ccdf, simp add: ccdfTx_differentiable_AE) lemma cdfTx_piecewise_differentiable[simp]: "(cdf (distr (\ \ alive x) borel (T x))) piecewise_differentiable_on UNIV" using piecewise_differentiable_diff piecewise_differentiable_const ccdfTx_piecewise_differentiable by (rewrite distrTx_RD.cdf_ccdf) blast lemma deriv_cdfTx_measurable[measurable]: "deriv (cdf (distr (\ \ alive x) borel (T x))) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (cdf (distr (\ \ alive x) borel (T x))))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Properties of Probability Density Function of T(x)\ lemma pdfTx_nonneg: "pdfT x t \ 0" for t::real proof - fix t show "pdfT x t \ 0" proof (cases \cdf (distr (\ \ alive x) borel (T x)) differentiable at t\) case True have "mono_on UNIV (cdf (distr (\ \ alive x) borel (T x)))" unfolding mono_on_def using distrTx_RD.cdf_nondecreasing by blast moreover have "(cdf (distr (\ \ alive x) borel (T x)) has_real_derivative pdfT x t) (at t)" unfolding pdfT_def using True DERIV_deriv_iff_real_differentiable by presburger ultimately show ?thesis by (rule mono_on_imp_deriv_nonneg) simp next case False thus ?thesis unfolding pdfT_def by simp qed qed lemma pdfTx_neg_0: "pdfT x t = 0" if "t < 0" for t::real proof - let ?e = "-t/2" have "(cdf (distr (\ \ alive x) borel (T x)) has_real_derivative 0) (at t)" apply (rewrite DERIV_cong_ev[of t t _ "\_. 0" 0 0], simp_all add: that) apply (rewrite eventually_nhds) apply (rule exI[of _ "ball t ?e"], auto simp add: that) by (rule cdfTx_nonpos_0; simp add: dist_real_def) thus ?thesis unfolding pdfT_def by (meson DERIV_imp_deriv) qed lemma pdfTx_0_0: "pdfT x 0 = 0" proof (cases \cdf (distr (\ \ alive x) borel (T x)) differentiable at 0\) case True let ?cdfx = "cdf (distr (\ \ alive x) borel (T x))" have "(?cdfx has_real_derivative 0) (at 0)" proof - from True obtain c where cdfx_deriv: "(?cdfx has_real_derivative c) (at 0)" unfolding differentiable_def using has_real_derivative by blast hence "(?cdfx has_real_derivative c) (at 0 within {..0})" by (rule has_field_derivative_at_within) moreover have "(?cdfx has_real_derivative 0) (at 0 within {..0})" proof - have "\\<^sub>F t in at 0 within {..0}. ?cdfx t = 0" proof - { fix t assume "t \ {..0::real}" "t \ 0" "dist t 0 < 1" hence "?cdfx t = 0" using cdfTx_nonpos_0 x_lt_psi by blast } hence "\d>0::real. \t\{..0}. t\0 \ dist t 0 < d \ ?cdfx t = 0" by (smt (verit)) thus ?thesis by (rewrite eventually_at) simp qed moreover have "?cdfx 0 = 0" proof - have "continuous (at 0 within {..0}) ?cdfx" using True differentiable_imp_continuous_within differentiable_subset by blast thus ?thesis by (simp add: cdfTx_nonpos_0) qed ultimately show ?thesis by (rewrite has_field_derivative_cong_eventually[of _ "\_. 0::real" 0 "{..0}" 0]; simp) qed ultimately have "c = 0" using has_real_derivative_iff_has_vector_derivative apply (intro vector_derivative_unique_within[of 0 "{..0}" ?cdfx c 0]; blast?) by (rewrite at_within_eq_bot_iff) (metis closure_lessThan islimpt_in_closure limpt_of_closure trivial_limit_at_left_real trivial_limit_within) thus "(?cdfx has_real_derivative 0) (at 0)" using cdfx_deriv by simp qed thus ?thesis unfolding pdfT_def by (meson DERIV_imp_deriv) next case False thus ?thesis unfolding pdfT_def by simp qed lemma pdfTx_nonpos_0: "pdfT x t = 0" if "t \ 0" for t::real using pdfTx_neg_0 pdfTx_0_0 that by force lemma pdfTx_beyond_0: "pdfT x t = 0" if "x+t \ $\" for t::real proof (cases \t \ 0\) case True thus ?thesis using pdfTx_nonpos_0 by simp next let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" case False hence t_pos: "t > 0" by simp thus ?thesis proof - have "(?cdfTx has_real_derivative 0) (at_right t)" apply (rewrite has_field_derivative_cong_eventually[where g="\_. 1"], simp_all) apply (rewrite eventually_at_right_field) using that cdfTx_1_equiv by (intro exI[of _"t+1"], auto simp add: le_ereal_less less_eq_ereal_def) thus ?thesis unfolding pdfT_def by (meson has_real_derivative_iff_has_vector_derivative has_vector_derivative_at_within DERIV_deriv_iff_real_differentiable trivial_limit_at_right_real vector_derivative_unique_within) qed qed lemma pdfTx_pdfX: "pdfT x t = pdfX (x+t) / \

(\ in \. X \ > x)" if "t > 0" for t::real proof (cases \cdf (distr \ borel X) differentiable at (x+t)\) case True let ?cdfX = "cdf (distr \ borel X)" let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" have [simp]: "?cdfTx differentiable at t" using differentiable_cdfX_cdfTx that True by simp have "pdfT x t = deriv ?cdfTx t" unfolding pdfT_def using that differentiable_cdfX_cdfTx by simp hence "(?cdfTx has_field_derivative (pdfT x t)) (at t)" using True DERIV_deriv_iff_real_differentiable by simp moreover have "\u. dist u t < t \ ?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1) = ?cdfTx u" proof - fix u::real assume "dist u t < t" hence [simp]: "u > 0" using dist_real_def by fastforce have "?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1) = (1 - \

(\ in \. X \ > x+u)) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1)" using MM_PS.ccdf_distr_P X_RV distrX_RD.cdf_ccdf distrX_RD.prob_space by presburger also have "\ = 1 - \

(\ in \. X \ > x+u) / \

(\ in \. X \ > x)" by (simp add: diff_divide_distrib) also have "\ = ?cdfTx u" apply (rewrite ccdfTx_PX[THEN sym], simp_all add: less_eq_real_def) using distrTx_RD.cdf_ccdf distrTx_RD.prob_space by presburger finally show "?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1) = ?cdfTx u" . qed ultimately have "((\u. ?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1)) has_field_derivative (pdfT x t)) (at t)" apply - by (rule has_field_derivative_transform_within[where d=t]; simp add: that) hence "((\u. ?cdfX (x+u) / \

(\ in \. X \ > x)) has_field_derivative (pdfT x t)) (at t)" unfolding has_field_derivative_def using has_derivative_add_const[where c="1 / \

(\ in \. X \ > x) - 1"] by force hence "((\u. ?cdfX (x+u) / \

(\ in \. X \ > x) * \

(\ in \. X \ > x)) has_field_derivative pdfT x t * \

(\ in \. X \ > x)) (at t)" using DERIV_cmult_right[where c="\

(\ in \. X \ > x)"] by force hence "((\u. ?cdfX (x+u)) has_field_derivative pdfT x t * \

(\ in \. X \ > x)) (at t)" unfolding has_field_derivative_def using has_derivative_transform PXx_pos x_lt_psi by (smt (verit) Collect_cong UNIV_I nonzero_mult_div_cancel_right times_divide_eq_left) hence "(?cdfX has_field_derivative pdfT x t * \

(\ in \. X \ > x)) (at (x+t))" using DERIV_at_within_shift by force moreover have "(?cdfX has_field_derivative deriv ?cdfX (x+t)) (at (x+t))" using True DERIV_deriv_iff_real_differentiable by blast ultimately have "pdfT x t * \

(\ in \. X \ > x) = deriv ?cdfX (x+t)" by (simp add: DERIV_imp_deriv) thus "pdfT x t = pdfX (x+t) / \

(\ in \. X \ > x)" unfolding pdfX_def using True by simp (metis PXx_pos nonzero_mult_div_cancel_right x_lt_psi zero_less_measure_iff) next case False hence [simp]: "\ cdf (distr (\ \ alive x) borel (T x)) differentiable at t" using differentiable_cdfX_cdfTx that by simp thus "pdfT x t = pdfX (x+t) / \

(\ in \. X \ > x)" unfolding pdfT_def pdfX_def using False by simp qed lemma pdfTx_measurable[measurable]: "pdfT x \ borel_measurable borel" proof - let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" have "countable {x. \ ?cdfTx differentiable at x}" using cdfX_nondifferentiable_finite_set uncountable_infinite by force thus ?thesis unfolding pdfT_def apply - by (rule measurable_discrete_difference [where X="{x. \ ?cdfTx differentiable at x}" and f="deriv ?cdfTx"]; simp) qed lemma distributed_pdfTx: "distributed (\ \ alive x) lborel (T x) (pdfT x)" proof - let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" obtain S where fin: "finite S" and diff: "\t. t \ S \ ?cdfTx differentiable at t" using cdfTx_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast { fix t::real assume t_notin: "t \ S" have "?cdfTx differentiable at t" using diff t_notin by simp hence "(?cdfTx has_real_derivative pdfT x t) (at t)" unfolding pdfT_def using DERIV_deriv_iff_real_differentiable by force } thus ?thesis by (intro alivex_PS.distributed_deriv_cdf[where S=S]; simp add: fin) qed lemma nn_integral_pdfTx_1: "(\\<^sup>+s. pdfT x s \lborel) = 1" proof - have "(\\<^sup>+s. pdfT x s \lborel) = emeasure (density lborel (pdfT x)) UNIV" by (rewrite emeasure_density) simp_all also have "\ = emeasure (distr (\ \ alive x) lborel (T x)) UNIV" using distributed_pdfTx unfolding distributed_def by simp also have "\ = 1" by (metis distrTx_RD.emeasure_space_1 distrTx_RD.space_eq_univ distr_cong sets_lborel) finally show ?thesis . qed corollary has_bochner_integral_pdfTx_1: "has_bochner_integral lborel (pdfT x) 1" by (rule has_bochner_integral_nn_integral; simp add: pdfTx_nonneg nn_integral_pdfTx_1) corollary LBINT_pdfTx_1: "LBINT s. pdfT x s = 1" using has_bochner_integral_pdfTx_1 by (simp add: has_bochner_integral_integral_eq) corollary pdfTx_has_integral_1: "(pdfT x has_integral 1) UNIV" by (rule nn_integral_has_integral; simp add: pdfTx_nonneg nn_integral_pdfTx_1) lemma set_nn_integral_pdfTx_1: "\\<^sup>+s\{0..}. pdfT x s \lborel = 1" apply (rewrite nn_integral_pdfTx_1[THEN sym]) apply (rule nn_integral_cong) using pdfTx_nonpos_0 by (metis atLeast_iff ennreal_0 indicator_simps(1) linorder_le_cases mult.commute mult_1 mult_zero_left) corollary has_bochner_integral_pdfTx_1_nonpos: "has_bochner_integral lborel (\s. pdfT x s * indicator {0..} s) 1" apply (rule has_bochner_integral_nn_integral, simp_all) using pdfTx_nonneg apply simp using set_nn_integral_pdfTx_1 by (simp add: nn_integral_set_ennreal) corollary set_LBINT_pdfTx_1: "(LBINT s:{0..}. pdfT x s) = 1" unfolding set_lebesgue_integral_def using has_bochner_integral_pdfTx_1_nonpos has_bochner_integral_integral_eq apply (simp, rewrite mult.commute) by (smt (verit) Bochner_Integration.integral_cong has_bochner_integral_integral_eq) corollary pdfTx_has_integral_1_nonpos: "(pdfT x has_integral 1) {0..}" proof - have "set_integrable lebesgue {0..} (pdfT x)" unfolding set_integrable_def apply (rewrite integrable_completion, simp_all) using has_bochner_integral_pdfTx_1_nonpos by (rewrite mult.commute, rule integrable.intros) moreover have "LINT s:{0..}|lebesgue. pdfT x s = 1" using set_LBINT_pdfTx_1 unfolding set_lebesgue_integral_def by (rewrite integral_completion; simp) ultimately show ?thesis using has_integral_set_lebesgue by force qed lemma set_nn_integral_pdfTx_PTx: "\\<^sup>+s\A. pdfT x s \lborel = \

(\ in \. T x \ \ A \ T x \ > 0)" if "A \ sets lborel" for A :: "real set" proof - have [simp]: "A \ sets borel" using that by simp have "\\<^sup>+s\A. pdfT x s \lborel = emeasure (density lborel (pdfT x)) A" using that by (rewrite emeasure_density; force) also have "\ = emeasure (distr (\ \ alive x) lborel (T x)) A" using distributed_pdfTx unfolding distributed_def by simp also have "\ = ennreal \

(\ in (\ \ alive x). T x \ \ A)" apply (rewrite emeasure_distr, simp_all) apply (rewrite finite_measure.emeasure_eq_measure, simp) by (smt (verit) Collect_cong vimage_eq Int_def) also have "\ = ennreal \

(\ in \. T x \ \ A \ T x \ > 0)" unfolding alive_def apply (rewrite MM_PS.cond_prob_space_cond_prob[THEN sym], simp_all add: pred_def futr_life_def) using borel_measurable_diff X_RV that by measurable finally show ?thesis . qed lemma pdfTx_set_integrable: "set_integrable lborel A (pdfT x)" if "A \ sets lborel" unfolding set_integrable_def - using that pdfTx_nonneg apply (intro integrableI_real_bounded, simp_all) + using that pdfTx_nonneg apply (intro integrableI_nonneg, simp_all) apply (rewrite mult.commute) using set_nn_integral_pdfTx_PTx that by (metis (no_types, lifting) ennreal_indicator ennreal_less_top ennreal_mult' nn_integral_cong) lemma set_integral_pdfTx_PTx: "LBINT s:A. pdfT x s = \

(\ in \. T x \ \ A \ T x \ > 0)" if "A \ sets lborel" for A :: "real set" unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral) using that apply (simp_all add: pdfTx_nonneg) apply (rewrite indicator_mult_ennreal[THEN sym], rewrite mult.commute) by (rewrite set_nn_integral_pdfTx_PTx; simp) lemma pdfTx_has_integral_PTx: "(pdfT x has_integral \

(\ in \. T x \ \ A \ T x \ > 0)) A" if "A \ sets lborel" for A :: "real set" proof - have "((\s. pdfT x s * indicat_real A s) has_integral \

(\ in \. T x \ \ A \ T x \ > 0)) UNIV" using that pdfTx_nonneg apply (intro nn_integral_has_integral, simp_all) using set_nn_integral_pdfTx_PTx that by (simp add: nn_integral_set_ennreal) thus ?thesis by (smt (verit) has_integral_cong has_integral_restrict_UNIV indicator_eq_0_iff indicator_simps(1) mult_cancel_left2 mult_eq_0_iff) qed corollary pdfTx_has_integral_PTx_Icc: "(pdfT x has_integral \

(\ in \. a \ T x \ \ T x \ \ b \ T x \ > 0)) {a..b}" for a b :: real using pdfTx_has_integral_PTx[of "{a..b}"] by simp corollary pdfTx_integrable_on_Icc: "pdfT x integrable_on {a..b}" for a b :: real using pdfTx_has_integral_PTx_Icc by auto end subsubsection \Properties of Probability Density Function of X\ lemma pdfX_nonneg: "pdfX x \ 0" for x::real using pdfTx_nonneg pdfT0_X psi_pos' by smt lemma pdfX_nonpos_0: "pdfX x = 0" if "x \ 0" for x::real using pdfTx_nonpos_0 pdfT0_X psi_pos' that by smt lemma pdfX_beyond_0: "pdfX x = 0" if "x \ $\" for x::real using pdfTx_beyond_0 pdfT0_X psi_pos' that by smt lemma nn_integral_pdfX_1: "integral\<^sup>N lborel pdfX = 1" using nn_integral_pdfTx_1 pdfT0_X psi_pos' by metis corollary has_bochner_integral_pdfX_1: "has_bochner_integral lborel pdfX 1" by (rule has_bochner_integral_nn_integral; simp add: pdfX_nonneg nn_integral_pdfX_1) corollary LBINT_pdfX_1: "LBINT s. pdfX s = 1" using has_bochner_integral_pdfX_1 by (simp add: has_bochner_integral_integral_eq) corollary pdfX_has_integral_1: "(pdfX has_integral 1) UNIV" by (rule nn_integral_has_integral; simp add: pdfX_nonneg nn_integral_pdfX_1) lemma set_nn_integral_pdfX_PX: "set_nn_integral lborel A pdfX = \

(\ in \. X \ \ A)" if "A \ sets lborel" for A :: "real set" using PT0_eq_PX_lborel that by (rewrite pdfT0_X[THEN sym], rewrite set_nn_integral_pdfTx_PTx; simp) lemma pdfX_set_integrable: "set_integrable lborel A pdfX" if "A \ sets lborel" for A :: "real set" using pdfTx_set_integrable pdfT0_X psi_pos' that by smt lemma set_integral_pdfX_PX: "LBINT s:A. pdfX s = \

(\ in \. X \ \ A)" if "A \ sets lborel" for A :: "real set" using PT0_eq_PX_lborel that by (rewrite pdfT0_X[THEN sym], rewrite set_integral_pdfTx_PTx; simp) lemma pdfX_has_integral_PX: "(pdfX has_integral \

(\ in \. X \ \ A)) A" if "A \ sets lborel" for A :: "real set" using that apply (rewrite pdfT0_X[THEN sym], rewrite PT0_eq_PX_lborel[THEN sym], simp) by (rule pdfTx_has_integral_PTx; simp) corollary pdfX_has_integral_PX_Icc: "(pdfX has_integral \

(\ in \. a \ X \ \ X \ \ b)) {a..b}" for a b :: real using pdfX_has_integral_PX[of "{a..b}"] by simp corollary pdfX_integrable_on_Icc: "pdfX integrable_on {a..b}" for a b :: real using pdfX_has_integral_PX_Icc by auto subsubsection \Relations between Life Expectation and Probability Density Function\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct; simp add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp proposition nn_integral_T_pdfT: "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) = \\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel" if "g \ borel_measurable lborel" for g :: "real \ real" proof - have "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) = \\<^sup>+s. ennreal (pdfT x s) * ennreal (g s) \lborel" proof - have "distributed (\ \ alive x) lborel (T x) (\s. ennreal (pdfT x s))" by (intro distributed_pdfTx) simp moreover have "(\s. ennreal (g s)) \ borel_measurable borel" using that by measurable ultimately show ?thesis by (rewrite distributed_nn_integral; simp) qed also have "\ = \\<^sup>+s. ennreal (pdfT x s * g s) \lborel" using ennreal_mult' pdfTx_nonneg by force also have "\ = \\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel" apply (rule nn_integral_cong, simp) by (metis atLeast_iff ennreal_0 indicator_simps linorder_not_le mult_1 mult_commute_abs mult_zero_left pdfTx_neg_0 x_lt_psi) finally show ?thesis . qed lemma expectation_LBINT_pdfT_nonneg: "alivex_PS.expectation (\\. g (T x \)) = LBINT s:{0..}. pdfT x s * g s" if "\s. s \ 0 \ g s \ 0" "g \ borel_measurable lborel" for g :: "real \ real" \ \Note that 0 = 0 holds when the integral diverges.\ using that apply (rewrite integral_eq_nn_integral, simp) apply (rule AE_I2, metis alivex_Tx_pos less_imp_le) unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral, simp_all) apply (rule AE_I2, metis indicator_pos_le pdfTx_nonpos_0 x_lt_psi zero_le_mult_iff zero_le_square pdfTx_nonneg) by (rewrite nn_integral_T_pdfT) (simp_all add: indicator_mult_ennreal mult.commute) corollary expectation_integral_pdfT_nonneg: "alivex_PS.expectation (\\. g (T x \)) = integral {0..} (\s. pdfT x s * g s)" if "\s. s \ 0 \ g s \ 0" "g \ borel_measurable lborel" for g :: "real \ real" \ \Note that 0 = 0 holds when the integral diverges.\ proof - have "alivex_PS.expectation (\\. g (T x \)) = LBINT s:{0..}. pdfT x s * g s" using expectation_LBINT_pdfT_nonneg that by simp also have "\ = integral {0..} (\s. pdfT x s * g s)" using that pdfTx_nonneg by (intro set_borel_integral_eq_integral_nonneg; simp) finally show ?thesis . qed proposition expectation_LBINT_pdfT: "alivex_PS.expectation (\\. g (T x \)) = LBINT s:{0..}. pdfT x s * g s" if "set_integrable lborel {0..} (\s. pdfT x s * g s)" "g \ borel_measurable lborel" for g :: "real \ real" proof - have pg_fin: "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) \ \" using that apply (rewrite nn_integral_T_pdfT, simp) using that unfolding set_integrable_def apply (rewrite in asm real_integrable_def, simp) by (simp add: indicator_mult_ennreal mult.commute) moreover have mg_fin: "(\\<^sup>+\. ennreal (- g (T x \)) \(\ \ alive x)) \ \" using that apply (rewrite nn_integral_T_pdfT[of "\s. - g s"], simp) using that unfolding set_integrable_def apply (rewrite in asm real_integrable_def, simp) by (simp add: indicator_mult_ennreal mult.commute) ultimately have [simp]: "integrable (\ \ alive x) (\\. g (T x \))" using that by (rewrite real_integrable_def; simp) have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (g s)) \lborel) = (\\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel)" using SPMF.ennreal_max_0 ennreal_mult' pdfTx_nonneg x_lt_psi by presburger also have "\ < \" using pg_fin nn_integral_T_pdfT that top.not_eq_extremum by auto finally have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (g s)) \lborel) < \" . hence [simp]: "set_integrable lborel {0..} (\s. pdfT x s * max 0 (g s))" unfolding set_integrable_def using that apply (intro integrableI_nonneg, simp_all) using pdfTx_nonneg apply (intro AE_I2, force) by (metis (no_types, lifting) indicator_mult_ennreal mult.commute nn_integral_cong) have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (- g s)) \lborel) = (\\<^sup>+s\{0..}. ennreal (pdfT x s * - g s) \lborel)" using SPMF.ennreal_max_0 ennreal_mult' pdfTx_nonneg x_lt_psi by presburger also have "\ < \" using mg_fin nn_integral_T_pdfT[of "\s. - g s"] that top.not_eq_extremum by force finally have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (- g s)) \lborel) < \" . hence [simp]: "set_integrable lborel {0..} (\s. pdfT x s * max 0 (- g s))" unfolding set_integrable_def using that apply (intro integrableI_nonneg, simp_all) using pdfTx_nonneg apply (intro AE_I2, force) by (metis (no_types, lifting) indicator_mult_ennreal mult.commute nn_integral_cong) have "alivex_PS.expectation (\\. g (T x \)) = alivex_PS.expectation (\\. max 0 (g (T x \))) - alivex_PS.expectation (\\. max 0 (- g (T x \)))" by (rewrite Bochner_Integration.integral_cong [where N="\ \ alive x" and g="\\. max 0 (g (T x \)) - max 0 (- g (T x \))"]; simp) moreover have "alivex_PS.expectation (\\. max 0 (g (T x \))) = (LBINT s:{0..}. pdfT x s * max 0 (g s))" using that by (rewrite expectation_LBINT_pdfT_nonneg[where g="\s. max 0 (g s)"]) simp_all moreover have "alivex_PS.expectation (\\. max 0 (- g (T x \))) = (LBINT s:{0..}. pdfT x s * max 0 (- g s))" using that by (rewrite expectation_LBINT_pdfT_nonneg[where g="\s. max 0 (- g s)"]) simp_all ultimately have "alivex_PS.expectation (\\. g (T x \)) = (LBINT s:{0..}. pdfT x s * max 0 (g s)) - (LBINT s:{0..}. pdfT x s * max 0 (- g s))" by simp also have "\ = LBINT s:{0..}. (pdfT x s * max 0 (g s) - pdfT x s * max 0 (- g s))" by (rewrite set_integral_diff; simp) also have "\ = LBINT s:{0..}. pdfT x s * (max 0 (g s) - max 0 (- g s))" by (simp add: right_diff_distrib) also have "\ = LBINT s:{0..}. pdfT x s * g s" using minus_max_eq_min by (metis (no_types, opaque_lifting) diff_zero max_def min_def minus_diff_eq) finally show ?thesis . qed corollary expectation_integral_pdfT: "alivex_PS.expectation (\\. g (T x \)) = integral {0..} (\s. pdfT x s * g s)" if "(\s. pdfT x s * g s) absolutely_integrable_on {0..}" "g \ borel_measurable lborel" for g :: "real \ real" proof - have [simp]: "set_integrable lborel {0..} (\s. pdfT x s * g s)" using that by (rewrite absolutely_integrable_on_iff_set_integrable; simp) show ?thesis apply (rewrite set_borel_integral_eq_integral(2)[THEN sym], simp) using that by (rewrite expectation_LBINT_pdfT; simp) qed corollary e_LBINT_pdfT: "$e`\_x = LBINT s:{0..}. pdfT x s * s" \ \Note that 0 = 0 holds when the life expectation diverges.\ unfolding life_expect_def using expectation_LBINT_pdfT_nonneg by force corollary e_integral_pdfT: "$e`\_x = integral {0..} (\s. pdfT x s * s)" \ \Note that 0 = 0 holds when the life expectation diverges.\ unfolding life_expect_def using expectation_integral_pdfT_nonneg by force end corollary e_LBINT_pdfX: "$e`\_0 = (LBINT x:{0..}. pdfX x * x)" \ \Note that 0 = 0 holds when the life expectation diverges.\ using e_LBINT_pdfT pdfT0_X psi_pos' by presburger corollary e_integral_pdfX: "$e`\_0 = integral {0..} (\x. pdfX x * x)" \ \Note that 0 = 0 holds when the life expectation diverges.\ using e_integral_pdfT pdfT0_X psi_pos' by presburger subsubsection \Introduction of Force of Mortality\ definition force_mortal :: "real \ real" ("$\'__" [101] 200) where "$\_x \ MM_PS.hazard_rate X x" lemma mu_pdfX: "$\_x = pdfX x / ccdf (distr \ borel X) x" if "(cdf (distr \ borel X)) differentiable at x" for x::real unfolding force_mortal_def pdfX_def by (rewrite MM_PS.hazard_rate_deriv_cdf, simp_all add: that) lemma mu_unborn_0: "$\_x = 0" if "x < 0" for x::real apply (rewrite mu_pdfX) using cdfX_has_real_derivative_0_unborn real_differentiable_def that apply blast using pdfX_nonpos_0 that by auto lemma mu_beyond_0: "$\_x = 0" if "x \ $\" for x::real \ \Note that division by 0 is defined as 0 in Isabelle/HOL.\ unfolding force_mortal_def using MM_PS.hazard_rate_0_ccdf_0 ccdfX_0_equiv that by simp lemma mu_nonneg_differentiable: "$\_x \ 0" if "(cdf (distr \ borel X)) differentiable at x" for x::real apply (rewrite mu_pdfX, simp add: that) using pdfX_nonneg distrX_RD.ccdf_nonneg by simp lemma mu_nonneg_AE: "AE x in lborel. $\_x \ 0" using cdfX_differentiable_AE mu_nonneg_differentiable by auto lemma mu_measurable[measurable]: "(\x. $\_x) \ borel_measurable borel" proof - obtain S where "finite S" and "\x. x \ S \ (cdf (distr \ borel X)) differentiable at x" using cdfX_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast thus ?thesis apply (rewrite measurable_discrete_difference [where f="\x. pdfX x / ccdf (distr \ borel X) x" and X=S], simp_all) by (simp_all add: MM_PS.ccdf_distr_measurable borel_measurable_divide countable_finite mu_pdfX) qed lemma mu_deriv_ccdf: "$\_x = - deriv (ccdf (distr \ borel X)) x / ccdf (distr \ borel X) x" if "(ccdf (distr \ borel X)) differentiable at x" "x < $\" for x::real unfolding force_mortal_def by (rewrite MM_PS.hazard_rate_deriv_ccdf; simp add: that) lemma mu_deriv_ln: "$\_x = - deriv (\x. ln (ccdf (distr \ borel X) x)) x" if "(ccdf (distr \ borel X)) differentiable at x" "x < $\" for x::real unfolding force_mortal_def apply (rewrite MM_PS.hazard_rate_deriv_ln_ccdf, simp_all add: that) using ccdfX_0_equiv that by force lemma p_exp_integral_mu: "$p_{t&x} = exp (- integral {x..x+t} (\y. $\_y))" if "x \ 0" "t \ 0" "x+t < $\" proof - have [simp]: "x < $\" using that by (meson add_increasing2 ereal_less linorder_not_le) have "$p_{t&x} = (ccdf (distr \ borel X) (x+t)) / (ccdf (distr \ borel X) x)" apply (rewrite p_PX, simp_all add: that) by (rewrite MM_PS.ccdf_distr_P, simp)+ simp also have "\ = exp (- integral {x..x+t} (MM_PS.hazard_rate X))" apply (rule MM_PS.ccdf_exp_cumulative_hazard, simp_all add: that) using ccdfX_piecewise_differentiable piecewise_differentiable_on_subset apply blast using ccdfX_continuous continuous_on_subset apply blast using ccdfX_0_equiv that ereal_less_le linorder_not_le by force also have "\ = exp (- integral {x..x+t} (\y. $\_y))" unfolding force_mortal_def by simp finally show ?thesis . qed corollary ccdfX_exp_integral_mu: "ccdf (distr \ borel X) x = exp (- integral {0..x} (\y. $\_y))" if "0 \ x \ x < $\" for x::real by (rewrite p_exp_integral_mu[where t=x and x=0, simplified, THEN sym]; simp add: that ccdfX_p) subsubsection \Properties of Force of Mortality\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct; simp add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma hazard_rate_Tx_mu: "alivex_PS.hazard_rate (T x) t = $\_(x+t)" if "t \ 0" "x+t < $\" for t::real proof - have [simp]: "\

(\ in \. X \ > x) > 0" by force moreover have [simp]: "\

(\ in \. X \ > x + t) > 0" using that by force moreover have [simp]: "{\ \ space (\ \ alive x). X \ > x + t} = {\ \ space \. X \ > x + t}" unfolding alive_def using that by (rewrite MM_PS.space_cond_prob_space, simp_all, force) hence [simp]: "{\ \ space (\ \ alive x). X \ > x + t} \ MM_PS.events" by simp ultimately have \[simp]: "\

(\ in (\ \ alive x). X \ > x + t) > 0" unfolding alive_def apply (rewrite MM_PS.cond_prob_space_cond_prob[THEN sym], (simp_all add: pred_def)[3]) unfolding cond_prob_def by (smt (verit) Collect_cong divide_pos_pos) have "alivex_PS.hazard_rate (T x) t = Lim (at_right 0) (\dt. \

(\ in (\ \ alive x). t < T x \ \ T x \ \ t + dt \ T x \ > t) / dt)" unfolding alivex_PS.hazard_rate_def by simp also have "\ = Lim (at_right 0) (\dt. \

(\ in (\ \ alive x). x + t < X \ \ X \ \ x + t + dt \ X \ > x + t) / dt)" apply (rule Lim_cong, rule eventually_at_rightI[of 0 1], simp_all) unfolding futr_life_def cond_prob_def using Collect_cong by smt also have "\ = Lim (at_right 0) (\dt. \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t) / dt)" proof - have "\dt. \

(\ in (\ \ alive x). x + t < X \ \ X \ \ x + t + dt \ X \ > x + t) = \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t)" proof - fix dt let ?rngX = "\\. x + t < X \ \ X \ \ x + t + dt" have "\

(\ in (\ \ alive x). ?rngX \ \ X \ > x + t) = \

(\ in ((\ \ alive x) \ {\ \ space (\ \ alive x). X \ > x + t}). ?rngX \)" using \ by (rewrite alivex_PS.cond_prob_space_cond_prob) simp_all also have "\ = \

(\ in (\ \ {\ \ space \. X \ > x + t}). ?rngX \)" proof - have "(\ \ alive x) \ {\ \ space (\ \ alive x). X \ > x + t} = \ \ {\ \ space (\ \ alive x). X \ > x + t}" apply (rewrite MM_PS.cond_cond_prob_space, simp_all) unfolding alive_def using that by force also have "\ = \ \ {\ \ space \. X \ > x + t}" by simp finally have "(\ \ alive x) \ {\ \ space (\ \ alive x). X \ > x + t} = \ \ {\ \ space \. X \ > x + t}" . thus ?thesis by simp qed also have "\ = \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t)" by (rewrite MM_PS.cond_prob_space_cond_prob, simp_all add: pred_def sets.sets_Collect_conj) finally show "\

(\ in (\ \ alive x). ?rngX \ \ X \ > x + t) = \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t)" . qed thus ?thesis apply - by (rule Lim_cong, rule eventually_at_rightI[of 0 1]; simp) qed also have "\ = $\_(x+t)" unfolding force_mortal_def MM_PS.hazard_rate_def by simp finally show ?thesis . qed lemma pdfTx_p_mu: "pdfT x t = $p_{t&x} * $\_(x+t)" if "(cdf (distr (\ \ alive x) borel (T x))) differentiable at t" "t > 0" for t::real proof (cases \x+t < $\\) case True hence [simp]: "t \ 0" and "(ccdf (distr (\ \ alive x) borel (T x))) t \ 0" using that p_0_equiv unfolding survive_def by auto have "deriv (cdf (distr (\ \ alive x) borel (T x))) t = ccdf (distr (\ \ alive x) borel (T x)) t * alivex_PS.hazard_rate (T x) t" by (rule alivex_PS.deriv_cdf_hazard_rate; simp add: that) thus ?thesis unfolding survive_def pdfT_def using hazard_rate_Tx_mu True that by simp next case False hence "x+t \ $\" by simp thus ?thesis using pdfTx_beyond_0 mu_beyond_0 by simp qed lemma deriv_t_p_mu: "deriv (\s. $p_{s&x}) t = - $p_{t&x} * $\_(x+t)" if "(\s. $p_{s&x}) differentiable at t" "t > 0" for t::real proof - let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" have diff: "?cdfTx differentiable at t" using that distrTx_RD.differentiable_cdf_ccdf unfolding survive_def by blast hence "deriv ?cdfTx t = $p_{t&x} * $\_(x+t)" using that pdfTx_p_mu pdfT_def by simp moreover have "deriv ?cdfTx t = - deriv (\s. $p_{s&x}) t" using that diff distrTx_RD.deriv_cdf_ccdf unfolding survive_def by presburger ultimately show ?thesis by simp qed lemma pdfTx_p_mu_AE: "AE s in lborel. s > 0 \ pdfT x s = $p_{s&x} * $\_(x+s)" proof - let ?cdfX = "cdf (distr \ borel X)" let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" from cdfX_differentiable_AE obtain N where diff: "\r. r \ space lborel - N \ ?cdfX differentiable at r" and null: "N \ null_sets lborel" using AE_E3 by blast let ?N' = "{s. x+s \ N}" have "?N' \ null_sets lborel" using null_sets_translation[of N "-x", simplified] null by (simp add: add.commute) moreover have "{s \ space lborel. \ (s > 0 \ pdfT x s = $p_{s&x} * $\_(x+s))} \ ?N'" proof (rewrite Compl_subset_Compl_iff[THEN sym], safe) fix s::real assume "s \ space lborel" and "x+s \ N" and "s > 0" thus "pdfT x s = $p_{s&x} * $\_(x+s)" apply (intro pdfTx_p_mu, simp_all) by (rewrite differentiable_cdfX_cdfTx[THEN sym]; simp add: diff) qed ultimately show ?thesis using AE_I'[of ?N'] by simp qed lemma LBINT_p_mu_q: "LBINT s:{f<..f+t}. $p_{s&x} * $\_(x+s) = $q_{f\t&x}" if "t \ 0" "f \ 0" for t f :: real proof - have "LBINT s:{f<..f+t}. $p_{s&x} * $\_(x+s) = LBINT s:{f<..f+t}. pdfT x s" apply (rule set_lebesgue_integral_cong_AE; simp) apply (simp add: survive_def) using pdfTx_p_mu_AE apply (rule AE_mp) using that by (intro always_eventually; simp add: ereal_less_le) also have "\ = enn2real (\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel)" proof - have "\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel < \" proof - have "\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel \ \\<^sup>+s. ennreal (pdfT x s) \lborel" by (smt (verit) indicator_simps le_zero_eq linorder_le_cases mult.right_neutral mult_zero_right nn_integral_mono) also have "\ < \" using nn_integral_pdfTx_1 by simp finally show ?thesis . qed thus ?thesis unfolding set_lebesgue_integral_def by (rewrite enn2real_nn_integral_eq_integral[where g="\s. pdfT x s * indicator {f<..f+t} s"]) (simp_all add: pdfTx_nonneg mult.commute ennreal_indicator ennreal_mult') qed also have "\ = measure (density lborel (pdfT x)) {f<..f+t}" unfolding measure_def by (rewrite emeasure_density; simp) also have "\ = measure (distr (\ \ alive x) lborel (T x)) {f<..f+t}" using distributed_pdfTx unfolding distributed_def by simp also have "\ = cdf (distr (\ \ alive x) lborel (T x)) (f+t) - cdf (distr (\ \ alive x) lborel (T x)) f" using that finite_borel_measure.cdf_diff_eq2 by (smt (verit) distrTx_RD.finite_borel_measure_axioms distr_cong sets_lborel) also have "\ = $q_{f\t&x}" using q_defer_q die_def that by (metis distr_cong sets_lborel x_lt_psi) finally show ?thesis . qed lemma set_integrable_p_mu: "set_integrable lborel {f<..f+t} (\s. $p_{s&x} * $\_(x+s))" if "t \ 0" "f \ 0" for t f :: real proof - have "(\s. $p_{s&x}) \ borel_measurable borel" unfolding survive_def by simp moreover have "AE s in lborel. f < s \ s \ f + t \ $p_{s&x} * $\_(x+s) = pdfT x s" using pdfTx_p_mu_AE apply (rule AE_mp) using that by (intro always_eventually; simp add: ereal_less_le) moreover have "set_integrable lborel {f<..f+t} (pdfT x)" using pdfTx_set_integrable by simp ultimately show ?thesis by (rewrite set_integrable_cong_AE[where g="pdfT x"]; simp) qed lemma p_mu_has_integral_q_defer_Ioc: "((\s. $p_{s&x} * $\_(x+s)) has_integral $q_{f\t&x}) {f<..f+t}" if "t \ 0" "f \ 0" for t f :: real apply (rewrite LBINT_p_mu_q[THEN sym], simp_all add: that) apply (rewrite set_borel_integral_eq_integral, simp add: set_integrable_p_mu that) by (rewrite has_integral_integral[THEN sym]; simp add: set_borel_integral_eq_integral set_integrable_p_mu that) lemma p_mu_has_integral_q_defer_Icc: "((\s. $p_{s&x} * $\_(x+s)) has_integral $q_{f\t&x}) {f..f+t}" if "t \ 0" "f \ 0" for t f :: real proof - have "negligible {f}" by simp hence [simp]: "negligible ({f..f+t} - {f<..f+t})" by (smt (verit) Diff_iff atLeastAtMost_iff greaterThanAtMost_iff insertI1 negligible_sing negligible_subset subsetI) have [simp]: "negligible ({f<..f+t} - {f..f+t})" by (simp add: subset_eq) show ?thesis apply (rewrite has_integral_spike_set_eq[where T="{f<..f+t}"]) apply (rule negligible_subset[of "{f..f+t} - {f<..f+t}"], simp, blast) apply (rule negligible_subset[of "{f<..f+t} - {f..f+t}"], simp, blast) using p_mu_has_integral_q_defer_Ioc that by simp qed corollary p_mu_has_integral_q_Icc: "((\s. $p_{s&x} * $\_(x+s)) has_integral $q_{t&x}) {0..t}" if "t \ 0" for t::real using p_mu_has_integral_q_defer_Icc[where f=0] that by simp corollary p_mu_integrable_on_Icc: "(\s. $p_{s&x} * $\_(x+s)) integrable_on {0..t}" if "t \ 0" for t::real using p_mu_has_integral_q_Icc that by auto lemma e_ennreal_p_mu: "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = \\<^sup>+s\{0..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel" proof - have [simp]: "sym_diff {0..} {0<..} = {0::real}" by force have "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = \\<^sup>+s\{0..}. ennreal (pdfT x s * s) \lborel" by (rewrite nn_integral_T_pdfT[where g="\s. s"]; simp) also have "\ = \\<^sup>+s\{0<..}. ennreal (pdfT x s * s) \lborel" by (rewrite nn_integral_null_delta; force) also have "\ = \\<^sup>+s\{0<..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel" apply (rule nn_integral_cong_AE) using pdfTx_p_mu_AE apply (rule AE_mp, intro AE_I2) by force also have "\ = \\<^sup>+s\{0..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel" by (rewrite nn_integral_null_delta[THEN sym]; force) finally show ?thesis . qed lemma e_LBINT_p_mu: "$e`\_x = LBINT s:{0..}. $p_{s&x} * $\_(x+s) * s" \ \Note that 0 = 0 holds when the life expectation diverges.\ proof - let ?f = "\s. $p_{s&x} * $\_(x+s) * s" have [simp]: "(\s. ?f s * indicat_real {0..} s) \ borel_measurable borel" by measurable (simp_all add: survive_def) hence [simp]: "(\s. indicat_real {0..} s * ?f s) \ borel_measurable borel" by (meson measurable_cong mult.commute) have [simp]: "AE s in lborel. ?f s * indicat_real {0..} s \ 0" proof - have "AE s in lborel. pdfT x s * s * indicat_real {0..} s \ 0" using pdfTx_nonneg pdfTx_nonpos_0 x_lt_psi by (intro AE_I2) (smt (verit, del_insts) indicator_pos_le mult_eq_0_iff mult_nonneg_nonneg) thus ?thesis apply (rule AE_mp) using pdfTx_p_mu_AE apply (rule AE_mp) by (rule AE_I2) (metis atLeast_iff indicator_simps(2) mult_eq_0_iff order_le_less) qed hence [simp]: "AE s in lborel. indicat_real {0..} s * ?f s \ 0" by (metis (no_types, lifting) AE_cong mult.commute) show ?thesis proof (cases \integrable (\ \ alive x) (T x)\) case True hence "ennreal ($e`\_x) = \\<^sup>+\. ennreal (T x \) \(\ \ alive x)" unfolding life_expect_def apply (rewrite nn_integral_eq_integral, simp_all) by (smt (verit) AE_I2 alivex_Tx_pos) also have "\ = \\<^sup>+s. ennreal (?f s * indicat_real {0..} s) \lborel" by (simp add: nn_integral_set_ennreal e_ennreal_p_mu) also have "\ = ennreal (LBINT s:{0..}. ?f s)" proof - have "integrable lborel (\s. ?f s * indicat_real {0..} s)" proof - have "(\\<^sup>+s. ennreal (?f s * indicat_real {0..} s) \lborel) < \" by (metis calculation ennreal_less_top infinity_ennreal_def) thus ?thesis by (intro integrableI_nonneg; simp) qed thus ?thesis unfolding set_lebesgue_integral_def by (rewrite nn_integral_eq_integral, simp_all) (meson mult.commute) qed finally have "ennreal ($e`\_x) = ennreal (LBINT s:{0..}. ?f s)" . moreover have "LBINT s:{0..}. ?f s \ 0" unfolding set_lebesgue_integral_def by (rule integral_nonneg_AE) simp ultimately show ?thesis using e_nonneg by simp next case False hence "$e`\_x = 0" unfolding life_expect_def using not_integrable_integral_eq by force also have "\ = LBINT s:{0..}. ?f s" proof - have "\ = \\<^sup>+\. ennreal (T x \) \(\ \ alive x)" using nn_integral_nonneg_infinite False by (smt (verit) AE_cong Tx_alivex_measurable alivex_PS.AE_prob_1 alivex_PS.prob_space alivex_Tx_pos nn_integral_cong) hence "0 = enn2real (\\<^sup>+s\{0..}. ennreal (?f s) \lborel)" using e_ennreal_p_mu by simp also have "\ = LBINT s:{0..}. ?f s" unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral, simp_all) by (simp add: indicator_mult_ennreal mult.commute) finally show ?thesis by simp qed finally show ?thesis . qed qed lemma e_integral_p_mu: "$e`\_x = integral {0..} (\s. $p_{s&x} * $\_(x+s) * s)" \ \Note that 0 = 0 holds when the life expectation diverges.\ proof - have "LBINT s:{0..}. $p_{s&x} * $\_(x+s) * s = integral {0..} (\s. $p_{s&x} * $\_(x+s) * s)" proof - have "AE s in lborel. s \ 0 \ $p_{s&x} * $\_(x+s) * s \ 0" proof - have "AE s in lborel. $\_(x+s) \ 0" by (rule AE_translation, rule mu_nonneg_AE) with p_nonneg show ?thesis by force qed moreover have "(\s. $p_{s&x} * $\_(x+s) * s) \ borel_measurable borel" unfolding survive_def by simp ultimately show ?thesis by (intro set_borel_integral_eq_integral_nonneg_AE; simp) qed thus ?thesis using e_LBINT_p_mu by simp qed end lemma p_has_real_derivative_x_ccdfX: "((\y. $p_{t&y}) has_real_derivative ((deriv svl (x+t) * svl x - svl (x+t) * deriv svl x) / (svl x)\<^sup>2)) (at x)" if "svl \ ccdf (distr \ borel X)" "svl differentiable at x" "svl differentiable at (x+t)" "t \ 0" "x < $\" for x t :: real proof - have "open {y. svl differentiable at y}" using ccdfX_differentiable_open_set that by simp with that obtain e0 where e0_pos: "e0 > 0" and ball_e0: "ball x e0 \ {y. svl differentiable at y}" using open_contains_ball by blast define e where "e \ if $\ < \ then min e0 (real_of_ereal $\ - x) else e0" have "e > 0 \ ball x e \ {y. svl y \ 0 \ svl differentiable at y}" proof (cases \$\ < \\) case True hence "e > 0" proof - have "real_of_ereal $\ - x > 0" using not_inftyI True that by force with e0_pos show ?thesis unfolding e_def using True by simp qed moreover have "ball x e \ {y. svl y \ 0}" proof - have "e \ real_of_ereal $\ - x" unfolding e_def using True by simp hence "ball x e \ {..< real_of_ereal $\}" unfolding ball_def dist_real_def by force thus ?thesis using that ccdfX_0_equiv using True ereal_less_real_iff psi_nonneg by auto qed moreover have "ball x e \ {y. svl differentiable at y}" proof - have "e \ e0" unfolding e_def using True by simp hence "ball x e \ ball x e0" by force with ball_e0 show ?thesis by simp qed ultimately show ?thesis by blast next case False hence "ball x e \ {y. svl y \ 0}" using ccdfX_0_equiv that by simp with False show ?thesis unfolding e_def using e0_pos ball_e0 by force qed hence e_pos: "e > 0" and ball_e: "ball x e \ {y. svl y \ 0 \ svl differentiable at y}" by auto hence ball_svl: "\y. y \ ball x e \ svl y \ 0 \ svl field_differentiable at y" using differentiable_eq_field_differentiable_real by blast hence "\y. y \ ball x e \ $p_{t&y} = svl (y+t) / svl y" unfolding survive_def using that ccdfX_0_equiv by (rewrite ccdfTx_ccdfX, simp_all) force moreover from ball_svl have "((\y. svl (y+t) / svl y) has_real_derivative ((deriv svl (x+t) * svl x - svl (x+t) * deriv svl x) / (svl x)\<^sup>2)) (at x)" apply (rewrite power2_eq_square) apply (rule DERIV_divide) using DERIV_deriv_iff_real_differentiable DERIV_shift that apply blast using that DERIV_deriv_iff_real_differentiable apply simp by (simp add: e_pos) ultimately show ?thesis using e_pos apply (rewrite has_field_derivative_cong_eventually[where g="\y. svl (y+t) / svl y"], simp_all) by (smt (verit) dist_commute eventually_at) qed lemma p_has_real_derivative_x_p_mu: "((\y. $p_{t&y}) has_real_derivative $p_{t&x} * ($\_x - $\_(x+t))) (at x)" if "ccdf (distr \ borel X) differentiable at x" "ccdf (distr \ borel X) differentiable at (x+t)" "t \ 0" "x < $\" for x t :: real proof (cases \x+t < $\\) case True let ?svl = "ccdf (distr \ borel X)" have [simp]: "?svl x \ 0" using that ccdfX_0_equiv by (smt (verit) le_ereal_le linorder_not_le) have [simp]: "?svl field_differentiable at (x+t)" using that differentiable_eq_field_differentiable_real by simp have "((\y. $p_{t&y}) has_real_derivative ((deriv ?svl (x+t) * ?svl x - ?svl (x+t) * deriv ?svl x) / (?svl x)\<^sup>2)) (at x)" using p_has_real_derivative_x_ccdfX that by simp moreover have "(deriv ?svl (x+t) * ?svl x - ?svl (x+t) * deriv ?svl x) / (?svl x)\<^sup>2 = $p_{t&x} * ($\_x - $\_(x+t))" (is "?LHS = ?RHS") proof - have "deriv ?svl (x+t) = deriv (\y. ?svl (y+t)) x" using that by (metis DERIV_deriv_iff_real_differentiable DERIV_imp_deriv DERIV_shift) hence "?LHS = (deriv (\y. ?svl (y+t)) x * ?svl x - ?svl (x+t) * deriv ?svl x) / (?svl x)\<^sup>2" by simp also have "\ = (deriv (\y. ?svl (y+t)) x - ?svl (x+t) * deriv ?svl x / ?svl x) / ?svl x" by (simp add: add_divide_eq_if_simps(4) power2_eq_square) also have "\ = (- ?svl (x+t) * $\_(x+t) + ?svl (x+t) * $\_x) / ?svl x" proof - have "deriv (\y. ?svl (y+t)) x = - ?svl (x+t) * $\_(x+t)" apply (rewrite add.commute, rewrite deriv_shift[THEN sym], rewrite add.commute, simp) using add.commute that by (metis MM_PS.deriv_ccdf_hazard_rate X_RV force_mortal_def) moreover have "- ?svl (x+t) * deriv ?svl x / ?svl x = ?svl (x+t) * $\_x" using that by (simp add: MM_PS.deriv_ccdf_hazard_rate force_mortal_def) ultimately show ?thesis by simp qed also have "\ = ?svl (x+t) * ($\_x - $\_(x+t)) / ?svl x" by (simp add: mult_diff_mult) also have "\ = ?RHS" unfolding survive_def using ccdfTx_ccdfX that by simp ultimately show ?thesis by simp qed ultimately show ?thesis by simp next case False hence ptx_0: "$p_{t&x} = 0" using p_0_equiv that by simp moreover have "((\y. $p_{t&y}) has_real_derivative 0) (at x)" proof - have "\y. x < y \ y < $\ \ $p_{t&y} = 0" using False p_0_equiv that by (smt (verit, ccfv_SIG) ereal_less_le linorder_not_le) hence "\\<^sub>F x in at_right x. $p_{t&x} = 0" apply (rewrite eventually_at_right_field) using that by (meson ereal_dense2 ereal_le_less less_eq_ereal_def less_ereal.simps) hence "((\y. $p_{t&y}) has_real_derivative 0) (at_right x)" using ptx_0 by (rewrite has_field_derivative_cong_eventually[where g="\_. 0"]; simp) thus ?thesis using vector_derivative_unique_within p_has_real_derivative_x_ccdfX that by (metis has_field_derivative_at_within has_real_derivative_iff_has_vector_derivative trivial_limit_at_right_real) qed ultimately show ?thesis by simp qed corollary deriv_x_p_mu: "deriv (\y. $p_{t&y}) x = $p_{t&x} * ($\_x - $\_(x+t))" if "ccdf (distr \ borel X) differentiable at x" "ccdf (distr \ borel X) differentiable at (x+t)" "t \ 0" "x < $\" for x t :: real using that p_has_real_derivative_x_p_mu DERIV_imp_deriv by blast lemma e_has_derivative_mu_e: "((\x. $e`\_x) has_real_derivative ($\_x * $e`\_x - 1)) (at x)" if "\x. x\{a<.. set_integrable lborel {x..} (ccdf (distr \ borel X))" "ccdf (distr \ borel X) differentiable at x" "x\{a<.. $\" for a b x :: real proof - let ?svl = "ccdf (distr \ borel X)" have x_lt_psi[simp]: "x < $\" using that ereal_le_less by simp hence svlx_neq0[simp]: "?svl x \ 0" by (simp add: ccdfX_0_equiv linorder_not_le) define d ::real where "d \ min (b-x) (x-a)" have d_pos: "d > 0" unfolding d_def using that ereal_less_real_iff by force have e_svl: "\y. y < $\ \ $e`\_y = (LBINT t:{0..}. ?svl (y+t)) / ?svl y" apply (rewrite e_LBINT_p, simp) apply (rewrite set_integral_divide_zero[THEN sym]) apply (rule set_lebesgue_integral_cong, simp_all) unfolding survive_def using ccdfTx_ccdfX by force have "((\y. LBINT t:{0..}. ?svl (y+t)) has_real_derivative (- ?svl x)) (at x)" proof - { fix y assume "dist y x < d" hence y_ab: "y \ {a<..) ?svl" by (rewrite set_integrable_discrete_difference[where X="{y}"]; simp) force moreover have "\u. ((\u. u-y) has_real_derivative (1-0)) (at u)" by (rule derivative_intros) auto moreover have "\u. y < u \ isCont (\t. ?svl (y+t)) (u-y)" apply (rewrite add.commute, rewrite isCont_shift, simp) using ccdfX_continuous continuous_on_eq_continuous_within by blast moreover have "((ereal \ (\u. u-y) \ real_of_ereal) \ 0) (at_right (ereal y))" by (smt (verit, ccfv_SIG) LIM_zero Lim_cong_within ereal_tendsto_simps1(2) ereal_tendsto_simps2(1) tendsto_ident_at zero_ereal_def) moreover have "((ereal \ (\u. u-y) \ real_of_ereal) \ \) (at_left \)" proof - have "\r u. r+y+1 \ u \ r < u-y" by auto hence "\r. \\<^sub>F u in at_top. r < u - y" by (rule eventually_at_top_linorderI) thus ?thesis by (rewrite ereal_tendsto_simps, rewrite tendsto_PInfty) simp qed ultimately have "(LBINT t=0..\. ?svl (y+t)) = (LBINT u=y..\. ?svl (y+(u-y)) * 1)" using distrX_RD.ccdf_nonneg by (intro interval_integral_substitution_nonneg(2); simp) moreover have "(LBINT t:{0..}. ?svl (y+t)) = (LBINT t=0..\. ?svl (y+t))" unfolding interval_lebesgue_integral_def einterval_def apply simp by (rule set_integral_discrete_difference[where X="{0}"]; force) moreover have "(LBINT u=y..\. ?svl (y+(u-y))*1) = (LBINT u:{y..}. ?svl u)" unfolding interval_lebesgue_integral_def einterval_def apply simp by (rule set_integral_discrete_difference[where X="{y}"]; force) ultimately have "(LBINT t:{0..}. ?svl (y+t)) = (LBINT u:{y..}. ?svl u)" by simp } hence "\\<^sub>F y in nhds x. LBINT t:{0..}. ?svl (y+t) = LBINT u:{y..}. ?svl u" using d_pos by (rewrite eventually_nhds_metric) auto moreover have "((\y. LBINT u:{y..}. ?svl u) has_real_derivative (- ?svl x)) (at x)" proof - have "((\y. integral {y..b} ?svl) has_real_derivative (- ?svl x)) (at x within {a..b})" using that apply (intro integral_has_real_derivative'; simp) using ccdfX_continuous continuous_on_subset by blast hence "((\y. integral {y..b} ?svl) has_real_derivative (- ?svl x)) (at x)" using that apply (rewrite at_within_open[where S="{a<..\<^sub>F y in nhds x. LBINT u:{y..b}. ?svl u = integral {y..b} ?svl" apply (rewrite eventually_nhds_metric) using d_pos by (metis ccdfX_integrable_Icc set_borel_integral_eq_integral(2)) ultimately have "((\y. LBINT u:{y..b}. ?svl u) has_real_derivative (- ?svl x)) (at x)" by (rewrite DERIV_cong_ev; simp) hence "((\y. (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)) has_real_derivative (- ?svl x)) (at x)" by (rewrite to "- ?svl x + 0" add_0_right[THEN sym], rule DERIV_add; simp) moreover have "\\<^sub>F y in nhds x. LBINT u:{y..}. ?svl u = (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)" proof - { fix y assume "dist y x < d" hence y_ab: "y \ {a<.. {b<..} = {}" "{y..} = {y..b} \ {b<..}" using y_ab by force+ ultimately have "LBINT u:{y..}. ?svl u = (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)" using set_integral_Un by simp } thus ?thesis using d_pos by (rewrite eventually_nhds_metric) blast qed ultimately show ?thesis by (rewrite has_field_derivative_cong_ev; simp) qed ultimately show ?thesis by (rewrite DERIV_cong_ev; simp) qed moreover have "(?svl has_real_derivative (deriv ?svl x)) (at x)" using that DERIV_deriv_iff_real_differentiable by blast ultimately have "((\y. (LBINT t:{0..}. ?svl (y+t)) / ?svl y) has_real_derivative ((- ?svl x) * ?svl x - (LBINT t:{0..}. ?svl (x+t)) * deriv ?svl x) / (?svl x * ?svl x)) (at x)" by (rule DERIV_divide) simp moreover have "eventually (\y. (LBINT t:{0..}. ?svl (y+t)) / ?svl y = $e`\_y) (at x)" proof - { fix y assume "dist y x < d" "y \ x" hence "y < $\" unfolding dist_real_def d_def using that ereal_le_less by fastforce hence "$e`\_y = (LBINT t:{0..}. ?svl (y+t)) / ?svl y" by (rule e_svl) } thus ?thesis apply (rewrite eventually_at_filter, rewrite eventually_nhds_metric) using d_pos that by metis qed ultimately have "((\y. $e`\_y) has_real_derivative ((- ?svl x) * ?svl x - (LBINT t:{0..}. ?svl (x+t)) * deriv ?svl x) / (?svl x * ?svl x)) (at x)" using e_svl by (rewrite has_field_derivative_cong_eventually[THEN sym]; simp) moreover have "((- ?svl x) * ?svl x - (LBINT t:{0..}. ?svl (x+t)) * deriv ?svl x) / (?svl x * ?svl x) = $\_x * $e`\_x - 1" (is "?LHS = ?RHS") proof - have "?LHS = -1 + (LBINT t:{0..}. ?svl (x+t)) / ?svl x * (- deriv ?svl x / ?svl x)" by simp (smt (verit) svlx_neq0 add_divide_distrib divide_eq_minus_1_iff mult_minus_left real_divide_square_eq) also have "\ = -1 + $e`\_x * $\_x" using e_svl mu_deriv_ccdf that by force also have "\ = ?RHS" by simp finally show ?thesis . qed ultimately show ?thesis by simp qed corollary e_has_derivative_mu_e': "((\x. $e`\_x) has_real_derivative ($\_x * $e`\_x - 1)) (at x)" if "\x. x\{a<.. ccdf (distr \ borel X) integrable_on {x..}" "ccdf (distr \ borel X) differentiable at x" "x\{a<.. $\" for a b x :: real using that apply (intro e_has_derivative_mu_e[where a=a], simp_all) using distrX_RD.ccdf_nonneg by (rewrite integrable_on_iff_set_integrable_nonneg; simp) subsubsection \Properties of Curtate Life Expectation\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin lemma isCont_p_nat: "isCont (\t. $p_{t&x}) (k + (1::real))" for k::nat proof - fix k::nat have "continuous_on {0<..} (\t. $p_{t&x})" unfolding survive_def using ccdfTx_continuous_on_nonneg continuous_on_subset Ioi_le_Ico x_lt_psi by blast hence "\t\{0<..}. isCont (\t. $p_{t&x}) t" using continuous_on_eq_continuous_at open_greaterThan by blast thus "isCont (\t. $p_{t&x}) (real k+1)" by force qed lemma curt_e_sum_p_smooth: "$e_x = (\k. $p_{k+1&x})" if "summable (\k. $p_{k+1&x})" using curt_e_sum_p isCont_p_nat that by simp lemma curt_e_rec_smooth: "$e_x = $p_x * (1 + $e_(x+1))" if "summable (\k. $p_{k+1&x})" "x+1 < $\" using curt_e_rec isCont_p_nat that by simp lemma curt_e_sum_p_finite_smooth: "$e_x = (\k $\" for n::nat using curt_e_sum_p_finite isCont_p_nat that by simp lemma temp_curt_e_sum_p_smooth: "$e_{x:n} = (\k" "n \ 0" for n::nat using temp_curt_e_rec isCont_p_nat that by simp end end subsection \Finite Survival Function\ locale finite_survival_function = survival_model + assumes psi_finite[simp]: "$\ < \" begin definition ult_age :: nat ("$\") where "$\ \ LEAST x::nat. ccdf (distr \ borel X) x = 0" \ \the conventional notation for ultimate age\ lemma ccdfX_ceil_psi_0: "ccdf (distr \ borel X) \real_of_ereal $\\ = 0" proof - have "real_of_ereal $\ \ \real_of_ereal $\\" by simp thus ?thesis using ccdfX_0_equiv psi_finite ccdfX_psi_0 le_ereal_le by presburger qed lemma ccdfX_omega_0: "ccdf (distr \ borel X) $\ = 0" unfolding ult_age_def proof (rule LeastI_ex) have "\real_of_ereal $\\ \ 0" using psi_nonneg real_of_ereal_pos by fastforce thus "\x::nat. ccdf (distr \ borel X) (real x) = 0" using ccdfX_ceil_psi_0 by (metis of_int_of_nat_eq zero_le_imp_eq_int) qed corollary psi_le_omega: "$\ \ $\" using ccdfX_0_equiv ccdfX_omega_0 by blast corollary omega_pos: "$\ > 0" using psi_le_omega order.strict_iff_not by fastforce lemma omega_ceil_psi: "$\ = \real_of_ereal $\\" proof (rule antisym) let ?cpsi = "\real_of_ereal $\\" have \: "?cpsi \ 0" using psi_nonneg real_of_ereal_pos by fastforce moreover have "ccdf (distr \ borel X) ?cpsi = 0" by (rule ccdfX_ceil_psi_0) ultimately have "$\ \ nat ?cpsi" unfolding ult_age_def using wellorder_Least_lemma of_nat_nat by smt thus "int $\ \ ?cpsi" using le_nat_iff \ by blast next show "\real_of_ereal $\\ \ int $\" using psi_le_omega by (simp add: ceiling_le_iff real_of_ereal_ord_simps(2)) qed lemma ccdfX_0_equiv_nat: "ccdf (distr \ borel X) x = 0 \ x \ $\" for x::nat proof assume "ccdf (distr \ borel X) (real x) = 0" thus "x \ $\" unfolding ult_age_def using wellorder_Least_lemma by fastforce next assume "x \ $\" hence "ereal (real x) \ $\" using psi_le_omega le_ereal_le of_nat_mono by blast thus "ccdf (distr \ borel X) (real x) = 0" using ccdfX_0_equiv by simp qed lemma psi_le_iff_omega_le: "$\ \ x \ $\ \ x" for x::nat using ccdfX_0_equiv ccdfX_0_equiv_nat by presburger context fixes x::nat assumes x_lt_omega[simp]: "x < $\" begin lemma x_lt_psi[simp]: "x < $\" using x_lt_omega psi_le_iff_omega_le by (meson linorder_not_less) lemma p_0_1_nat: "$p_{0&x} = 1" using p_0_1 by simp lemma p_0_equiv_nat: "$p_{t&x} = 0 \ x+t \ $\" for t::nat using psi_le_iff_omega_le p_0_equiv by (metis of_nat_add x_lt_psi) lemma q_0_0_nat: "$q_{0&x} = 0" using p_q_1 p_0_1_nat by (smt (verit) x_lt_psi) lemma q_1_equiv_nat: "$q_{t&x} = 1 \ x+t \ $\" for t::nat using p_q_1 p_0_equiv_nat by (smt (verit) x_lt_psi) lemma q_defer_old_0_nat: "$q_{f\t&x} = 0" if "$\ \ x+f" for f t :: nat using q_defer_old_0 psi_le_iff_omega_le that by (metis of_nat_0_le_iff of_nat_add x_lt_psi) lemma curt_e_sum_P_finite_nat: "$e_x = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" if "x+n \ $\" for n::nat apply (rule curt_e_sum_P_finite, simp) using that psi_le_iff_omega_le by (smt (verit) le_ereal_less of_nat_add) lemma curt_e_sum_p_finite_nat: "$e_x = (\kk::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" "x+n \ $\" for n::nat apply (rule curt_e_sum_p_finite, simp_all add: that) using that psi_le_iff_omega_le by (smt (verit) le_ereal_less of_nat_add) end lemma q_omega_1: "$q_($\-1) = 1" using q_1_equiv_nat by (metis diff_less dual_order.refl le_diff_conv of_nat_1 omega_pos zero_less_one) end end diff --git a/thys/Affine_Arithmetic/Intersection.thy b/thys/Affine_Arithmetic/Intersection.thy --- a/thys/Affine_Arithmetic/Intersection.thy +++ b/thys/Affine_Arithmetic/Intersection.thy @@ -1,2814 +1,2814 @@ section \Intersection\ theory Intersection imports "HOL-Library.Monad_Syntax" Polygon Counterclockwise_2D_Arbitrary Affine_Form begin text \\label{sec:intersection}\ subsection \Polygons and @{term ccw}, @{term lex}, @{term psi}, @{term coll}\ lemma polychain_of_ccw_conjunction: assumes sorted: "ccw'.sortedP 0 Ps" assumes z: "z \ set (polychain_of Pc Ps)" shows "list_all (\(xi, xj). ccw xi xj (fst z) \ ccw xi xj (snd z)) (polychain_of Pc Ps)" using assms proof (induction Ps arbitrary: Pc z rule: list.induct) case (Cons P Ps) { assume "set Ps = {}" hence ?case using Cons by simp } moreover { assume "set Ps \ {}" hence "Ps \ []" by simp { fix a assume "a \ set Ps" hence "ccw' 0 P a" using Cons.prems by (auto elim!: linorder_list0.sortedP_Cons) } note ccw' = this have sorted': "linorder_list0.sortedP (ccw' 0) Ps" using Cons.prems by (auto elim!: linorder_list0.sortedP_Cons) from in_set_polychain_of_imp_sum_list[OF Cons(3)] obtain d where d: "z = (Pc + sum_list (take d (P # Ps)), Pc + sum_list (take (Suc d) (P # Ps)))" . from Cons(3) have disj: "z = (Pc, Pc + P) \ z \ set (polychain_of (Pc + P) Ps)" by auto let ?th = "\(xi, xj). ccw xi xj Pc \ ccw xi xj (Pc + P)" have la: "list_all ?th (polychain_of (Pc + P) Ps)" proof (rule list_allI) fix x assume x: "x \ set (polychain_of (Pc + P) Ps)" from in_set_polychain_of_imp_sum_list[OF this] obtain e where e: "x = (Pc + P + sum_list (take e Ps), Pc + P + sum_list (take (Suc e) Ps))" by auto { assume "e \ length Ps" hence "?th x" by (auto simp: e) } moreover { assume "e < length Ps" have 0: "\e. e < length Ps \ ccw' 0 P (Ps ! e)" by (rule ccw') simp have 2: "0 < e \ ccw' 0 (P + sum_list (take e Ps)) (Ps ! e)" using \e < length Ps\ by (auto intro!: ccw'.add1 0 ccw'.sum2 sorted' ccw'.sorted_nth_less simp: sum_list_sum_nth) have "ccw Pc (Pc + P + sum_list (take e Ps)) (Pc + P + sum_list (take (Suc e) Ps))" by (cases "e = 0") (auto simp add: ccw_translate_origin take_Suc_eq add.assoc[symmetric] 0 2 intro!: ccw'_imp_ccw intro: cyclic) hence "ccw (Pc + P + sum_list (take e Ps)) (Pc + P + sum_list (take (Suc e) Ps)) Pc" by (rule cyclic) moreover have "0 < e \ ccw 0 (Ps ! e) (- sum_list (take e Ps))" using \e < length Ps\ by (auto simp add: take_Suc_eq add.assoc[symmetric] sum_list_sum_nth intro!: ccw'_imp_ccw ccw'.sum2 sorted' ccw'.sorted_nth_less) hence "ccw (Pc + P + sum_list (take e Ps)) (Pc + P + sum_list (take (Suc e) Ps)) (Pc + P)" by (cases "e = 0") (simp_all add: ccw_translate_origin take_Suc_eq) ultimately have "?th x" by (auto simp add: e) } ultimately show "?th x" by arith qed from disj have ?case proof assume z: "z \ set (polychain_of (Pc + P) Ps)" have "ccw 0 P (sum_list (take d (P # Ps)))" proof (cases d) case (Suc e) note e = this show ?thesis proof (cases e) case (Suc f) have "ccw 0 P (P + sum_list (take (Suc f) Ps))" using z by (force simp add: sum_list_sum_nth intro!: ccw'.sum intro: ccw' ccw'_imp_ccw) thus ?thesis by (simp add: e Suc) qed (simp add: e) qed simp hence "ccw Pc (Pc + P) (fst z)" by (simp add: d ccw_translate_origin) moreover from z have "ccw 0 P (P + sum_list (take d Ps))" by (cases d, force) (force simp add: sum_list_sum_nth intro!: ccw'_imp_ccw ccw'.sum intro: ccw')+ hence "ccw Pc (Pc + P) (snd z)" by (simp add: d ccw_translate_origin) moreover from z Cons.prems have "list_all (\(xi, xj). ccw xi xj (fst z) \ ccw xi xj (snd z)) (polychain_of (Pc + P) Ps)" by (intro Cons.IH) (auto elim!: linorder_list0.sortedP_Cons) ultimately show ?thesis by simp qed (simp add: la) } ultimately show ?case by blast qed simp lemma lex_polychain_of_center: "d \ set (polychain_of x0 xs) \ list_all (\x. lex x 0) xs \ lex (snd d) x0" proof (induction xs arbitrary: x0) case (Cons x xs) thus ?case by (auto simp add: lex_def lex_translate_origin dest!: Cons.IH) qed (auto simp: lex_translate_origin) lemma lex_polychain_of_last: "(c, d) \ set (polychain_of x0 xs) \ list_all (\x. lex x 0) xs \ lex (snd (last (polychain_of x0 xs))) d" proof (induction xs arbitrary: x0 c d) case (Cons x xs) let ?c1 = "c = x0 \ d = x0 + x" let ?c2 = "(c, d) \ set (polychain_of (x0 + x) xs)" from Cons(2) have "?c1 \ ?c2" by auto thus ?case proof assume ?c1 with Cons.prems show ?thesis by (auto intro!: lex_polychain_of_center) next assume ?c2 from Cons.IH[OF this] Cons.prems show ?thesis by auto qed qed (auto simp: lex_translate_origin) lemma distinct_fst_polychain_of: assumes "list_all (\x. x \ 0) xs" assumes "list_all (\x. lex x 0) xs" shows "distinct (map fst (polychain_of x0 xs))" using assms proof (induction xs arbitrary: x0) case Nil thus ?case by simp next case (Cons x xs) hence "\d. list_all (\x. lex x 0) (x # take d xs)" by (auto simp: list_all_iff dest!: in_set_takeD) from sum_list_nlex_eq_zero_iff[OF this] Cons.prems show ?case by (cases "xs = []") (auto intro!: Cons.IH elim!: in_set_polychain_of_imp_sum_list) qed lemma distinct_snd_polychain_of: assumes "list_all (\x. x \ 0) xs" assumes "list_all (\x. lex x 0) xs" shows "distinct (map snd (polychain_of x0 xs))" using assms proof (induction xs arbitrary: x0) case Nil thus ?case by simp next case (Cons x xs) have contra: "\d. xs \ [] \ list_all (\x. x \ 0) xs \ list_all ((=) 0) (take (Suc d) xs) \ False" by (auto simp: neq_Nil_conv) from Cons have "\d. list_all (\x. lex x 0) (take (Suc d) xs)" by (auto simp: list_all_iff dest!: in_set_takeD) from sum_list_nlex_eq_zero_iff[OF this] Cons.prems contra show ?case by (cases "xs = []") (auto intro!: Cons.IH elim!: in_set_polychain_of_imp_sum_list dest!: contra) qed subsection \Orient all entries\ lift_definition nlex_pdevs::"point pdevs \ point pdevs" is "\x n. if lex 0 (x n) then - x n else x n" by simp lemma pdevs_apply_nlex_pdevs[simp]: "pdevs_apply (nlex_pdevs x) n = (if lex 0 (pdevs_apply x n) then - pdevs_apply x n else pdevs_apply x n)" by transfer simp lemma nlex_pdevs_zero_pdevs[simp]: "nlex_pdevs zero_pdevs = zero_pdevs" by (auto intro!: pdevs_eqI) lemma pdevs_domain_nlex_pdevs[simp]: "pdevs_domain (nlex_pdevs x) = pdevs_domain x" by (auto simp: pdevs_domain_def) lemma degree_nlex_pdevs[simp]: "degree (nlex_pdevs x) = degree x" by (rule degree_cong) auto lemma pdevs_val_nlex_pdevs: assumes "e \ UNIV \ I" "uminus ` I = I" obtains e' where "e' \ UNIV \ I" "pdevs_val e x = pdevs_val e' (nlex_pdevs x)" using assms by (atomize_elim, intro exI[where x="\n. if lex 0 (pdevs_apply x n) then - e n else e n"]) (force simp: pdevs_val_pdevs_domain intro!: sum.cong) lemma pdevs_val_nlex_pdevs2: assumes "e \ UNIV \ I" "uminus ` I = I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (nlex_pdevs x) = pdevs_val e' x" using assms by (atomize_elim, intro exI[where x="\n. (if lex 0 (pdevs_apply x n) then - e n else e n)"]) (force simp: pdevs_val_pdevs_domain intro!: sum.cong) lemma pdevs_val_selsort_ccw: assumes "distinct xs" assumes "e \ UNIV \ I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (ccw.selsort 0 xs))" proof - have "set xs = set (ccw.selsort 0 xs)" "distinct xs" "distinct (ccw.selsort 0 xs)" by (simp_all add: assms) from this assms(2) obtain e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (ccw.selsort 0 xs))" by (rule pdevs_val_permute) thus thesis .. qed lemma pdevs_val_selsort_ccw2: assumes "distinct xs" assumes "e \ UNIV \ I" obtains e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list (ccw.selsort 0 xs)) = pdevs_val e' (pdevs_of_list xs)" proof - have "set (ccw.selsort 0 xs) = set xs" "distinct (ccw.selsort 0 xs)" "distinct xs" by (simp_all add: assms) from this assms(2) obtain e' where "e' \ UNIV \ I" "pdevs_val e (pdevs_of_list (ccw.selsort 0 xs)) = pdevs_val e' (pdevs_of_list xs)" by (rule pdevs_val_permute) thus thesis .. qed lemma lex_nlex_pdevs: "lex (pdevs_apply (nlex_pdevs x) i) 0" by (auto simp: lex_def algebra_simps prod_eq_iff) subsection \Lowest Vertex\ definition lowest_vertex::"'a::ordered_euclidean_space aform \ 'a" where "lowest_vertex X = fst X - sum_list (map snd (list_of_pdevs (snd X)))" lemma snd_abs: "snd (abs x) = abs (snd x)" by (metis abs_prod_def snd_conv) lemma lowest_vertex: fixes X Y::"(real*real) aform" assumes "e \ UNIV \ {-1 .. 1}" assumes "\i. snd (pdevs_apply (snd X) i) \ 0" assumes "\i. abs (snd (pdevs_apply (snd Y) i)) = abs (snd (pdevs_apply (snd X) i))" assumes "degree_aform Y = degree_aform X" assumes "fst Y = fst X" shows "snd (lowest_vertex X) \ snd (aform_val e Y)" proof - from abs_pdevs_val_le_tdev[OF assms(1), of "snd Y"] have "snd \pdevs_val e (snd Y)\ \ (\isnd (pdevs_apply (snd X) i)\)" unfolding lowest_vertex_def by (auto simp: aform_val_def tdev_def less_eq_prod_def snd_sum snd_abs assms) also have "\ = (\i \ snd (sum_list (map snd (list_of_pdevs (snd X))))" by (simp add: sum_list_list_of_pdevs dense_list_of_pdevs_def sum_list_distinct_conv_sum_set snd_sum atLeast0LessThan) finally show ?thesis by (auto simp: aform_val_def lowest_vertex_def minus_le_iff snd_abs abs_real_def assms split: if_split_asm) qed lemma sum_list_nonposI: fixes xs::"'a::ordered_comm_monoid_add list" shows "list_all (\x. x \ 0) xs \ sum_list xs \ 0" by (induct xs) (auto simp: intro!: add_nonpos_nonpos) lemma center_le_lowest: "fst (fst X) \ fst (lowest_vertex (fst X, nlex_pdevs (snd X)))" by (auto simp: lowest_vertex_def fst_sum_list intro!: sum_list_nonposI) (auto simp: lex_def list_all_iff list_of_pdevs_def dest!: in_set_butlastD split: if_split_asm) lemma lowest_vertex_eq_center_iff: "lowest_vertex (x0, nlex_pdevs (snd X)) = x0 \ snd X = zero_pdevs" proof assume "lowest_vertex (x0, nlex_pdevs (snd X)) = x0" then have "sum_list (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = 0" by (simp add: lowest_vertex_def) moreover have "list_all (\x. Counterclockwise_2D_Arbitrary.lex x 0) (map snd (list_of_pdevs (nlex_pdevs (snd X))))" by (auto simp add: list_all_iff list_of_pdevs_def) ultimately have "\x\set (list_of_pdevs (nlex_pdevs (snd X))). snd x = 0" by (simp add: sum_list_nlex_eq_zero_iff list_all_iff) then have "pdevs_apply (snd X) i = pdevs_apply zero_pdevs i" for i by (simp add: list_of_pdevs_def split: if_splits) then show "snd X = zero_pdevs" by (rule pdevs_eqI) qed (simp add: lowest_vertex_def) subsection \Collinear Generators\ lemma scaleR_le_self_cancel: fixes c::"'a::ordered_real_vector" shows "a *\<^sub>R c \ c \ (1 < a \ c \ 0 \ a < 1 \ 0 \ c \ a = 1)" using scaleR_le_0_iff[of "a - 1" c] by (simp add: algebra_simps) lemma pdevs_val_coll: assumes coll: "list_all (coll 0 x) xs" assumes nlex: "list_all (\x. lex x 0) xs" assumes "x \ 0" assumes "f \ UNIV \ {-1 .. 1}" obtains e where "e \ {-1 .. 1}" "pdevs_val f (pdevs_of_list xs) = e *\<^sub>R (sum_list xs)" proof cases assume "sum_list xs = 0" have "pdevs_of_list xs = zero_pdevs" by (auto intro!: pdevs_eqI sum_list_nlex_eq_zeroI[OF nlex \sum_list xs = 0\] simp: pdevs_apply_pdevs_of_list list_all_iff dest!: nth_mem) hence "0 \ {-1 .. 1::real}" "pdevs_val f (pdevs_of_list xs) = 0 *\<^sub>R sum_list xs" by simp_all thus ?thesis .. next assume "sum_list xs \ 0" have "sum_list (map abs xs) \ 0" by (auto intro!: sum_list_nonneg) hence [simp]: "\sum_list (map abs xs) \ 0" by (metis \sum_list xs \ 0\ abs_le_zero_iff antisym_conv sum_list_abs) have collist: "list_all (coll 0 (sum_list xs)) xs" proof (rule list_allI) fix y assume "y \ set xs" hence "coll 0 x y" using coll by (auto simp: list_all_iff) also have "coll 0 x (sum_list xs)" using coll by (auto simp: list_all_iff intro!: coll_sum_list) finally (coll_trans) show "coll 0 (sum_list xs) y" by (simp add: coll_commute \x \ 0\) qed { fix i assume "i < length xs" hence "\r. xs ! i = r *\<^sub>R (sum_list xs)" by (metis (mono_tags) coll_scale nth_mem \sum_list xs \ 0\ list_all_iff collist) } then obtain r where r: "\i. i < length xs \ (xs ! i) = r i *\<^sub>R (sum_list xs)" by metis let ?coll = "pdevs_of_list xs" have "pdevs_val f (pdevs_of_list xs) = (\iR xs ! i)" unfolding pdevs_val_sum by (simp add: pdevs_apply_pdevs_of_list less_degree_pdevs_of_list_imp_less_length) also have "\ = (\iR (sum_list xs))" by (simp add: r less_degree_pdevs_of_list_imp_less_length) also have "\ = (\iR (sum_list xs)" by (simp add: algebra_simps scaleR_sum_left) finally have eq: "pdevs_val f ?coll = (\iR (sum_list xs)" (is "_ = ?e *\<^sub>R _") . have "abs (pdevs_val f ?coll) \ tdev ?coll" using assms(4) by (intro abs_pdevs_val_le_tdev) (auto simp: Pi_iff less_imp_le) also have "\ = sum_list (map abs xs)" using assms by simp also note eq finally have less: "\?e\ *\<^sub>R abs (sum_list xs) \ sum_list (map abs xs)" by (simp add: abs_scaleR) also have "\sum_list xs\ = sum_list (map abs xs)" using coll \x \ 0\ nlex by (rule abs_sum_list_coll) finally have "?e \ {-1 .. 1}" by (auto simp add: less_le scaleR_le_self_cancel) thus ?thesis using eq .. qed lemma scaleR_eq_self_cancel: fixes x::"'a::real_vector" shows "a *\<^sub>R x = x \ a = 1 \ x = 0" using scaleR_cancel_right[of a x 1] by simp lemma abs_pdevs_val_less_tdev: assumes "e \ UNIV \ {-1 <..< 1}" "degree x > 0" shows "\pdevs_val e x\ < tdev x" proof - have bnds: "\i. \e i\ < 1" "\i. \e i\ \ 1" using assms by (auto simp: Pi_iff abs_less_iff order.strict_implies_order) moreover let ?w = "degree x - 1" have witness: "\e ?w\ *\<^sub>R \pdevs_apply x ?w\ < \pdevs_apply x ?w\" using degree_least_nonzero[of x] assms bnds by (intro neq_le_trans) (auto simp: scaleR_eq_self_cancel Pi_iff intro!: scaleR_left_le_one_le neq_le_trans intro: abs_leI less_imp_le dest!: order.strict_implies_not_eq) ultimately show ?thesis using assms witness bnds by (auto simp: pdevs_val_sum tdev_def abs_scaleR intro!: le_less_trans[OF sum_abs] sum_strict_mono_ex1 scaleR_left_le_one_le) qed lemma pdevs_val_coll_strict: assumes coll: "list_all (coll 0 x) xs" assumes nlex: "list_all (\x. lex x 0) xs" assumes "x \ 0" assumes "f \ UNIV \ {-1 <..< 1}" obtains e where "e \ {-1 <..< 1}" "pdevs_val f (pdevs_of_list xs) = e *\<^sub>R (sum_list xs)" proof cases assume "sum_list xs = 0" have "pdevs_of_list xs = zero_pdevs" by (auto intro!: pdevs_eqI sum_list_nlex_eq_zeroI[OF nlex \sum_list xs = 0\] simp: pdevs_apply_pdevs_of_list list_all_iff dest!: nth_mem) hence "0 \ {-1 <..< 1::real}" "pdevs_val f (pdevs_of_list xs) = 0 *\<^sub>R sum_list xs" by simp_all thus ?thesis .. next assume "sum_list xs \ 0" have "sum_list (map abs xs) \ 0" by (auto intro!: sum_list_nonneg) hence [simp]: "\sum_list (map abs xs) \ 0" by (metis \sum_list xs \ 0\ abs_le_zero_iff antisym_conv sum_list_abs) have "\x \ set xs. x \ 0" proof (rule ccontr) assume "\ (\x\set xs. x \ 0)" hence "\x. x \ set xs \ x = 0" by auto hence "sum_list xs = 0" by (auto simp: sum_list_eq_0_iff_nonpos list_all_iff less_eq_prod_def prod_eq_iff fst_sum_list snd_sum_list) thus False using \sum_list xs \ 0\ by simp qed then obtain i where i: "i < length xs" "xs ! i \ 0" by (metis in_set_conv_nth) hence "i < degree (pdevs_of_list xs)" by (auto intro!: degree_gt simp: pdevs_apply_pdevs_of_list) hence deg_pos: "0 < degree (pdevs_of_list xs)" by simp have collist: "list_all (coll 0 (sum_list xs)) xs" proof (rule list_allI) fix y assume "y \ set xs" hence "coll 0 x y" using coll by (auto simp: list_all_iff) also have "coll 0 x (sum_list xs)" using coll by (auto simp: list_all_iff intro!: coll_sum_list) finally (coll_trans) show "coll 0 (sum_list xs) y" by (simp add: coll_commute \x \ 0\) qed { fix i assume "i < length xs" hence "\r. xs ! i = r *\<^sub>R (sum_list xs)" by (metis (mono_tags, lifting) \sum_list xs \ 0\ coll_scale collist list_all_iff nth_mem) } then obtain r where r: "\i. i < length xs \ (xs ! i) = r i *\<^sub>R (sum_list xs)" by metis let ?coll = "pdevs_of_list xs" have "pdevs_val f (pdevs_of_list xs) = (\iR xs ! i)" unfolding pdevs_val_sum by (simp add: less_degree_pdevs_of_list_imp_less_length pdevs_apply_pdevs_of_list) also have "\ = (\iR (sum_list xs))" by (simp add: r less_degree_pdevs_of_list_imp_less_length) also have "\ = (\iR (sum_list xs)" by (simp add: algebra_simps scaleR_sum_left) finally have eq: "pdevs_val f ?coll = (\iR (sum_list xs)" (is "_ = ?e *\<^sub>R _") . have "abs (pdevs_val f ?coll) < tdev ?coll" using assms(4) by (intro abs_pdevs_val_less_tdev) (auto simp: Pi_iff less_imp_le deg_pos) also have "\ = sum_list (map abs xs)" using assms by simp also note eq finally have less: "\?e\ *\<^sub>R abs (sum_list xs) < sum_list (map abs xs)" by (simp add: abs_scaleR) also have "\sum_list xs\ = sum_list (map abs xs)" using coll \x \ 0\ nlex by (rule abs_sum_list_coll) finally have "?e \ {-1 <..< 1}" by (auto simp add: less_le scaleR_le_self_cancel) thus ?thesis using eq .. qed subsection \Independent Generators\ fun independent_pdevs::"point list \ point list" where "independent_pdevs [] = []" | "independent_pdevs (x#xs) = (let (cs, is) = List.partition (coll 0 x) xs; s = x + sum_list cs in (if s = 0 then [] else [s]) @ independent_pdevs is)" lemma in_set_independent_pdevsE: assumes "y \ set (independent_pdevs xs)" obtains x where "x\set xs" "coll 0 x y" proof atomize_elim show "\x. x \ set xs \ coll 0 x y" using assms proof (induct xs rule: independent_pdevs.induct) case 1 thus ?case by simp next case (2 z zs) let ?c1 = "y = z + sum_list (filter (coll 0 z) zs)" let ?c2 = "y \ set (independent_pdevs (filter (Not \ coll 0 z) zs))" from 2 have "?c1 \ ?c2" by (auto simp: Let_def split: if_split_asm) thus ?case proof assume ?c2 hence "y \ set (independent_pdevs (snd (partition (coll 0 z) zs)))" by simp from 2(1)[OF refl prod.collapse refl this] show ?case by auto next assume y: ?c1 show ?case unfolding y by (rule exI[where x="z"]) (auto intro!: coll_add coll_sum_list ) qed qed qed lemma in_set_independent_pdevs_nonzero: "x \ set (independent_pdevs xs) \ x \ 0" proof (induct xs rule: independent_pdevs.induct) case (2 y ys) from 2(1)[OF refl prod.collapse refl] 2(2) show ?case by (auto simp: Let_def split: if_split_asm) qed simp lemma independent_pdevs_pairwise_non_coll: assumes "x \ set (independent_pdevs xs)" assumes "y \ set (independent_pdevs xs)" assumes "\x. x \ set xs \ x \ 0" assumes "x \ y" shows "\ coll 0 x y" using assms proof (induct xs rule: independent_pdevs.induct) case 1 thus ?case by simp next case (2 z zs) from 2 have "z \ 0" by simp from 2(2) have "x \ 0" by (rule in_set_independent_pdevs_nonzero) from 2(3) have "y \ 0" by (rule in_set_independent_pdevs_nonzero) let ?c = "filter (coll 0 z) zs" let ?nc = "filter (Not \ coll 0 z) zs" { assume "x \ set (independent_pdevs ?nc)" "y \ set (independent_pdevs ?nc)" hence "\coll 0 x y" by (intro 2(1)[OF refl prod.collapse refl _ _ 2(4) 2(5)]) auto } note IH = this { fix x assume "x \ 0" "z + sum_list ?c \ 0" "coll 0 x (z + sum_list ?c)" hence "x \ set (independent_pdevs ?nc)" using sum_list_filter_coll_ex_scale[OF \z \ 0\, of "z#zs"] by (auto elim!: in_set_independent_pdevsE simp: coll_commute) (metis (no_types) \x \ 0\ coll_scale coll_scaleR) } note nc = this from 2(2,3,4,5) nc[OF \x \ 0\] nc[OF \y \ 0\] show ?case by (auto simp: Let_def IH coll_commute split: if_split_asm) qed lemma distinct_independent_pdevs[simp]: shows "distinct (independent_pdevs xs)" proof (induct xs rule: independent_pdevs.induct) case 1 thus ?case by simp next case (2 x xs) let ?is = "independent_pdevs (filter (Not \ coll 0 x) xs)" have "distinct ?is" by (rule 2) (auto intro!: 2) thus ?case proof (clarsimp simp add: Let_def) let ?s = "x + sum_list (filter (coll 0 x) xs)" assume s: "?s \ 0" "?s \ set ?is" from in_set_independent_pdevsE[OF s(2)] obtain y where y: "y \ set (filter (Not \ coll 0 x) xs)" "coll 0 y (x + sum_list (filter (coll 0 x) xs))" by auto { assume "y = 0 \ x = 0 \ sum_list (filter (coll 0 x) xs) = 0" hence False using s y by (auto simp: coll_commute) } moreover { assume "y \ 0" "x \ 0" "sum_list (filter (coll 0 x) xs) \ 0" "sum_list (filter (coll 0 x) xs) + x \ 0" have *: "coll 0 (sum_list (filter (coll 0 x) xs)) x" by (auto intro!: coll_sum_list simp: coll_commute) have "coll 0 y (sum_list (filter (coll 0 x) xs) + x)" using s y by (simp add: add.commute) hence "coll 0 y x" using * by (rule coll_add_trans) fact+ hence False using s y by (simp add: coll_commute) } ultimately show False using s y by (auto simp: add.commute) qed qed lemma in_set_independent_pdevs_invariant_nlex: "x \ set (independent_pdevs xs) \ (\x. x \ set xs \ lex x 0) \ (\x. x \ set xs \ x \ 0) \ Counterclockwise_2D_Arbitrary.lex x 0" proof (induction xs arbitrary: x rule: independent_pdevs.induct ) case 1 thus ?case by simp next case (2 y ys) from 2 have "y \ 0" by auto from 2(2) have "x \ set (independent_pdevs (filter (Not \ coll 0 y) ys)) \ x = y + sum_list (filter (coll 0 y) ys)" by (auto simp: Let_def split: if_split_asm) thus ?case proof assume "x \ set (independent_pdevs (filter (Not \ coll 0 y) ys))" from 2(1)[OF refl prod.collapse refl, simplified, OF this 2(3,4)] show ?case by simp next assume "x = y + sum_list (filter (coll 0 y) ys)" also have "lex \ 0" by (force intro: nlex_add nlex_sum simp: sum_list_sum_nth dest: nth_mem intro: 2(3)) finally show ?case . qed qed lemma pdevs_val_independent_pdevs2: assumes "e \ UNIV \ I" shows "\e'. e' \ UNIV \ I \ pdevs_val e (pdevs_of_list (independent_pdevs xs)) = pdevs_val e' (pdevs_of_list xs)" using assms proof (induct xs arbitrary: e rule: independent_pdevs.induct) case 1 thus ?case by force next case (2 x xs) let ?coll = "(filter (coll 0 x) (x#xs))" let ?ncoll = "(filter (Not o coll 0 x) (x#xs))" let ?e0 = "if sum_list ?coll = 0 then e else e \ (+) (Suc 0)" have "pdevs_val e (pdevs_of_list (independent_pdevs (x#xs))) = e 0 *\<^sub>R (sum_list ?coll) + pdevs_val ?e0 (pdevs_of_list (independent_pdevs ?ncoll))" (is "_ = ?vc + ?vnc") by (simp add: Let_def) also have e_split: "(\_. e 0) \ UNIV \ I" "?e0 \ UNIV \ I" using 2(2) by auto from 2(1)[OF refl prod.collapse refl e_split(2)] obtain e' where e': "e' \ UNIV \ I" and "?vnc = pdevs_val e' (pdevs_of_list ?ncoll)" by (auto simp add: o_def) note this(2) also have "?vc = pdevs_val (\_. e 0) (pdevs_of_list ?coll)" by (simp add: pdevs_val_const_pdevs_of_list) also from pdevs_val_pdevs_of_list_append[OF e_split(1) e'] obtain e'' where e'': "e'' \ UNIV \ I" and "pdevs_val (\_. e 0) (pdevs_of_list ?coll) + pdevs_val e' (pdevs_of_list ?ncoll) = pdevs_val e'' (pdevs_of_list (?coll @ ?ncoll))" by metis note this(2) also from pdevs_val_perm[OF partition_permI e'', of "coll 0 x" "x#xs"] obtain e''' where e''': "e''' \ UNIV \ I" and "\ = pdevs_val e''' (pdevs_of_list (x # xs))" by metis note this(2) finally show ?case using e''' by auto qed lemma list_all_filter: "list_all p (filter p xs)" by (induct xs) auto lemma pdevs_val_independent_pdevs: assumes "list_all (\x. lex x 0) xs" assumes "list_all (\x. x \ 0) xs" assumes "e \ UNIV \ {-1 .. 1}" shows "\e'. e' \ UNIV \ {-1 .. 1} \ pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (independent_pdevs xs))" using assms(1,2,3) proof (induct xs arbitrary: e rule: independent_pdevs.induct) case 1 thus ?case by force next case (2 x xs) let ?coll = "(filter (coll 0 x) (x#xs))" let ?ncoll = "(filter (Not o coll 0 x) xs)" from 2 have "x \ 0" by simp from pdevs_val_partition[OF 2(4), of "x#xs" "coll 0 x"] obtain f g where part: "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val f (pdevs_of_list ?coll) + pdevs_val g (pdevs_of_list (filter (Not o coll 0 x) (x#xs)))" and f: "f \ UNIV \ {-1 .. 1}" and g: "g \ UNIV \ {-1 .. 1}" by blast note part also have "list_all (\x. lex x 0) (filter (coll 0 x) (x#xs))" using 2(2) by (auto simp: inner_prod_def list_all_iff) from pdevs_val_coll[OF list_all_filter this \x \ 0\ f] obtain f' where "pdevs_val f (pdevs_of_list ?coll) = f' *\<^sub>R sum_list (filter (coll 0 x) (x#xs))" and f': "f' \ {-1 .. 1}" by blast note this(1) also have "filter (Not o coll 0 x) (x#xs) = ?ncoll" by simp also from 2(2) have "list_all (\x. lex x 0) ?ncoll" "list_all (\x. x \ 0) ?ncoll" by (auto simp: list_all_iff) from 2(1)[OF refl partition_filter_conv[symmetric] refl this g] obtain g' where "pdevs_val g (pdevs_of_list ?ncoll) = pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll))" and g': "g' \ UNIV \ {-1 .. 1}" by metis note this(1) also define h where "h = (if sum_list ?coll \ 0 then rec_nat f' (\i _. g' i) else g')" have "f' *\<^sub>R sum_list ?coll + pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll)) = pdevs_val h (pdevs_of_list (independent_pdevs (x#xs)))" by (simp add: h_def o_def Let_def) finally have "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val h (pdevs_of_list (independent_pdevs (x # xs)))" . moreover have "h \ UNIV \ {-1 .. 1}" proof fix i show "h i \ {-1 .. 1}" using f' g' by (cases i) (auto simp: h_def) qed ultimately show ?case by blast qed lemma pdevs_val_independent_pdevs_strict: assumes "list_all (\x. lex x 0) xs" assumes "list_all (\x. x \ 0) xs" assumes "e \ UNIV \ {-1 <..< 1}" shows "\e'. e' \ UNIV \ {-1 <..< 1} \ pdevs_val e (pdevs_of_list xs) = pdevs_val e' (pdevs_of_list (independent_pdevs xs))" using assms(1,2,3) proof (induct xs arbitrary: e rule: independent_pdevs.induct) case 1 thus ?case by force next case (2 x xs) let ?coll = "(filter (coll 0 x) (x#xs))" let ?ncoll = "(filter (Not o coll 0 x) xs)" from 2 have "x \ 0" by simp from pdevs_val_partition[OF 2(4), of "x#xs" "coll 0 x"] obtain f g where part: "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val f (pdevs_of_list ?coll) + pdevs_val g (pdevs_of_list (filter (Not o coll 0 x) (x#xs)))" and f: "f \ UNIV \ {-1 <..< 1}" and g: "g \ UNIV \ {-1 <..< 1}" by blast note part also have "list_all (\x. lex x 0) (filter (coll 0 x) (x#xs))" using 2(2) by (auto simp: inner_prod_def list_all_iff) from pdevs_val_coll_strict[OF list_all_filter this \x \ 0\ f] obtain f' where "pdevs_val f (pdevs_of_list ?coll) = f' *\<^sub>R sum_list (filter (coll 0 x) (x#xs))" and f': "f' \ {-1 <..< 1}" by blast note this(1) also have "filter (Not o coll 0 x) (x#xs) = ?ncoll" by simp also from 2(2) have "list_all (\x. lex x 0) ?ncoll" "list_all (\x. x \ 0) ?ncoll" by (auto simp: list_all_iff) from 2(1)[OF refl partition_filter_conv[symmetric] refl this g] obtain g' where "pdevs_val g (pdevs_of_list ?ncoll) = pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll))" and g': "g' \ UNIV \ {-1 <..< 1}" by metis note this(1) also define h where "h = (if sum_list ?coll \ 0 then rec_nat f' (\i _. g' i) else g')" have "f' *\<^sub>R sum_list ?coll + pdevs_val g' (pdevs_of_list (independent_pdevs ?ncoll)) = pdevs_val h (pdevs_of_list (independent_pdevs (x#xs)))" by (simp add: h_def o_def Let_def) finally have "pdevs_val e (pdevs_of_list (x # xs)) = pdevs_val h (pdevs_of_list (independent_pdevs (x # xs)))" . moreover have "h \ UNIV \ {-1 <..< 1}" proof fix i show "h i \ {-1 <..< 1}" using f' g' by (cases i) (auto simp: h_def) qed ultimately show ?case by blast qed lemma sum_list_independent_pdevs: "sum_list (independent_pdevs xs) = sum_list xs" proof (induct xs rule: independent_pdevs.induct) case (2 y ys) from 2[OF refl prod.collapse refl] show ?case using sum_list_partition[of "coll 0 y" ys, symmetric] by (auto simp: Let_def) qed simp lemma independent_pdevs_eq_Nil_iff: "list_all (\x. lex x 0) xs \ list_all (\x. x \ 0) xs \ independent_pdevs xs = [] \ xs = []" proof (induct xs) case Nil thus ?case by simp next case (Cons x xs) from Cons(2) have "list_all (\x. lex x 0) (x # filter (coll 0 x) xs)" by (auto simp: list_all_iff) from sum_list_nlex_eq_zero_iff[OF this] Cons(3) show ?case by (auto simp: list_all_iff) qed subsection \Independent Oriented Generators\ definition "inl p = independent_pdevs (map snd (list_of_pdevs (nlex_pdevs p)))" lemma distinct_inl[simp]: "distinct (inl (snd X))" by (auto simp: inl_def) lemma in_set_inl_nonzero: "x \ set (inl xs) \ x \ 0" by (auto simp: inl_def in_set_independent_pdevs_nonzero) lemma inl_ncoll: "y \ set (inl (snd X)) \ z \ set (inl (snd X)) \ y \ z \ \coll 0 y z" unfolding inl_def by (rule independent_pdevs_pairwise_non_coll, assumption+) (auto simp: inl_def list_of_pdevs_nonzero) lemma in_set_inl_lex: "x \ set (inl xs) \ lex x 0" by (auto simp: inl_def list_of_pdevs_def dest!: in_set_independent_pdevs_invariant_nlex split: if_split_asm) interpretation ccw0: linorder_list "ccw 0" "set (inl (snd X))" proof unfold_locales fix a b c show "a \ b \ ccw 0 a b \ ccw 0 b a" by (metis UNIV_I ccw_self(1) nondegenerate) assume a: "a \ set (inl (snd X))" show "ccw 0 a a" by simp assume b: "b \ set (inl (snd X))" show "ccw 0 a b \ ccw 0 b a \ a = b" by (metis ccw_self(1) in_set_inl_nonzero mem_Collect_eq not_ccw_eq a b) assume c: "c \ set (inl (snd X))" assume distinct: "a \ b" "b \ c" "a \ c" assume ab: "ccw 0 a b" and bc: "ccw 0 b c" show "ccw 0 a c" using a b c ab bc proof (cases "a = (0, 1)" "b = (0, 1)" "c = (0, 1)" rule: case_split[case_product case_split[case_product case_split]]) assume nu: "a \ (0, 1)" "b \ (0, 1)" "c \ (0, 1)" have "distinct5 a b c (0, 1) 0" "in5 UNIV a b c (0, 1) 0" using a b c distinct nu by (simp_all add: in_set_inl_nonzero) moreover have "ccw 0 (0, 1) a" "ccw 0 (0, 1) b" "ccw 0 (0, 1) c" by (auto intro!: nlex_ccw_left in_set_inl_lex a b c) ultimately show ?thesis using ab bc by (rule ccw.transitive[where S=UNIV and s="(0, 1)"]) next assume "a \ (0, 1)" "b = (0, 1)" "c \ (0, 1)" thus ?thesis using ccw_switch23 in_set_inl_lex inl_ncoll nlex_ccw_left a b ab by blast next assume "a \ (0, 1)" "b \ (0, 1)" "c = (0, 1)" thus ?thesis using ccw_switch23 in_set_inl_lex inl_ncoll nlex_ccw_left b c bc by blast qed (auto simp add: nlex_ccw_left in_set_inl_lex) qed lemma sorted_inl: "ccw.sortedP 0 (ccw.selsort 0 (inl (snd X)))" by (rule ccw0.sortedP_selsort) auto lemma sorted_scaled_inl: "ccw.sortedP 0 (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))" using sorted_inl by (rule ccw_sorted_scaleR) simp lemma distinct_selsort_inl: "distinct (ccw.selsort 0 (inl (snd X)))" by simp lemma distinct_map_scaleRI: fixes xs::"'a::real_vector list" shows "distinct xs \ c \ 0 \ distinct (map ((*\<^sub>R) c) xs)" by (induct xs) auto lemma distinct_scaled_inl: "distinct (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))" using distinct_selsort_inl by (rule distinct_map_scaleRI) simp lemma ccw'_sortedP_scaled_inl: "ccw'.sortedP 0 (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))" using ccw_sorted_implies_ccw'_sortedP by (rule ccw'_sorted_scaleR) (auto simp: sorted_inl inl_ncoll) lemma pdevs_val_pdevs_of_list_inl2E: assumes "e \ UNIV \ {-1 .. 1}" obtains e' where "pdevs_val e X = pdevs_val e' (pdevs_of_list (inl X))" "e' \ UNIV \ {-1 .. 1}" proof - let ?l = "map snd (list_of_pdevs (nlex_pdevs X))" have l: "list_all (\x. Counterclockwise_2D_Arbitrary.lex x 0) ?l" "list_all (\x. x \ 0) (map snd (list_of_pdevs (nlex_pdevs X)))" by (auto simp: list_all_iff list_of_pdevs_def) from pdevs_val_nlex_pdevs[OF assms(1)] obtain e' where "e' \ UNIV \ {-1 .. 1}" "pdevs_val e X = pdevs_val e' (nlex_pdevs X)" by auto note this(2) also from pdevs_val_of_list_of_pdevs2[OF \e' \ _\] obtain e'' where "e'' \ UNIV \ {-1 .. 1}" "\ = pdevs_val e'' (pdevs_of_list ?l)" by metis note this(2) also from pdevs_val_independent_pdevs[OF l \e'' \ _\] obtain e''' where "e''' \ UNIV \ {-1 .. 1}" and "\ = pdevs_val e''' (pdevs_of_list (independent_pdevs ?l))" by metis note this(2) also have "\ = pdevs_val e''' (pdevs_of_list (inl X))" by (simp add: inl_def) finally have "pdevs_val e X = pdevs_val e''' (pdevs_of_list (inl X))" . thus thesis using \e''' \ _\ .. qed lemma pdevs_val_pdevs_of_list_inlE: assumes "e \ UNIV \ I" "uminus ` I = I" "0 \ I" obtains e' where "pdevs_val e (pdevs_of_list (inl X)) = pdevs_val e' X" "e' \ UNIV \ I" proof - let ?l = "map snd (list_of_pdevs (nlex_pdevs X))" have "pdevs_val e (pdevs_of_list (inl X)) = pdevs_val e (pdevs_of_list (independent_pdevs ?l))" by (simp add: inl_def) also from pdevs_val_independent_pdevs2[OF \e \ _\] obtain e' where "pdevs_val e (pdevs_of_list (independent_pdevs ?l)) = pdevs_val e' (pdevs_of_list ?l)" and "e' \ UNIV \ I" by auto note this(1) also from pdevs_val_of_list_of_pdevs[OF \e' \ _\ \0 \ I\, of "nlex_pdevs X"] obtain e'' where "pdevs_val e' (pdevs_of_list ?l) = pdevs_val e'' (nlex_pdevs X)" and "e'' \ UNIV \ I" by metis note this(1) also from pdevs_val_nlex_pdevs2[OF \e'' \ _\ \_ = I\] obtain e''' where "pdevs_val e'' (nlex_pdevs X) = pdevs_val e''' X" "e''' \ UNIV \ I" by metis note this(1) finally have "pdevs_val e (pdevs_of_list (inl X)) = pdevs_val e''' X" . thus ?thesis using \e''' \ UNIV \ I\ .. qed lemma sum_list_nlex_eq_sum_list_inl: "sum_list (map snd (list_of_pdevs (nlex_pdevs X))) = sum_list (inl X)" by (auto simp: inl_def sum_list_list_of_pdevs sum_list_independent_pdevs) lemma Affine_inl: "Affine (fst X, pdevs_of_list (inl (snd X))) = Affine X" by (auto simp: Affine_def valuate_def aform_val_def elim: pdevs_val_pdevs_of_list_inlE[of _ _ "snd X"] pdevs_val_pdevs_of_list_inl2E[of _ "snd X"]) subsection \Half Segments\ definition half_segments_of_aform::"point aform \ (point*point) list" where "half_segments_of_aform X = (let x0 = lowest_vertex (fst X, nlex_pdevs (snd X)) in polychain_of x0 (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))))" lemma subsequent_half_segments: fixes X assumes "Suc i < length (half_segments_of_aform X)" shows "snd (half_segments_of_aform X ! i) = fst (half_segments_of_aform X ! Suc i)" using assms by (cases i) (auto simp: half_segments_of_aform_def Let_def polychain_of_subsequent_eq) lemma polychain_half_segments_of_aform: "polychain (half_segments_of_aform X)" by (auto simp: subsequent_half_segments intro!: polychainI) lemma fst_half_segments: "half_segments_of_aform X \ [] \ fst (half_segments_of_aform X ! 0) = lowest_vertex (fst X, nlex_pdevs (snd X))" by (auto simp: half_segments_of_aform_def Let_def o_def split_beta') lemma nlex_half_segments_of_aform: "(a, b) \ set (half_segments_of_aform X) \ lex b a" by (auto simp: half_segments_of_aform_def prod_eq_iff lex_def dest!: in_set_polychain_ofD in_set_inl_lex) lemma ccw_half_segments_of_aform_all: assumes cd: "(c, d) \ set (half_segments_of_aform X)" shows "list_all (\(xi, xj). ccw xi xj c \ ccw xi xj d) (half_segments_of_aform X)" proof - have "list_all (\(xi, xj). ccw xi xj (fst (c, d)) \ ccw xi xj (snd (c, d))) (polychain_of (lowest_vertex (fst X, nlex_pdevs (snd X))) ((map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X))))))" using ccw'_sortedP_scaled_inl cd[unfolded half_segments_of_aform_def Let_def] by (rule polychain_of_ccw_conjunction) thus ?thesis unfolding half_segments_of_aform_def[unfolded Let_def, symmetric] fst_conv snd_conv . qed lemma ccw_half_segments_of_aform: assumes ij: "(xi, xj) \ set (half_segments_of_aform X)" assumes c: "(c, d) \ set (half_segments_of_aform X)" shows "ccw xi xj c" "ccw xi xj d" using ccw_half_segments_of_aform_all[OF c] ij by (auto simp add: list_all_iff) lemma half_segments_of_aform1: assumes ch: "x \ convex hull set (map fst (half_segments_of_aform X))" assumes ab: "(a, b) \ set (half_segments_of_aform X)" shows "ccw a b x" using finite_set _ ch proof (rule ccw.convex_hull) fix c assume "c \ set (map fst (half_segments_of_aform X))" then obtain d where "(c, d) \ set (half_segments_of_aform X)" by auto with ab show "ccw a b c" by (rule ccw_half_segments_of_aform(1)) qed (insert ab, simp add: nlex_half_segments_of_aform) lemma half_segments_of_aform2: assumes ch: "x \ convex hull set (map snd (half_segments_of_aform X))" assumes ab: "(a, b) \ set (half_segments_of_aform X)" shows "ccw a b x" using finite_set _ ch proof (rule ccw.convex_hull) fix d assume "d \ set (map snd (half_segments_of_aform X))" then obtain c where "(c, d) \ set (half_segments_of_aform X)" by auto with ab show "ccw a b d" by (rule ccw_half_segments_of_aform(2)) qed (insert ab, simp add: nlex_half_segments_of_aform) lemma in_set_half_segments_of_aform_aform_valE: assumes "(x2, y2) \ set (half_segments_of_aform X)" obtains e where "y2 = aform_val e X" "e \ UNIV \ {-1 .. 1}" proof - from assms obtain d where "y2 = lowest_vertex (fst X, nlex_pdevs (snd X)) + sum_list (take (Suc d) (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))))" by (auto simp: half_segments_of_aform_def elim!: in_set_polychain_of_imp_sum_list) also have "lowest_vertex (fst X, nlex_pdevs (snd X)) = fst X - sum_list (map snd (list_of_pdevs (nlex_pdevs (snd X))))" by (simp add: lowest_vertex_def) also have "sum_list (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = pdevs_val (\_. 1) (nlex_pdevs (snd X))" by (auto simp: pdevs_val_sum_list) also have "sum_list (take (Suc d) (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))) = pdevs_val (\i. if i \ d then 2 else 0) (pdevs_of_list (ccw.selsort 0 (inl (snd X))))" (is "_ = pdevs_val ?e _") by (subst sum_list_take_pdevs_val_eq) (auto simp: pdevs_val_sum if_distrib pdevs_apply_pdevs_of_list degree_pdevs_of_list_scaleR intro!: sum.cong ) also obtain e'' where "\ = pdevs_val e'' (pdevs_of_list (inl (snd X)))" "e'' \ UNIV \ {0..2}" by (auto intro: pdevs_val_selsort_ccw2[of "inl (snd X)" ?e "{0 .. 2}"]) note this(1) also note inl_def also let ?l = "map snd (list_of_pdevs (nlex_pdevs (snd X)))" from pdevs_val_independent_pdevs2[OF \e'' \ _\] obtain e''' where "pdevs_val e'' (pdevs_of_list (independent_pdevs ?l)) = pdevs_val e''' (pdevs_of_list ?l)" and "e''' \ UNIV \ {0..2}" by auto note this(1) also have "0 \ {0 .. 2::real}" by simp from pdevs_val_of_list_of_pdevs[OF \e''' \ _\ this, of "nlex_pdevs (snd X)"] obtain e'''' where "pdevs_val e''' (pdevs_of_list ?l) = pdevs_val e'''' (nlex_pdevs (snd X))" and "e'''' \ UNIV \ {0 .. 2}" by metis note this(1) finally have "y2 = fst X + (pdevs_val e'''' (nlex_pdevs (snd X)) - pdevs_val (\_. 1) (nlex_pdevs (snd X)))" by simp also have "pdevs_val e'''' (nlex_pdevs (snd X)) - pdevs_val (\_. 1) (nlex_pdevs (snd X)) = pdevs_val (\i. e'''' i - 1) (nlex_pdevs (snd X))" by (simp add: pdevs_val_minus) also have "(\i. e'''' i - 1) \ UNIV \ {-1 .. 1}" using \e'''' \ _\ by auto from pdevs_val_nlex_pdevs2[OF this] obtain f where "f \ UNIV \ {-1 .. 1}" and "pdevs_val (\i. e'''' i - 1) (nlex_pdevs (snd X)) = pdevs_val f (snd X)" by auto note this(2) finally have "y2 = aform_val f X" by (simp add: aform_val_def) thus ?thesis using \f \ _\ .. qed lemma fst_hd_half_segments_of_aform: assumes "half_segments_of_aform X \ []" shows "fst (hd (half_segments_of_aform X)) = lowest_vertex (fst X, (nlex_pdevs (snd X)))" using assms by (auto simp: half_segments_of_aform_def Let_def fst_hd_polychain_of) lemma "linorder_list0.sortedP (ccw' (lowest_vertex (fst X, nlex_pdevs (snd X)))) (map snd (half_segments_of_aform X))" (is "linorder_list0.sortedP (ccw' ?x0) ?ms") unfolding half_segments_of_aform_def Let_def by (rule ccw'_sortedP_polychain_of_snd) (rule ccw'_sortedP_scaled_inl) lemma rev_zip: "length xs = length ys \ rev (zip xs ys) = zip (rev xs) (rev ys)" by (induct xs ys rule: list_induct2) auto lemma zip_upt_self_aux: "zip [0..i. (i, xs ! i)) [0.. UNIV \ {-1 <..< 1}" assumes "seg \ set (half_segments_of_aform X)" assumes "length (half_segments_of_aform X) \ 1" shows "ccw' (fst seg) (snd seg) (aform_val e X)" using assms unfolding half_segments_of_aform_def Let_def proof - have len: "length (map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X)))) \ 1" using assms by (auto simp: half_segments_of_aform_def) have "aform_val e X = fst X + pdevs_val e (snd X)" by (simp add: aform_val_def) also obtain e' where "e' \ UNIV \ {-1 <..< 1}" "pdevs_val e (snd X) = pdevs_val e' (nlex_pdevs (snd X))" using pdevs_val_nlex_pdevs[OF \e \ _\] by auto note this(2) also obtain e'' where "e'' \ UNIV \ {-1 <..< 1}" "\ = pdevs_val e'' (pdevs_of_list (map snd (list_of_pdevs (nlex_pdevs (snd X)))))" by (metis pdevs_val_of_list_of_pdevs2[OF \e' \ _\]) note this(2) also obtain e''' where "e''' \ UNIV \ {-1 <..< 1}" "\ = pdevs_val e''' (pdevs_of_list (inl (snd X)))" unfolding inl_def using pdevs_val_independent_pdevs_strict[OF list_all_list_of_pdevsI, OF lex_nlex_pdevs list_of_pdevs_all_nonzero \e'' \ _\] by metis note this(2) also from pdevs_val_selsort_ccw[OF distinct_inl \e''' \ _\] obtain f where "f \ UNIV \ {-1 <..< 1}" "\ = pdevs_val f (pdevs_of_list (linorder_list0.selsort (ccw 0) (inl (snd X))))" (is "_ = pdevs_val _ (pdevs_of_list ?sl)") by metis note this(2) also have "\ = pdevs_val (\i. f i + 1) (pdevs_of_list ?sl) + lowest_vertex (fst X, nlex_pdevs (snd X)) - fst X" proof - have "sum_list (dense_list_of_pdevs (nlex_pdevs (snd X))) = sum_list (dense_list_of_pdevs (pdevs_of_list (ccw.selsort 0 (inl (snd X)))))" by (subst dense_list_of_pdevs_pdevs_of_list) (auto simp: in_set_independent_pdevs_nonzero dense_list_of_pdevs_pdevs_of_list inl_def sum_list_distinct_selsort sum_list_independent_pdevs sum_list_list_of_pdevs) thus ?thesis by (auto simp add: pdevs_val_add lowest_vertex_def algebra_simps pdevs_val_sum_list sum_list_list_of_pdevs in_set_inl_nonzero dense_list_of_pdevs_pdevs_of_list) qed also have "pdevs_val (\i. f i + 1) (pdevs_of_list ?sl) = pdevs_val (\i. 1/2 * (f i + 1)) (pdevs_of_list (map ((*\<^sub>R) 2) ?sl))" (is "_ = pdevs_val ?f' (pdevs_of_list ?ssl)") by (subst pdevs_val_cmul) (simp add: pdevs_of_list_map_scaleR) also have "distinct ?ssl" "?f' \ UNIV \ {0<..<1}" using \f \ _\ by (auto simp: distinct_map_scaleRI Pi_iff algebra_simps real_0_less_add_iff) from pdevs_of_list_sum[OF this] obtain g where "g \ UNIV \ {0<..<1}" "pdevs_val ?f' (pdevs_of_list ?ssl) = (\P\set ?ssl. g P *\<^sub>R P)" by blast note this(2) finally have "aform_val e X = lowest_vertex (fst X, nlex_pdevs (snd X)) + (\P\set ?ssl. g P *\<^sub>R P)" by simp also have "ccw' (fst seg) (snd seg) \" using \g \ _\ _ len \seg \ _\[unfolded half_segments_of_aform_def Let_def] by (rule in_polychain_of_ccw) (simp add: ccw'_sortedP_scaled_inl) finally show ?thesis . qed lemma half_segments_of_aform_strict_all: assumes "e \ UNIV \ {-1 <..< 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. ccw' (fst seg) (snd seg) (aform_val e X)) (half_segments_of_aform X)" using assms by (auto intro!: half_segments_of_aform_strict simp: list_all_iff) lemma list_allD: "list_all P xs \ x \ set xs \ P x" by (auto simp: list_all_iff) lemma minus_one_less_multI: fixes e x::real shows "- 1 \ e \ 0 < x \ x < 1 \ - 1 < e * x" by (metis abs_add_one_gt_zero abs_real_def le_less_trans less_not_sym mult_less_0_iff mult_less_cancel_left1 real_0_less_add_iff) lemma less_one_multI: fixes e x::real shows "e \ 1 \ 0 < x \ x < 1 \ e * x < 1" by (metis (erased, opaque_lifting) less_eq_real_def monoid_mult_class.mult.left_neutral mult_strict_mono zero_less_one) lemma ccw_half_segments_of_aform_strictI: assumes "e \ UNIV \ {-1 <..< 1}" assumes "(s1, s2) \ set (half_segments_of_aform X)" assumes "length (half_segments_of_aform X) \ 1" assumes "x = (aform_val e X)" shows "ccw' s1 s2 x" using half_segments_of_aform_strict[OF assms(1-3)] assms(4) by simp lemma ccw'_sortedP_subsequent: assumes "Suc i < length xs" "ccw'.sortedP 0 (map dirvec xs)" "fst (xs ! Suc i) = snd (xs ! i)" shows "ccw' (fst (xs ! i)) (snd (xs ! i)) (snd (xs ! Suc i))" using assms proof (induct xs arbitrary: i) case Nil thus ?case by simp next case (Cons x xs) thus ?case by (auto simp: nth_Cons dirvec_minus split: nat.split elim!: ccw'.sortedP_Cons) (metis (no_types, lifting) ccw'.renormalize length_greater_0_conv nth_mem prod.case_eq_if) qed lemma ccw'_sortedP_uminus: "ccw'.sortedP 0 xs \ ccw'.sortedP 0 (map uminus xs)" by (induct xs) (auto intro!: ccw'.sortedP.Cons elim!: ccw'.sortedP_Cons simp del: uminus_Pair) lemma subsequent_half_segments_ccw: fixes X assumes "Suc i < length (half_segments_of_aform X)" shows "ccw' (fst (half_segments_of_aform X ! i)) (snd (half_segments_of_aform X ! i)) (snd (half_segments_of_aform X ! Suc i))" using assms by (intro ccw'_sortedP_subsequent ) (auto simp: subsequent_half_segments half_segments_of_aform_def sorted_inl polychain_of_subsequent_eq intro!: ccw_sorted_implies_ccw'_sortedP[OF inl_ncoll] ccw'_sorted_scaleR) lemma convex_polychain_half_segments_of_aform: "convex_polychain (half_segments_of_aform X)" proof cases assume "length (half_segments_of_aform X) = 1" thus ?thesis by (auto simp: length_Suc_conv convex_polychain_def polychain_def) next assume len: "length (half_segments_of_aform X) \ 1" show ?thesis by (rule convex_polychainI) (simp_all add: polychain_half_segments_of_aform subsequent_half_segments_ccw ccw'_def[symmetric]) qed lemma hd_distinct_neq_last: "distinct xs \ length xs > 1 \ hd xs \ last xs" by (metis One_nat_def add_Suc_right distinct.simps(2) last.simps last_in_set less_irrefl list.exhaust list.sel(1) list.size(3) list.size(4) add.right_neutral nat_neq_iff not_less0) lemma ccw_hd_last_half_segments_dirvec: assumes "length (half_segments_of_aform X) > 1" shows "ccw' 0 (dirvec (hd (half_segments_of_aform X))) (dirvec (last (half_segments_of_aform X)))" proof - let ?i = "ccw.selsort 0 (inl (snd X))" let ?s = "map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))" from assms have l: "1 < length (inl (snd X))" "inl (snd X) \ []" using assms by (auto simp add: half_segments_of_aform_def) hence "hd ?i \ set ?i" "last ?i \ set ?i" by (auto intro!: hd_in_set last_in_set simp del: ccw.set_selsort) hence "\coll 0 (hd ?i) (last ?i)" using l by (intro inl_ncoll[of _ X]) (auto simp: hd_distinct_neq_last) hence "\coll 0 (hd ?s) (last ?s)" using l by (auto simp: hd_map last_map) hence "ccw' 0 (hd (map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X))))) (last (map ((*\<^sub>R) 2) (linorder_list0.selsort (ccw 0) (inl (snd X)))))" using assms by (auto simp add: half_segments_of_aform_def intro!: sorted_inl ccw_sorted_scaleR ccw.hd_last_sorted ccw_ncoll_imp_ccw) with assms show ?thesis by (auto simp add: half_segments_of_aform_def Let_def dirvec_hd_polychain_of dirvec_last_polychain_of length_greater_0_conv[symmetric] simp del: polychain_of.simps length_greater_0_conv split: if_split_asm) qed lemma map_fst_half_segments_aux_eq: "[] \ half_segments_of_aform X \ map fst (half_segments_of_aform X) = fst (hd (half_segments_of_aform X))#butlast (map snd (half_segments_of_aform X))" by (rule nth_equalityI) (auto simp: nth_Cons hd_conv_nth nth_butlast subsequent_half_segments split: nat.split) lemma le_less_Suc_eq: "x - Suc 0 \ (i::nat) \ i < x \ x - Suc 0 = i" by simp lemma map_snd_half_segments_aux_eq: "half_segments_of_aform X \ [] \ map snd (half_segments_of_aform X) = tl (map fst (half_segments_of_aform X)) @ [snd (last (half_segments_of_aform X))]" by (rule nth_equalityI) (auto simp: nth_Cons hd_conv_nth nth_append nth_tl subsequent_half_segments not_less last_conv_nth algebra_simps dest!: le_less_Suc_eq split: nat.split) lemma ccw'_sortedP_snd_half_segments_of_aform: "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (map snd (half_segments_of_aform X))" by (auto simp: half_segments_of_aform_def Let_def intro!: ccw'.sortedP.Cons ccw'_sortedP_polychain_of_snd ccw'_sortedP_scaled_inl) lemma lex_half_segments_lowest_vertex: assumes "(c, d) \ set (half_segments_of_aform X)" shows "lex d (lowest_vertex (fst X, nlex_pdevs (snd X)))" unfolding half_segments_of_aform_def Let_def by (rule lex_polychain_of_center[OF assms[unfolded half_segments_of_aform_def Let_def], unfolded snd_conv]) (auto simp: list_all_iff lex_def dest!: in_set_inl_lex) lemma lex_half_segments_lowest_vertex': assumes "d \ set (map snd (half_segments_of_aform X))" shows "lex d (lowest_vertex (fst X, nlex_pdevs (snd X)))" using assms by (auto intro: lex_half_segments_lowest_vertex) lemma lex_half_segments_last: assumes "(c, d) \ set (half_segments_of_aform X)" shows "lex (snd (last (half_segments_of_aform X))) d" using assms unfolding half_segments_of_aform_def Let_def by (rule lex_polychain_of_last) (auto simp: list_all_iff lex_def dest!: in_set_inl_lex) lemma lex_half_segments_last': assumes "d \ set (map snd (half_segments_of_aform X))" shows "lex (snd (last (half_segments_of_aform X))) d" using assms by (auto intro: lex_half_segments_last) lemma ccw'_half_segments_lowest_last: assumes set_butlast: "(c, d) \ set (butlast (half_segments_of_aform X))" assumes ne: "inl (snd X) \ []" shows "ccw' (lowest_vertex (fst X, nlex_pdevs (snd X))) d (snd (last (half_segments_of_aform X)))" using set_butlast unfolding half_segments_of_aform_def Let_def by (rule ccw'_polychain_of_sorted_center_last) (auto simp: ne ccw'_sortedP_scaled_inl) lemma distinct_fst_half_segments: "distinct (map fst (half_segments_of_aform X))" by (auto simp: half_segments_of_aform_def list_all_iff lex_scale1_zero simp del: scaleR_Pair intro!: distinct_fst_polychain_of dest: in_set_inl_nonzero in_set_inl_lex) lemma distinct_snd_half_segments: "distinct (map snd (half_segments_of_aform X))" by (auto simp: half_segments_of_aform_def list_all_iff lex_scale1_zero simp del: scaleR_Pair intro!: distinct_snd_polychain_of dest: in_set_inl_nonzero in_set_inl_lex) subsection \Mirror\ definition "mirror_point x y = 2 *\<^sub>R x - y" lemma ccw'_mirror_point3[simp]: "ccw' (mirror_point x y) (mirror_point x z) (mirror_point x w) \ ccw' y z w " by (auto simp: mirror_point_def det3_def' ccw'_def algebra_simps) lemma mirror_point_self_inverse[simp]: fixes x::"'a::real_vector" shows "mirror_point p (mirror_point p x) = x" by (auto simp: mirror_point_def scaleR_2) lemma mirror_half_segments_of_aform: assumes "e \ UNIV \ {-1 <..< 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. ccw' (fst seg) (snd seg) (aform_val e X)) (map (pairself (mirror_point (fst X))) (half_segments_of_aform X))" unfolding list_all_length proof safe let ?m = "map (pairself (mirror_point (fst X))) (half_segments_of_aform X)" fix n assume "n < length ?m" hence "ccw' (fst (half_segments_of_aform X ! n)) (snd (half_segments_of_aform X ! n)) (aform_val (- e) X)" using assms by (auto dest!: nth_mem intro!: half_segments_of_aform_strict) also have "aform_val (-e) X = 2 *\<^sub>R fst X - aform_val e X" by (auto simp: aform_val_def pdevs_val_sum algebra_simps scaleR_2 sum_negf) finally have le: "ccw' (fst (half_segments_of_aform X ! n)) (snd (half_segments_of_aform X ! n)) (2 *\<^sub>R fst X - aform_val e X)" . have eq: "(map (pairself (mirror_point (fst X))) (half_segments_of_aform X) ! n) = (2 *\<^sub>R fst X - fst ((half_segments_of_aform X) ! n), 2 *\<^sub>R fst X - snd ((half_segments_of_aform X) ! n))" using \n < length ?m\ by (cases "half_segments_of_aform X ! n") (auto simp add: mirror_point_def) show "ccw' (fst (?m ! n)) (snd (?m ! n)) (aform_val e X)" using le unfolding eq by (auto simp: algebra_simps ccw'_def det3_def') qed lemma last_half_segments: assumes "half_segments_of_aform X \ []" shows "snd (last (half_segments_of_aform X)) = mirror_point (fst X) (lowest_vertex (fst X, nlex_pdevs (snd X)))" using assms by (auto simp add: half_segments_of_aform_def Let_def lowest_vertex_def mirror_point_def scaleR_2 scaleR_sum_list[symmetric] last_polychain_of sum_list_distinct_selsort inl_def sum_list_independent_pdevs sum_list_list_of_pdevs) lemma convex_polychain_map_mirror: assumes "convex_polychain hs" shows "convex_polychain (map (pairself (mirror_point x)) hs)" proof (rule convex_polychainI) qed (insert assms, auto simp: convex_polychain_def polychain_map_pairself pairself_apply mirror_point_def det3_def' algebra_simps) lemma ccw'_mirror_point0: "ccw' (mirror_point x y) z w \ ccw' y (mirror_point x z) (mirror_point x w)" by (metis ccw'_mirror_point3 mirror_point_self_inverse) lemma ccw'_sortedP_mirror: "ccw'.sortedP x0 (map (mirror_point p0) xs) \ ccw'.sortedP (mirror_point p0 x0) xs" by (induct xs) (simp_all add: linorder_list0.sortedP.Nil linorder_list0.sortedP_Cons_iff ccw'_mirror_point0) lemma ccw'_sortedP_mirror2: "ccw'.sortedP (mirror_point p0 x0) (map (mirror_point p0) xs) \ ccw'.sortedP x0 xs" using ccw'_sortedP_mirror[of "mirror_point p0 x0" p0 xs] by simp lemma map_mirror_o_snd_polychain_of_eq: "map (mirror_point x0 \ snd) (polychain_of y xs) = map snd (polychain_of (mirror_point x0 y) (map uminus xs))" by (induct xs arbitrary: x0 y) (auto simp: mirror_point_def o_def algebra_simps) lemma lowest_vertex_eq_mirror_last: "half_segments_of_aform X \ [] \ (lowest_vertex (fst X, nlex_pdevs (snd X))) = mirror_point (fst X) (snd (last (half_segments_of_aform X)))" using last_half_segments[of X] by simp lemma snd_last: "xs \ [] \ snd (last xs) = last (map snd xs)" by (induct xs) auto lemma mirror_point_snd_last_eq_lowest: "half_segments_of_aform X \ [] \ mirror_point (fst X) (last (map snd (half_segments_of_aform X))) = lowest_vertex (fst X, nlex_pdevs (snd X))" by (simp add: lowest_vertex_eq_mirror_last snd_last) lemma lex_mirror_point: "lex (mirror_point x0 a) (mirror_point x0 b) \ lex b a" by (auto simp: mirror_point_def lex_def) lemma ccw'_mirror_point: "ccw' (mirror_point x0 a) (mirror_point x0 b) (mirror_point x0 c) \ ccw' a b c" by (auto simp: mirror_point_def) lemma inj_mirror_point: "inj (mirror_point (x::'a::real_vector))" by (auto simp: mirror_point_def inj_on_def algebra_simps) lemma distinct_map_mirror_point_eq: "distinct (map (mirror_point (x::'a::real_vector)) xs) = distinct xs" by (auto simp: distinct_map intro!: subset_inj_on[OF inj_mirror_point]) lemma eq_self_mirror_iff: fixes x::"'a::real_vector" shows "x = mirror_point y x \ x = y" by (auto simp: mirror_point_def algebra_simps scaleR_2[symmetric]) subsection \Full Segments\ definition segments_of_aform::"point aform \ (point * point) list" where "segments_of_aform X = (let hs = half_segments_of_aform X in hs @ map (pairself (mirror_point (fst X))) hs)" lemma segments_of_aform_strict: assumes "e \ UNIV \ {-1 <..< 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. ccw' (fst seg) (snd seg) (aform_val e X)) (segments_of_aform X)" using assms by (auto simp: segments_of_aform_def Let_def mirror_half_segments_of_aform half_segments_of_aform_strict_all) lemma mirror_point_aform_val[simp]: "mirror_point (fst X) (aform_val e X) = aform_val (- e) X" by (auto simp: mirror_point_def aform_val_def pdevs_val_sum algebra_simps scaleR_2 sum_negf) lemma in_set_segments_of_aform_aform_valE: assumes "(x2, y2) \ set (segments_of_aform X)" obtains e where "y2 = aform_val e X" "e \ UNIV \ {-1 .. 1}" using assms proof (auto simp: segments_of_aform_def Let_def) assume "(x2, y2) \ set (half_segments_of_aform X)" from in_set_half_segments_of_aform_aform_valE[OF this] obtain e where "y2 = aform_val e X" "e \ UNIV \ {- 1..1}" by auto thus ?thesis .. next fix a b aa ba assume "((a, b), aa, ba) \ set (half_segments_of_aform X)" from in_set_half_segments_of_aform_aform_valE[OF this] obtain e where e: "(aa, ba) = aform_val e X" "e \ UNIV \ {- 1..1}" by auto assume "y2 = mirror_point (fst X) (aa, ba)" hence "y2 = aform_val (-e) X" "(-e) \ UNIV \ {-1 .. 1}" using e by auto thus ?thesis .. qed lemma last_half_segments_eq_mirror_hd: assumes "half_segments_of_aform X \ []" shows "snd (last (half_segments_of_aform X)) = mirror_point (fst X) (fst (hd (half_segments_of_aform X)))" by (simp add: last_half_segments assms fst_hd_half_segments_of_aform) lemma polychain_segments_of_aform: "polychain (segments_of_aform X)" by (auto simp: segments_of_aform_def Let_def polychain_half_segments_of_aform polychain_map_pairself last_half_segments_eq_mirror_hd hd_map pairself_apply intro!: polychain_appendI) lemma segments_of_aform_closed: assumes "segments_of_aform X \ []" shows "fst (hd (segments_of_aform X)) = snd (last (segments_of_aform X))" using assms by (auto simp: segments_of_aform_def Let_def fst_hd_half_segments_of_aform last_map pairself_apply last_half_segments mirror_point_def) lemma convex_polychain_segments_of_aform: assumes "1 < length (half_segments_of_aform X)" shows "convex_polychain (segments_of_aform X)" unfolding segments_of_aform_def Let_def using ccw_hd_last_half_segments_dirvec[OF assms] by (intro convex_polychain_appendI) (auto simp: convex_polychain_half_segments_of_aform convex_polychain_map_mirror dirvec_minus hd_map pairself_apply last_half_segments mirror_point_def fst_hd_half_segments_of_aform det3_def' algebra_simps ccw'_def intro!: polychain_appendI polychain_half_segments_of_aform polychain_map_pairself) lemma convex_polygon_segments_of_aform: assumes "1 < length (half_segments_of_aform X)" shows "convex_polygon (segments_of_aform X)" proof - from assms have hne: "half_segments_of_aform X \ []" by auto from convex_polychain_segments_of_aform[OF assms] have "convex_polychain (segments_of_aform X)" . thus ?thesis by (auto simp: convex_polygon_def segments_of_aform_closed) qed lemma previous_segments_of_aformE: assumes "(y, z) \ set (segments_of_aform X)" obtains x where "(x, y) \ set (segments_of_aform X)" proof - from assms have ne[simp]: "segments_of_aform X \ []" by auto from assms obtain i where i: "i set (segments_of_aform X)" by (metis fst_conv hd_conv_nth i(2) last_in_set ne snd_conv surj_pair) thus ?thesis .. next case (Suc j) have "(fst (segments_of_aform X ! j), snd (segments_of_aform X ! j)) \ set (segments_of_aform X)" using Suc i(1) by auto also from i have "y = fst (segments_of_aform X ! i)" by auto hence "snd (segments_of_aform X ! j) = y" using polychain_segments_of_aform[of X] i(1) Suc by (auto simp: polychain_def Suc) finally have "(fst (segments_of_aform X ! j), y) \ set (segments_of_aform X)" . thus ?thesis .. qed qed lemma fst_compose_pairself: "fst o pairself f = f o fst" and snd_compose_pairself: "snd o pairself f = f o snd" by (auto simp: pairself_apply) lemma in_set_butlastI: "xs \ [] \ x \ set xs \ x \ last xs \ x \ set (butlast xs)" by (induct xs) (auto split: if_splits) lemma distinct_in_set_butlastD: "x \ set (butlast xs) \ distinct xs \ x \ last xs" by (induct xs) auto lemma distinct_in_set_tlD: "x \ set (tl xs) \ distinct xs \ x \ hd xs" by (induct xs) auto lemma ccw'_sortedP_snd_segments_of_aform: assumes "length (inl (snd X)) > 1" shows "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (butlast (map snd (segments_of_aform X)))" proof cases assume "[] = half_segments_of_aform X" from this show ?thesis by (simp add: segments_of_aform_def Let_def ccw'.sortedP.Nil) next assume H: "[] \ half_segments_of_aform X" let ?m = "mirror_point (fst X)" have ne: "inl (snd X) \ []" using assms by auto have "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (map snd (half_segments_of_aform X) @ butlast (map (?m \ snd) (half_segments_of_aform X)))" proof (rule ccw'.sortedP_appendI) show "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (map snd (half_segments_of_aform X))" by (rule ccw'_sortedP_snd_half_segments_of_aform) have "butlast (map (?m \ snd) (half_segments_of_aform X)) = butlast (map (?m \ snd) (polychain_of (lowest_vertex (fst X, nlex_pdevs (snd X))) (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X))))))" by (simp add: half_segments_of_aform_def) also have "\ = map snd (butlast (polychain_of (?m (lowest_vertex (fst X, nlex_pdevs (snd X)))) (map uminus (map ((*\<^sub>R) 2) (ccw.selsort 0 (inl (snd X)))))))" (is "_ = map snd (butlast (polychain_of ?x ?xs))") by (simp add: map_mirror_o_snd_polychain_of_eq map_butlast) also { have "ccw'.sortedP 0 ?xs" by (intro ccw'_sortedP_uminus ccw'_sortedP_scaled_inl) moreover have "ccw'.sortedP ?x (map snd (polychain_of ?x ?xs))" unfolding ccw'_sortedP_mirror[symmetric] map_map map_mirror_o_snd_polychain_of_eq by (auto simp add: o_def intro!: ccw'_sortedP_polychain_of_snd ccw'_sortedP_scaled_inl) ultimately have "ccw'.sortedP (snd (last (polychain_of ?x ?xs))) (map snd (butlast (polychain_of ?x ?xs)))" by (rule ccw'_sortedP_convex_rotate_aux) } also have "(snd (last (polychain_of ?x ?xs))) = ?m (last (map snd (half_segments_of_aform X)))" by (simp add: half_segments_of_aform_def ne map_mirror_o_snd_polychain_of_eq last_map[symmetric, where f="?m"] last_map[symmetric, where f="snd"]) also have "\ = lowest_vertex (fst X, nlex_pdevs (snd X))" using ne H by (auto simp: lowest_vertex_eq_mirror_last snd_last) finally show "ccw'.sortedP (lowest_vertex (fst X, nlex_pdevs (snd X))) (butlast (map (?m \ snd) (half_segments_of_aform X)))" . next fix x y assume seg: "x \ set (map snd (half_segments_of_aform X))" and mseg: "y \ set (butlast (map (?m \ snd) (half_segments_of_aform X)))" from seg assms have neq_Nil: "inl (snd X) \ []" "half_segments_of_aform X \ []" by auto from seg obtain a where a: "(a, x) \ set (half_segments_of_aform X)" by auto from mseg obtain b where mirror_y: "(b, ?m y) \ set (butlast (half_segments_of_aform X))" by (auto simp: map_butlast[symmetric]) let ?l = "lowest_vertex (fst X, nlex_pdevs (snd X))" let ?ml = "?m ?l" have mirror_eq_last: "?ml = snd (last (half_segments_of_aform X))" using seg H by (intro last_half_segments[symmetric]) simp define r where "r = ?l + (0, abs (snd x - snd ?l) + abs (snd y - snd ?l) + abs (snd ?ml - snd ?l) + 1)" have d1: "x \ r" "y \ r" "?l \ r" "?ml \ r" by (auto simp: r_def plus_prod_def prod_eq_iff) have "distinct (map (?m \ snd) (half_segments_of_aform X))" unfolding map_comp_map[symmetric] unfolding o_def distinct_map_mirror_point_eq by (rule distinct_snd_half_segments) from distinct_in_set_butlastD[OF \y \ _\ this] have "?l \ y" by (simp add: neq_Nil lowest_vertex_eq_mirror_last last_map) moreover have "?l \ ?ml" using neq_Nil by (auto simp add: eq_self_mirror_iff lowest_vertex_eq_center_iff inl_def) ultimately have d2: "?l \ y" "?l \ ?ml" by auto have *: "ccw' ?l (?m y) ?ml" by (metis mirror_eq_last ccw'_half_segments_lowest_last mirror_y neq_Nil(1)) have "ccw' ?ml y ?l" by (rule ccw'_mirror_point[of "fst X"]) (simp add: *) hence lmy: "ccw' ?l ?ml y" by (simp add: ccw'_def det3_def' algebra_simps) let ?ccw = "ccw' (lowest_vertex (fst X, nlex_pdevs (snd X))) x y" { assume "x \ ?ml" hence x_butlast: "(a, x) \ set (butlast (half_segments_of_aform X))" unfolding mirror_eq_last using a by (auto intro!: in_set_butlastI simp: prod_eq_iff) have "ccw' ?l x ?ml" by (metis mirror_eq_last ccw'_half_segments_lowest_last x_butlast neq_Nil(1)) } note lxml = this { assume "x = ?ml" hence ?ccw by (simp add: lmy) } moreover { assume "x \ ?ml" "y = ?ml" hence ?ccw by (simp add: lxml) } moreover { assume d3: "x \ ?ml" "y \ ?ml" from \x \ set _\ have "x \ set (map snd (half_segments_of_aform X))" by force hence "x \ set (tl (map fst (half_segments_of_aform X)))" using d3 unfolding map_snd_half_segments_aux_eq[OF neq_Nil(2)] by (auto simp: mirror_eq_last) from distinct_in_set_tlD[OF this distinct_fst_half_segments] have "?l \ x" by (simp add: fst_hd_half_segments_of_aform neq_Nil hd_map) from lxml[OF \x \ ?ml\] \ccw' ?l ?ml y\ have d4: "x \ y" by (rule neq_left_right_of lxml) have "distinct5 x ?ml y r ?l" using d1 d2 \?l \ x\ d3 d4 by simp_all moreover note _ moreover have "lex x ?l" by (rule lex_half_segments_lowest_vertex) fact hence "ccw ?l r x" unfolding r_def by (rule lex_ccw_left) simp moreover have "lex ?ml ?l" using last_in_set[OF H[symmetric]] by (auto simp: mirror_eq_last intro: lex_half_segments_lowest_vertex') hence "ccw ?l r ?ml" unfolding r_def by (rule lex_ccw_left) simp moreover have "lex (?m (lowest_vertex (fst X, nlex_pdevs (snd X)))) (?m y)" using mirror_y by (force dest!: in_set_butlastD intro: lex_half_segments_last' simp: mirror_eq_last ) hence "lex y ?l" by (rule lex_mirror_point) hence "ccw ?l r y" unfolding r_def by (rule lex_ccw_left) simp moreover from \x \ ?ml\ have "ccw ?l x ?ml" by (simp add: ccw_def lxml) moreover from lmy have "ccw ?l ?ml y" by (simp add: ccw_def) ultimately have "ccw ?l x y" by (rule ccw.transitive[where S=UNIV]) simp moreover { have "x \ ?l" using \?l \ x\ by simp moreover have "lex (?m y) (?m ?ml)" using mirror_y by (force intro: lex_half_segments_lowest_vertex in_set_butlastD) hence "lex ?ml y" by (rule lex_mirror_point) moreover from a have "lex ?ml x" unfolding mirror_eq_last by (rule lex_half_segments_last) moreover note \lex y ?l\ \lex x ?l\ \ccw' ?l x ?ml\ \ccw' ?l ?ml y\ ultimately have ncoll: "\ coll ?l x y" by (rule not_coll_ordered_lexI) } ultimately have ?ccw by (simp add: ccw_def) } ultimately show ?ccw by blast qed thus ?thesis using H by (simp add: segments_of_aform_def Let_def butlast_append snd_compose_pairself) qed lemma polychain_of_segments_of_aform1: assumes "length (segments_of_aform X) = 1" shows "False" using assms by (auto simp: segments_of_aform_def Let_def half_segments_of_aform_def add_is_1 split: if_split_asm) lemma polychain_of_segments_of_aform2: assumes "segments_of_aform X = [x, y]" assumes "between (fst x, snd x) p" shows "p \ convex hull set (map fst (segments_of_aform X))" proof - from polychain_segments_of_aform[of X] segments_of_aform_closed[of X] assms have "fst y = snd x" "snd y = fst x" by (simp_all add: polychain_def) thus ?thesis using assms by (cases x) (auto simp: between_mem_convex_hull) qed lemma append_eq_2: assumes "length xs = length ys" shows "xs @ ys = [x, y] \ xs = [x] \ ys = [y]" using assms proof (cases xs) case (Cons z zs) thus ?thesis using assms by (cases zs) auto qed simp lemma segments_of_aform_line_segment: assumes "segments_of_aform X = [x, y]" assumes "e \ UNIV \ {-1 .. 1}" shows "aform_val e X \ closed_segment (fst x) (snd x)" proof - from pdevs_val_pdevs_of_list_inl2E[OF \e \ _\, of "snd X"] obtain e' where e': "pdevs_val e (snd X) = pdevs_val e' (pdevs_of_list (inl (snd X)))" "e' \ UNIV \ {- 1..1}" . from e' have "0 \ 1 + e' 0" by (auto simp: Pi_iff dest!: spec[where x=0]) with assms e' show ?thesis by (auto simp: segments_of_aform_def Let_def append_eq_2 half_segments_of_aform_def polychain_of_singleton_iff mirror_point_def ccw.selsort_singleton_iff lowest_vertex_def aform_val_def sum_list_nlex_eq_sum_list_inl closed_segment_def Pi_iff intro!: exI[where x="(1 + e' 0) / 2"]) (auto simp: algebra_simps) qed subsection \Continuous Generalization\ lemma LIMSEQ_minus_fract_mult: "(\n. r * (1 - 1 / real (Suc (Suc n)))) \ r" by (rule tendsto_eq_rhs[OF tendsto_mult[where a=r and b = 1]]) (auto simp: inverse_eq_divide[symmetric] simp del: of_nat_Suc intro: filterlim_compose[OF LIMSEQ_inverse_real_of_nat filterlim_Suc] tendsto_eq_intros) lemma det3_nonneg_segments_of_aform: assumes "e \ UNIV \ {-1 .. 1}" assumes "length (half_segments_of_aform X) \ 1" shows "list_all (\seg. det3 (fst seg) (snd seg) (aform_val e X) \ 0) (segments_of_aform X)" unfolding list_all_iff proof safe fix a b c d assume seg: "((a, b), c, d) \ set (segments_of_aform X)" (is "?seg \ _") define normal_of_segment where "normal_of_segment = (\((a0, a1), b0, b1). (b1 - a1, a0 - b0)::real*real)" define support_of_segment where "support_of_segment = (\(a, b). normal_of_segment (a, b) \ a)" have "closed ((\x. x \ normal_of_segment ?seg) -` {..support_of_segment ?seg})" (is "closed ?cl") by (auto intro!: continuous_intros closed_vimage) moreover define f where "f n i = e i * ( 1 - 1 / (Suc (Suc n)))" for n i have "\n. aform_val (f n) X \ ?cl" proof fix n have "f n \ UNIV \ {-1 <..< 1}" using assms by (auto simp: f_def Pi_iff intro!: less_one_multI minus_one_less_multI) from list_allD[OF segments_of_aform_strict[OF this assms(2)] seg] show "aform_val (f n) X \ (\x. x \ normal_of_segment ((a, b), c, d)) -` {..support_of_segment ((a, b), c, d)}" by (auto simp: list_all_iff normal_of_segment_def support_of_segment_def det3_def' field_simps inner_prod_def ccw'_def) qed moreover have "\i. (\n. f n i) \ e i" unfolding f_def by (rule LIMSEQ_minus_fract_mult) hence "(\n. aform_val (f n) X) \ aform_val e X" by (auto simp: aform_val_def pdevs_val_sum intro!: tendsto_intros) ultimately have "aform_val e X \ ?cl" - by (rule closed_sequentially) + by (meson closed_sequentially) thus "det3 (fst ?seg) (snd ?seg) (aform_val e X) \ 0" by (auto simp: list_all_iff normal_of_segment_def support_of_segment_def det3_def' field_simps inner_prod_def) qed lemma det3_nonneg_segments_of_aformI: assumes "e \ UNIV \ {-1 .. 1}" assumes "length (half_segments_of_aform X) \ 1" assumes "seg \ set (segments_of_aform X)" shows "det3 (fst seg) (snd seg) (aform_val e X) \ 0" using assms det3_nonneg_segments_of_aform by (auto simp: list_all_iff) subsection \Intersection of Vertical Line with Segment\ fun intersect_segment_xline'::"nat \ point * point \ real \ point option" where "intersect_segment_xline' p ((x0, y0), (x1, y1)) xl = (if x0 \ xl \ xl \ x1 then if x0 = x1 then Some ((min y0 y1), (max y0 y1)) else let yl = truncate_down p (truncate_down p (real_divl p (y1 - y0) (x1 - x0) * (xl - x0)) + y0); yr = truncate_up p (truncate_up p (real_divr p (y1 - y0) (x1 - x0) * (xl - x0)) + y0) in Some (yl, yr) else None)" lemma norm_pair_fst0[simp]: "norm (0, x) = norm x" by (auto simp: norm_prod_def) lemmas add_right_mono_le = order_trans[OF add_right_mono] lemmas mult_right_mono_le = order_trans[OF mult_right_mono] lemmas add_right_mono_ge = order_trans[OF _ add_right_mono] lemmas mult_right_mono_ge = order_trans[OF _ mult_right_mono] lemma dest_segment: fixes x b::real assumes "(x, b) \ closed_segment (x0, y0) (x1, y1)" assumes "x0 \ x1" shows "b = (y1 - y0) * (x - x0) / (x1 - x0) + y0" proof - from assms obtain u where u: "x = x0 *\<^sub>R (1 - u) + u * x1" "b = y0 *\<^sub>R (1 - u) + u * y1" "0 \ u" "u \ 1" by (auto simp: closed_segment_def algebra_simps) show "b = (y1 - y0) * (x - x0) / (x1 - x0) + y0 " using assms by (auto simp: closed_segment_def field_simps u) qed lemma intersect_segment_xline': assumes "intersect_segment_xline' prec (p0, p1) x = Some (m, M)" shows "closed_segment p0 p1 \ {p. fst p = x} \ {(x, m) .. (x, M)}" using assms proof (cases p0) case (Pair x0 y0) note p0 = this show ?thesis proof (cases p1) case (Pair x1 y1) note p1 = this { assume "x0 = x1" "x = x1" "m = min y0 y1" "M = max y0 y1" hence ?thesis by (force simp: abs_le_iff p0 p1 min_def max_def algebra_simps dest: segment_bound split: if_split_asm) } thus ?thesis using assms by (auto simp: abs_le_iff p0 p1 split: if_split_asm intro!: truncate_up_le truncate_down_le add_right_mono_le[OF truncate_down] add_right_mono_le[OF real_divl] add_right_mono_le[OF mult_right_mono_le[OF real_divl]] add_right_mono_ge[OF _ truncate_up] add_right_mono_ge[OF _ mult_right_mono_ge[OF _ real_divr]] dest!: dest_segment) qed qed lemma in_segment_fst_le: fixes x0 x1 b::real assumes "x0 \ x1" "(x, b) \ closed_segment (x0, y0) (x1, y1)" shows "x \ x1" using assms using mult_left_mono[OF \x0 \ x1\, where c="1 - u" for u] by (force simp add: min_def max_def split: if_split_asm simp add: algebra_simps not_le closed_segment_def) lemma in_segment_fst_ge: fixes x0 x1 b::real assumes "x0 \ x1" "(x, b) \ closed_segment (x0, y0) (x1, y1)" shows "x0 \ x" using assms using mult_left_mono[OF \x0 \ x1\] by (force simp add: min_def max_def split: if_split_asm simp add: algebra_simps not_le closed_segment_def) lemma intersect_segment_xline'_eq_None: assumes "intersect_segment_xline' prec (p0, p1) x = None" assumes "fst p0 \ fst p1" shows "closed_segment p0 p1 \ {p. fst p = x} = {}" using assms by (cases p0, cases p1) (auto simp: abs_le_iff split: if_split_asm dest: in_segment_fst_le in_segment_fst_ge) fun intersect_segment_xline where "intersect_segment_xline prec ((a, b), (c, d)) x = (if a \ c then intersect_segment_xline' prec ((a, b), (c, d)) x else intersect_segment_xline' prec ((c, d), (a, b)) x)" lemma closed_segment_commute: "closed_segment a b = closed_segment b a" by (meson convex_contains_segment convex_closed_segment dual_order.antisym ends_in_segment) lemma intersect_segment_xline: assumes "intersect_segment_xline prec (p0, p1) x = Some (m, M)" shows "closed_segment p0 p1 \ {p. fst p = x} \ {(x, m) .. (x, M)}" using assms by (cases p0, cases p1) (auto simp: closed_segment_commute split: if_split_asm simp del: intersect_segment_xline'.simps dest!: intersect_segment_xline') lemma intersect_segment_xline_fst_snd: assumes "intersect_segment_xline prec seg x = Some (m, M)" shows "closed_segment (fst seg) (snd seg) \ {p. fst p = x} \ {(x, m) .. (x, M)}" using intersect_segment_xline[of prec "fst seg" "snd seg" x m M] assms by simp lemma intersect_segment_xline_eq_None: assumes "intersect_segment_xline prec (p0, p1) x = None" shows "closed_segment p0 p1 \ {p. fst p = x} = {}" using assms by (cases p0, cases p1) (auto simp: closed_segment_commute split: if_split_asm simp del: intersect_segment_xline'.simps dest!: intersect_segment_xline'_eq_None) lemma inter_image_empty_iff: "(X \ {p. f p = x} = {}) \ (x \ f ` X)" by auto lemma fst_closed_segment[simp]: "fst ` closed_segment a b = closed_segment (fst a) (fst b)" by (force simp: closed_segment_def) lemma intersect_segment_xline_eq_empty: fixes p0 p1::"real * real" assumes "closed_segment p0 p1 \ {p. fst p = x} = {}" shows "intersect_segment_xline prec (p0, p1) x = None" using assms by (cases p0, cases p1) (auto simp: inter_image_empty_iff closed_segment_eq_real_ivl split: if_split_asm) lemma intersect_segment_xline_le: assumes "intersect_segment_xline prec y xl = Some (m0, M0)" shows "m0 \ M0" using assms by (cases y) (auto simp: min_def split: if_split_asm intro!: truncate_up_le truncate_down_le order_trans[OF real_divl] order_trans[OF _ real_divr] mult_right_mono) lemma intersect_segment_xline_None_iff: fixes p0 p1::"real * real" shows "intersect_segment_xline prec (p0, p1) x = None \ closed_segment p0 p1 \ {p. fst p = x} = {}" by (auto intro!: intersect_segment_xline_eq_empty dest!: intersect_segment_xline_eq_None) subsection \Bounds on Vertical Intersection with Oriented List of Segments\ primrec bound_intersect_2d where "bound_intersect_2d prec [] x = None" | "bound_intersect_2d prec (X # Xs) xl = (case bound_intersect_2d prec Xs xl of None \ intersect_segment_xline prec X xl | Some (m, M) \ (case intersect_segment_xline prec X xl of None \ Some (m, M) | Some (m', M') \ Some (min m' m, max M' M)))" lemma bound_intersect_2d_eq_None: assumes "bound_intersect_2d prec Xs x = None" assumes "X \ set Xs" shows "intersect_segment_xline prec X x = None" using assms by (induct Xs) (auto split: option.split_asm) lemma bound_intersect_2d_upper: assumes "bound_intersect_2d prec Xs x = Some (m, M)" obtains X m' where "X \ set Xs" "intersect_segment_xline prec X x = Some (m', M)" "\X m' M' . X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ M' \ M" proof atomize_elim show "\X m'. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M) \ (\X m' M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ M' \ M)" using assms proof (induct Xs arbitrary: m M) case Nil thus ?case by (simp add: bound_intersect_2d_def) next case (Cons X Xs) show ?case proof (cases "bound_intersect_2d prec Xs x") case None thus ?thesis using Cons by (intro exI[where x=X] exI[where x=m]) (simp del: intersect_segment_xline.simps add: bound_intersect_2d_eq_None) next case (Some mM) note Some1=this then obtain m' M' where mM: "mM = (m', M')" by (cases mM) from Cons(1)[OF Some[unfolded mM]] obtain X' m'' where X': "X' \ set Xs" and m'': "intersect_segment_xline prec X' x = Some (m'', M')" and max: "\X m' M'a. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M'a) \ M'a \ M'" by auto show ?thesis proof (cases "intersect_segment_xline prec X x") case None thus ?thesis using Some1 mM Cons(2) X' m'' max by (intro exI[where x= X'] exI[where x= m'']) (auto simp del: intersect_segment_xline.simps split: option.split_asm) next case (Some mM''') thus ?thesis using Some1 mM Cons(2) X' m'' by (cases mM''') (force simp: max_def min_def simp del: intersect_segment_xline.simps split: option.split_asm if_split_asm dest!: max intro!: exI[where x= "if M' \ snd mM''' then X' else X"] exI[where x= "if M' \ snd mM''' then m'' else fst mM'''"]) qed qed qed qed lemma bound_intersect_2d_lower: assumes "bound_intersect_2d prec Xs x = Some (m, M)" obtains X M' where "X \ set Xs" "intersect_segment_xline prec X x = Some (m, M')" "\X m' M' . X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ m \ m'" proof atomize_elim show "\X M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m, M') \ (\X m' M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m', M') \ m \ m')" using assms proof (induct Xs arbitrary: m M) case Nil thus ?case by (simp add: bound_intersect_2d_def) next case (Cons X Xs) show ?case proof (cases "bound_intersect_2d prec Xs x") case None thus ?thesis using Cons by (intro exI[where x= X]) (simp del: intersect_segment_xline.simps add: bound_intersect_2d_eq_None) next case (Some mM) note Some1=this then obtain m' M' where mM: "mM = (m', M')" by (cases mM) from Cons(1)[OF Some[unfolded mM]] obtain X' M'' where X': "X' \ set Xs" and M'': "intersect_segment_xline prec X' x = Some (m', M'')" and min: "\X m'a M'. X \ set Xs \ intersect_segment_xline prec X x = Some (m'a, M') \ m' \ m'a" by auto show ?thesis proof (cases "intersect_segment_xline prec X x") case None thus ?thesis using Some1 mM Cons(2) X' M'' min by (intro exI[where x= X'] exI[where x= M'']) (auto simp del: intersect_segment_xline.simps split: option.split_asm) next case (Some mM''') thus ?thesis using Some1 mM Cons(2) X' M'' min by (cases mM''') (force simp: max_def min_def simp del: intersect_segment_xline.simps split: option.split_asm if_split_asm dest!: min intro!: exI[where x= "if m' \ fst mM''' then X' else X"] exI[where x= "if m' \ fst mM''' then M'' else snd mM'''"]) qed qed qed qed lemma bound_intersect_2d: assumes "bound_intersect_2d prec Xs x = Some (m, M)" shows "(\(p1, p2) \ set Xs. closed_segment p1 p2) \ {p. fst p = x} \ {(x, m) .. (x, M)}" proof (clarsimp, safe) fix b x0 y0 x1 y1 assume H: "((x0, y0), x1, y1) \ set Xs" "(x, b) \ closed_segment (x0, y0) (x1, y1)" hence "intersect_segment_xline prec ((x0, y0), x1, y1) x \ None" by (intro notI) (auto dest!: intersect_segment_xline_eq_None simp del: intersect_segment_xline.simps) then obtain e f where ef: "intersect_segment_xline prec ((x0, y0), x1, y1) x = Some (e, f)" by auto with H have "m \ e" by (auto intro: bound_intersect_2d_lower[OF assms]) also have "\ \ b" using intersect_segment_xline[OF ef] H by force finally show "m \ b" . have "b \ f" using intersect_segment_xline[OF ef] H by force also have "\ \ M" using H ef by (auto intro: bound_intersect_2d_upper[OF assms]) finally show "b \ M" . qed lemma bound_intersect_2d_eq_None_iff: "bound_intersect_2d prec Xs x = None \ (\X\set Xs. intersect_segment_xline prec X x = None)" by (induct Xs) (auto simp: split: option.split_asm) lemma bound_intersect_2d_nonempty: assumes "bound_intersect_2d prec Xs x = Some (m, M)" shows "(\(p1, p2) \ set Xs. closed_segment p1 p2) \ {p. fst p = x} \ {}" proof - from assms have "bound_intersect_2d prec Xs x \ None" by simp then obtain p1 p2 where "(p1, p2) \ set Xs" "intersect_segment_xline prec (p1, p2) x \ None" unfolding bound_intersect_2d_eq_None_iff by auto hence "closed_segment p1 p2 \ {p. fst p = x} \ {}" by (simp add: intersect_segment_xline_None_iff) thus ?thesis using \(p1, p2) \ set Xs\ by auto qed lemma bound_intersect_2d_le: assumes "bound_intersect_2d prec Xs x = Some (m, M)" shows "m \ M" proof - from bound_intersect_2d_nonempty[OF assms] bound_intersect_2d[OF assms] show "m \ M" by auto qed subsection \Bounds on Vertical Intersection with General List of Segments\ definition "bound_intersect_2d_ud prec X xl = (case segments_of_aform X of [] \ if fst (fst X) = xl then let m = snd (fst X) in Some (m, m) else None | [x, y] \ intersect_segment_xline prec x xl | xs \ (case bound_intersect_2d prec (filter (\((x1, y1), x2, y2). x1 < x2) xs) xl of Some (m, M') \ (case bound_intersect_2d prec (filter (\((x1, y1), x2, y2). x1 > x2) xs) xl of Some (m', M) \ Some (min m m', max M M') | None \ None) | None \ None))" lemma list_cases4: "\x P. (x = [] \ P) \ (\y. x = [y] \ P) \ (\y z. x = [y, z] \ P) \ (\w y z zs. x = w # y # z # zs \ P) \ P" by (metis list.exhaust) lemma bound_intersect_2d_ud_segments_of_aform_le: "bound_intersect_2d_ud prec X xl = Some (m0, M0) \ m0 \ M0" by (cases "segments_of_aform X" rule: list_cases4) (auto simp: Let_def bound_intersect_2d_ud_def min_def max_def intersect_segment_xline_le if_split_eq1 split: option.split_asm prod.split_asm list.split_asm dest!: bound_intersect_2d_le) lemma pdevs_domain_eq_empty_iff[simp]: "pdevs_domain (snd X) = {} \ snd X = zero_pdevs" by (auto simp: intro!: pdevs_eqI) lemma ccw_contr_on_line_left: assumes "det3 (a, b) (x, c) (x, d) \ 0" "a > x" shows "d \ c" proof - from assms have "d * (a - x) \ c * (a - x)" by (auto simp: det3_def' algebra_simps) with assms show "c \ d" by auto qed lemma ccw_contr_on_line_right: assumes "det3 (a, b) (x, c) (x, d) \ 0" "a < x" shows "d \ c" proof - from assms have "c * (x - a) \ d * (x - a)" by (auto simp: det3_def' algebra_simps) with assms show "d \ c" by auto qed lemma singleton_contrE: assumes "\x y. x \ y \ x \ X \ y \ X \ False" assumes "X \ {}" obtains x where "X = {x}" using assms by blast lemma segment_intersection_singleton: fixes xl and a b::"real * real" defines "i \ closed_segment a b \ {p. fst p = xl}" assumes ne1: "fst a \ fst b" assumes upper: "i \ {}" obtains p1 where "i = {p1}" proof (rule singleton_contrE[OF _ upper]) fix x y assume H: "x \ y" "x \ i" "y \ i" then obtain u v where uv: "x = u *\<^sub>R b + (1 - u) *\<^sub>R a" "y = v *\<^sub>R b + (1 - v) *\<^sub>R a" "0 \ u" "u \ 1" "0 \ v" "v \ 1" by (auto simp add: closed_segment_def i_def field_simps) then have "x + u *\<^sub>R a = a + u *\<^sub>R b" "y + v *\<^sub>R a = a + v *\<^sub>R b" by simp_all then have "fst (x + u *\<^sub>R a) = fst (a + u *\<^sub>R b)" "fst (y + v *\<^sub>R a) = fst (a + v *\<^sub>R b)" by simp_all then have "u = v * (fst a - fst b) / (fst a - fst b)" using ne1 H(2,3) \0 \ u\ \u \ 1\ \0 \ v\ \v \ 1\ by (simp add: closed_segment_def i_def field_simps) then have "u = v" by (simp add: ne1) then show False using H uv by simp qed lemma bound_intersect_2d_ud_segments_of_aform: assumes "bound_intersect_2d_ud prec X xl = Some (m0, M0)" assumes "e \ UNIV \ {-1 .. 1}" shows "{aform_val e X} \ {x. fst x = xl} \ {(xl, m0) .. (xl, M0)}" proof safe fix a b assume safeassms: "(a, b) = aform_val e X" "xl = fst (a, b)" hence fst_aform_val: "fst (aform_val e X) = xl" by simp { assume len: "length (segments_of_aform X) > 2" with assms obtain m M m' M' where lb: "bound_intersect_2d prec [((x1, y1), x2, y2)\segments_of_aform X . x1 < x2] xl = Some (m, M')" and ub: "bound_intersect_2d prec [((x1, y1), x2, y2)\segments_of_aform X . x2 < x1] xl = Some (m', M)" and minmax: "m0 = min m m'" "M0 = max M M'" by (auto simp: bound_intersect_2d_ud_def split: option.split_asm list.split_asm ) from bound_intersect_2d_upper[OF ub] obtain X1 m1 where upper: "X1 \ set [((x1, y1), x2, y2)\segments_of_aform X . x2 < x1]" "intersect_segment_xline prec X1 xl = Some (m1, M)" by metis from bound_intersect_2d_lower[OF lb] obtain X2 M2 where lower: "X2 \ set [((x1, y1), x2, y2)\segments_of_aform X . x1 < x2]" "intersect_segment_xline prec X2 xl = Some (m, M2)" by metis from upper(1) lower(1) have X1: "X1 \ set (segments_of_aform X)" "fst (fst X1) > fst (snd X1)" and X2: "X2 \ set (segments_of_aform X)" "fst (fst X2) < fst (snd X2)" by auto note upper_seg = intersect_segment_xline_fst_snd[OF upper(2)] note lower_seg = intersect_segment_xline_fst_snd[OF lower(2)] from len have lh: "length (half_segments_of_aform X) \ 1" by (auto simp: segments_of_aform_def Let_def) from X1 have ne1: "fst (fst X1) \ fst (snd X1)" by simp moreover have "closed_segment (fst X1) (snd X1) \ {p. fst p = xl} \ {}" using upper(2) by (simp add: intersect_segment_xline_None_iff[of prec, symmetric]) ultimately obtain p1 where p1: "closed_segment (fst X1) (snd X1) \ {p. fst p = xl} = {p1}" by (rule segment_intersection_singleton) then obtain u where u: "p1 = (1 - u) *\<^sub>R fst X1 + u *\<^sub>R (snd X1)" "0 \ u" "u \ 1" by (auto simp: closed_segment_def algebra_simps) have coll1: "det3 (fst X1) p1 (aform_val e X) \ 0" and coll1': "det3 p1 (snd X1) (aform_val e X) \ 0" unfolding atomize_conj using u by (cases "u = 0 \ u = 1") (auto simp: u(1) intro: det3_nonneg_scaleR_segment1 det3_nonneg_scaleR_segment2 det3_nonneg_segments_of_aformI[OF \e \ _\ lh X1(1)]) from X2 have ne2: "fst (fst X2) \ fst (snd X2)" by simp moreover have "closed_segment (fst X2) (snd X2) \ {p. fst p = xl} \ {}" using lower(2) by (simp add: intersect_segment_xline_None_iff[of prec, symmetric]) ultimately obtain p2 where p2: "closed_segment (fst X2) (snd X2) \ {p. fst p = xl} = {p2}" by (rule segment_intersection_singleton) then obtain v where v: "p2 = (1 - v) *\<^sub>R fst X2 + v *\<^sub>R (snd X2)" "0 \ v" "v \ 1" by (auto simp: closed_segment_def algebra_simps) have coll2: "det3 (fst X2) p2 (aform_val e X) \ 0" and coll2': "det3 p2 (snd X2) (aform_val e X) \ 0" unfolding atomize_conj using v by (cases "v = 0 \ v = 1") (auto simp: v(1) intro: det3_nonneg_scaleR_segment1 det3_nonneg_scaleR_segment2 det3_nonneg_segments_of_aformI[OF \e \ _\ lh X2(1)]) from in_set_segments_of_aform_aform_valE [of "fst X1" "snd X1" X, unfolded prod.collapse, OF X1(1)] obtain e1s where e1s: "snd X1 = aform_val e1s X" "e1s \ UNIV \ {- 1..1}" . from previous_segments_of_aformE [of "fst X1" "snd X1" X, unfolded prod.collapse, OF X1(1)] obtain fX0 where "(fX0, fst X1) \ set (segments_of_aform X)" . from in_set_segments_of_aform_aform_valE[OF this] obtain e1f where e1f: "fst X1 = aform_val e1f X" "e1f \ UNIV \ {- 1..1}" . have "p1 \ closed_segment (aform_val e1f X) (aform_val e1s X)" using p1 by (auto simp: e1s e1f) with segment_in_aform_val[OF e1s(2) e1f(2), of X] obtain ep1 where ep1: "ep1 \ UNIV \ {-1 .. 1}" "p1 = aform_val ep1 X" by (auto simp: Affine_def valuate_def closed_segment_commute) from in_set_segments_of_aform_aform_valE [of "fst X2" "snd X2" X, unfolded prod.collapse, OF X2(1)] obtain e2s where e2s: "snd X2 = aform_val e2s X" "e2s \ UNIV \ {- 1..1}" . from previous_segments_of_aformE [of "fst X2" "snd X2" X, unfolded prod.collapse, OF X2(1)] obtain fX02 where "(fX02, fst X2) \ set (segments_of_aform X)" . from in_set_segments_of_aform_aform_valE[OF this] obtain e2f where e2f: "fst X2 = aform_val e2f X" "e2f \ UNIV \ {- 1..1}" . have "p2 \ closed_segment (aform_val e2f X) (aform_val e2s X)" using p2 by (auto simp: e2s e2f) with segment_in_aform_val[OF e2f(2) e2s(2), of X] obtain ep2 where ep2: "ep2 \ UNIV \ {-1 .. 1}" "p2 = aform_val ep2 X" by (auto simp: Affine_def valuate_def) from det3_nonneg_segments_of_aformI[OF ep2(1), of X "(fst X1, snd X1)", unfolded prod.collapse, OF lh X1(1), unfolded ep2(2)[symmetric]] have c2: "det3 (fst X1) (snd X1) p2 \ 0" . hence c12: "det3 (fst X1) p1 p2 \ 0" using u by (cases "u = 0") (auto simp: u(1) det3_nonneg_scaleR_segment2) from det3_nonneg_segments_of_aformI[OF ep1(1), of X "(fst X2, snd X2)", unfolded prod.collapse, OF lh X2(1), unfolded ep1(2)[symmetric]] have c1: "det3 (fst X2) (snd X2) p1 \ 0" . hence c21: "det3 (fst X2) p2 p1 \ 0" using v by (cases "v = 0") (auto simp: v(1) det3_nonneg_scaleR_segment2) from p1 p2 have p1p2xl: "fst p1 = xl" "fst p2 = xl" by (auto simp: det3_def') from upper_seg p1 have "snd p1 \ M" by (auto simp: less_eq_prod_def) from lower_seg p2 have "m \ snd p2" by (auto simp: less_eq_prod_def) { have *: "(fst p1, snd (aform_val e X)) = aform_val e X" by (simp add: prod_eq_iff p1p2xl fst_aform_val) hence coll: "det3 (fst (fst X1), snd (fst X1)) (fst p1, snd p1) (fst p1, snd (aform_val e X)) \ 0" and coll': "det3 (fst (snd X1), snd (snd X1)) (fst p1, snd (aform_val e X)) (fst p1, snd p1) \ 0" using coll1 coll1' by (auto simp: det3_rotate) have "snd (aform_val e X) \ M" proof (cases "fst (fst X1) = xl") case False have "fst (fst X1) \ fst p1" unfolding u using X1 by (auto simp: algebra_simps intro!: mult_left_mono u) moreover have "fst (fst X1) \ fst p1" using False by (simp add: p1p2xl) ultimately have "fst (fst X1) > fst p1" by simp from ccw_contr_on_line_left[OF coll this] show ?thesis using \snd p1 \ M\ by simp next case True have "fst (snd X1) * (1 - u) \ fst (fst X1) * (1 - u)" using X1 u by (auto simp: intro!: mult_right_mono) hence "fst (snd X1) \ fst p1" unfolding u by (auto simp: algebra_simps) moreover have "fst (snd X1) \ fst p1" using True ne1 by (simp add: p1p2xl) ultimately have "fst (snd X1) < fst p1" by simp from ccw_contr_on_line_right[OF coll' this] show ?thesis using \snd p1 \ M\ by simp qed } moreover { have "(fst p2, snd (aform_val e X)) = aform_val e X" by (simp add: prod_eq_iff p1p2xl fst_aform_val) hence coll: "det3 (fst (fst X2), snd (fst X2)) (fst p2, snd p2) (fst p2, snd (aform_val e X)) \ 0" and coll': "det3 (fst (snd X2), snd (snd X2)) (fst p2, snd (aform_val e X)) (fst p2, snd p2) \ 0" using coll2 coll2' by (auto simp: det3_rotate) have "m \ snd (aform_val e X)" proof (cases "fst (fst X2) = xl") case False have "fst (fst X2) \ fst p2" unfolding v using X2 by (auto simp: algebra_simps intro!: mult_left_mono v) moreover have "fst (fst X2) \ fst p2" using False by (simp add: p1p2xl) ultimately have "fst (fst X2) < fst p2" by simp from ccw_contr_on_line_right[OF coll this] show ?thesis using \m \ snd p2\ by simp next case True have "(1 - v) * fst (snd X2) \ (1 - v) * fst (fst X2)" using X2 v by (auto simp: intro!: mult_left_mono) hence "fst (snd X2) \ fst p2" unfolding v by (auto simp: algebra_simps) moreover have "fst (snd X2) \ fst p2" using True ne2 by (simp add: p1p2xl) ultimately have "fst (snd X2) > fst p2" by simp from ccw_contr_on_line_left[OF coll' this] show ?thesis using \m \ snd p2\ by simp qed } ultimately have "aform_val e X \ {(xl, m) .. (xl, M)}" by (auto simp: less_eq_prod_def fst_aform_val) hence "aform_val e X \ {(xl, m0) .. (xl, M0)}" by (auto simp: minmax less_eq_prod_def) } moreover { assume "length (segments_of_aform X) = 2" then obtain a b where s: "segments_of_aform X = [a, b]" by (auto simp: numeral_2_eq_2 length_Suc_conv) from segments_of_aform_line_segment[OF this assms(2)] have "aform_val e X \ closed_segment (fst a) (snd a)" . moreover from assms have "intersect_segment_xline prec a xl = Some (m0, M0)" by (auto simp: bound_intersect_2d_ud_def s) note intersect_segment_xline_fst_snd[OF this] ultimately have "aform_val e X \ {(xl, m0) .. (xl, M0)}" by (auto simp: less_eq_prod_def fst_aform_val) } moreover { assume "length (segments_of_aform X) = 1" from polychain_of_segments_of_aform1[OF this] have "aform_val e X \ {(xl, m0) .. (xl, M0)}" by auto } moreover { assume len: "length (segments_of_aform X) = 0" hence "independent_pdevs (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = []" by (simp add: segments_of_aform_def Let_def half_segments_of_aform_def inl_def) hence "snd X = zero_pdevs" by (subst (asm) independent_pdevs_eq_Nil_iff) (auto simp: list_all_iff list_of_pdevs_def) hence "aform_val e X = fst X" by (simp add: aform_val_def) with len assms have "aform_val e X \ {(xl, m0) .. (xl, M0)}" by (auto simp: bound_intersect_2d_ud_def Let_def split: if_split_asm) } ultimately have "aform_val e X \ {(xl, m0)..(xl, M0)}" by arith thus "(a, b) \ {(fst (a, b), m0)..(fst (a, b), M0)}" using safeassms by simp qed subsection \Approximation from Orthogonal Directions\ definition inter_aform_plane_ortho:: "nat \ 'a::executable_euclidean_space aform \ 'a \ real \ 'a aform option" where "inter_aform_plane_ortho p Z n g = do { mMs \ those (map (\b. bound_intersect_2d_ud p (inner2_aform Z n b) g) Basis_list); let l = (\(b, m)\zip Basis_list (map fst mMs). m *\<^sub>R b); let u = (\(b, M)\zip Basis_list (map snd mMs). M *\<^sub>R b); Some (aform_of_ivl l u) }" lemma those_eq_SomeD: assumes "those (map f xs) = Some ys" shows "ys = map (the o f) xs \ (\i.\y. i < length xs \ f (xs ! i) = Some y)" using assms by (induct xs arbitrary: ys) (auto split: option.split_asm simp: o_def nth_Cons split: nat.split) lemma sum_list_zip_map: assumes "distinct xs" shows "(\(x, y)\zip xs (map g xs). f x y) = (\x\set xs. f x (g x))" by (force simp add: sum_list_distinct_conv_sum_set assms distinct_zipI1 split_beta' in_set_zip in_set_conv_nth inj_on_convol_ident intro!: sum.reindex_cong[where l="\x. (x, g x)"]) lemma inter_aform_plane_ortho_overappr: assumes "inter_aform_plane_ortho p Z n g = Some X" shows "{x. \i \ Basis. x \ i \ {y. (g, y) \ (\x. (x \ n, x \ i)) ` Affine Z}} \ Affine X" proof - let ?inter = "(\b. bound_intersect_2d_ud p (inner2_aform Z n b) g)" obtain xs where xs: "those (map ?inter Basis_list) = Some xs" using assms by (cases "those (map ?inter Basis_list)") (auto simp: inter_aform_plane_ortho_def) from those_eq_SomeD[OF this] obtain y' where xs_eq: "xs = map (the \ ?inter) Basis_list" and y': "\i. i < length (Basis_list::'a list) \ ?inter (Basis_list ! i) = Some (y' i)" by metis have "\(i::'a) \ Basis. \ji::'a. i\Basis \ j i < length (Basis_list::'a list)" "\i::'a. i\Basis \ i = Basis_list ! j i" by metis define y where "y = y' o j" with y' j have y: "\i. i \ Basis \ ?inter i = Some (y i)" by (metis comp_def) hence y_le: "\i. i \ Basis \ fst (y i) \ snd (y i)" by (auto intro!: bound_intersect_2d_ud_segments_of_aform_le) hence "(\b\Basis. fst (y b) *\<^sub>R b) \ (\b\Basis. snd (y b) *\<^sub>R b)" by (auto simp: eucl_le[where 'a='a]) with assms have X: "Affine X = {\b\Basis. fst (y b) *\<^sub>R b..\b\Basis. snd (y b) *\<^sub>R b}" by (auto simp: inter_aform_plane_ortho_def sum_list_zip_map xs xs_eq y Affine_aform_of_ivl) show ?thesis proof safe fix x assume x: "\i\Basis. x \ i \ {y. (g, y) \ (\x. (x \ n, x \ i)) ` Affine Z}" { fix i::'a assume i: "i \ Basis" from x i have x_in2: "(g, x \ i) \ (\x. (x \ n, x \ i)) ` Affine Z" by auto from x_in2 obtain e where e: "e \ UNIV \ {- 1..1}" and g: "g = aform_val e Z \ n" and x: "x \ i = aform_val e Z \ i" by (auto simp: Affine_def valuate_def) have "{aform_val e (inner2_aform Z n i)} = {aform_val e (inner2_aform Z n i)} \ {x. fst x = g}" by (auto simp: g inner2_def) also from y[OF \i \ Basis\] have "?inter i = Some (fst (y i), snd (y i))" by simp note bound_intersect_2d_ud_segments_of_aform[OF this e] finally have "x \ i \ {fst (y i) .. snd (y i)}" by (auto simp: x inner2_def) } thus "x \ Affine X" unfolding X by (auto simp: eucl_le[where 'a='a]) qed qed lemma inter_proj_eq: fixes n g l defines "G \ {x. x \ n = g}" shows "(\x. x \ l) ` (Z \ G) = {y. (g, y) \ (\x. (x \ n, x \ l)) ` Z}" by (auto simp: G_def) lemma inter_overappr: fixes n \ l shows "Z \ {x. x \ n = g} \ {x. \i \ Basis. x \ i \ {y. (g, y) \ (\x. (x \ n, x \ i)) ` Z}}" by auto lemma inter_inter_aform_plane_ortho: assumes "inter_aform_plane_ortho p Z n g = Some X" shows "Affine Z \ {x. x \ n = g} \ Affine X" proof - note inter_overappr[of "Affine Z" n g] also note inter_aform_plane_ortho_overappr[OF assms] finally show ?thesis . qed subsection \``Completeness'' of Intersection\ abbreviation "encompasses x seg \ det3 (fst seg) (snd seg) x \ 0" lemma encompasses_cases: "encompasses x seg \ encompasses x (snd seg, fst seg)" by (auto simp: det3_def' algebra_simps) lemma list_all_encompasses_cases: assumes "list_all (encompasses p) (x # y # zs)" obtains "list_all (encompasses p) [x, y, (snd y, fst x)]" | "list_all (encompasses p) ((fst x, snd y)#zs)" using encompasses_cases proof assume "encompasses p (snd y, fst x)" hence "list_all (encompasses p) [x, y, (snd y, fst x)]" using assms by (auto simp: list_all_iff) thus ?thesis .. next assume "encompasses p (snd (snd y, fst x), fst (snd y, fst x))" hence "list_all (encompasses p) ((fst x, snd y)#zs)" using assms by (auto simp: list_all_iff) thus ?thesis .. qed lemma triangle_encompassing_polychain_of: assumes "det3 p a b \ 0" "det3 p b c \ 0" "det3 p c a \ 0" assumes "ccw' a b c" shows "p \ convex hull {a, b, c}" proof - from assms have nn: "det3 b c p \ 0" "det3 c a p \ 0" "det3 a b p \ 0" "det3 a b c \ 0" by (auto simp: det3_def' algebra_simps) have "det3 a b c *\<^sub>R p = det3 b c p *\<^sub>R a + det3 c a p *\<^sub>R b + det3 a b p *\<^sub>R c" by (auto simp: det3_def' algebra_simps prod_eq_iff) hence "inverse (det3 a b c) *\<^sub>R (det3 a b c *\<^sub>R p) = inverse (det3 a b c) *\<^sub>R (det3 b c p *\<^sub>R a + det3 c a p *\<^sub>R b + det3 a b p *\<^sub>R c)" by simp with assms have p_eq: "p = (det3 b c p / det3 a b c) *\<^sub>R a + (det3 c a p / det3 a b c) *\<^sub>R b + (det3 a b p / det3 a b c) *\<^sub>R c" (is "_ = scaleR ?u _ + scaleR ?v _ + scaleR ?w _") by (simp add: inverse_eq_divide algebra_simps ccw'_def) have det_eq: "det3 b c p / det3 a b c + det3 c a p / det3 a b c + det3 a b p / det3 a b c = 1" using assms(4) by (simp add: add_divide_distrib[symmetric] det3_def' algebra_simps ccw'_def) show ?thesis unfolding convex_hull_3 using assms(4) by (blast intro: exI[where x="?u"] exI[where x="?v"] exI[where x="?w"] intro!: p_eq divide_nonneg_nonneg nn det_eq) qed lemma encompasses_convex_polygon3: assumes "list_all (encompasses p) (x#y#z#zs)" assumes "convex_polygon (x#y#z#zs)" assumes "ccw'.sortedP (fst x) (map snd (butlast (x#y#z#zs)))" shows "p \ convex hull (set (map fst (x#y#z#zs)))" using assms proof (induct zs arbitrary: x y z p) case Nil thus ?case by (auto simp: det3_def' algebra_simps elim!: ccw'.sortedP_Cons ccw'.sortedP_Nil intro!: triangle_encompassing_polychain_of) next case (Cons w ws) from Cons.prems(2) have "snd y = fst z" by auto from Cons.prems(1) show ?case proof (rule list_all_encompasses_cases) assume "list_all (encompasses p) [x, y, (snd y, fst x)]" hence "p \ convex hull {fst x, fst y, snd y}" using Cons.prems by (auto simp: det3_def' algebra_simps elim!: ccw'.sortedP_Cons ccw'.sortedP_Nil intro!: triangle_encompassing_polychain_of) thus ?case by (rule rev_subsetD[OF _ hull_mono]) (auto simp: \snd y = fst z\) next assume *: "list_all (encompasses p) ((fst x, snd y) # z # w # ws)" from Cons.prems have enc: "ws \ [] \ encompasses p (last ws)" by (auto simp: list_all_iff) have "set (map fst ((fst x, snd y)#z#w#ws)) \ set (map fst (x # y # z # w # ws))" by auto moreover { note * moreover have "convex_polygon ((fst x, snd y) # z # w # ws)" by (metis convex_polygon_skip Cons.prems(2,3)) moreover have "ccw'.sortedP (fst (fst x, snd y)) (map snd (butlast ((fst x, snd y) # z # w # ws)))" using Cons.prems(3) by (auto elim!: ccw'.sortedP_Cons intro!: ccw'.sortedP.Cons ccw'.sortedP.Nil split: if_split_asm) ultimately have "p \ convex hull set (map fst ((fst x, snd y)#z#w#ws))" by (rule Cons.hyps) } ultimately show "p \ convex hull set (map fst (x # y # z # w # ws))" by (rule subsetD[OF hull_mono]) qed qed lemma segments_of_aform_empty_Affine_eq: assumes "segments_of_aform X = []" shows "Affine X = {fst X}" proof - have "independent_pdevs (map snd (list_of_pdevs (nlex_pdevs (snd X)))) = [] \ (list_of_pdevs (nlex_pdevs (snd X))) = []" by (subst independent_pdevs_eq_Nil_iff) (auto simp: list_all_iff list_of_pdevs_def ) with assms show ?thesis by (force simp: aform_val_def list_of_pdevs_def Affine_def valuate_def segments_of_aform_def Let_def half_segments_of_aform_def inl_def) qed lemma not_segments_of_aform_singleton: "segments_of_aform X \ [x]" by (auto simp: segments_of_aform_def Let_def add_is_1 dest!: arg_cong[where f=length]) lemma encompasses_segments_of_aform_in_AffineI: assumes "length (segments_of_aform X) > 2" assumes "list_all (encompasses p) (segments_of_aform X)" shows "p \ Affine X" proof - from assms(1) obtain x y z zs where eq: "segments_of_aform X = x#y#z#zs" by (cases "segments_of_aform X" rule: list_cases4) auto hence "fst x = fst (hd (half_segments_of_aform X))" by (metis segments_of_aform_def hd_append list.map_disc_iff list.sel(1)) also have "\ = lowest_vertex (fst X, nlex_pdevs (snd X))" using assms by (intro fst_hd_half_segments_of_aform) (auto simp: segments_of_aform_def) finally have fstx: "fst x = lowest_vertex (fst X, nlex_pdevs (snd X))" . have "p \ convex hull (set (map fst (segments_of_aform X)))" using assms(2) unfolding eq proof (rule encompasses_convex_polygon3) show "convex_polygon (x # y # z # zs)" using assms(1) unfolding eq[symmetric] by (intro convex_polygon_segments_of_aform) (simp add: segments_of_aform_def Let_def) show "ccw'.sortedP (fst x) (map snd (butlast (x # y # z # zs)))" using assms(1) unfolding fstx map_butlast eq[symmetric] by (intro ccw'_sortedP_snd_segments_of_aform) (simp add: segments_of_aform_def Let_def half_segments_of_aform_def) qed also have "\ \ convex hull (Affine X)" proof (rule hull_mono, safe) fix a b assume "(a, b) \ set (map fst (segments_of_aform X))" then obtain c d where "((a, b), c, d) \ set (segments_of_aform X)" by auto from previous_segments_of_aformE[OF this] obtain x where "(x, a, b) \ set (segments_of_aform X)" by auto from in_set_segments_of_aform_aform_valE[OF this] obtain e where "(a, b) = aform_val e X" "e \ UNIV \ {- 1..1}" by auto thus "(a, b) \ Affine X" by (auto simp: Affine_def valuate_def image_iff) qed also have "\ = Affine X" by (simp add: convex_Affine convex_hull_eq) finally show ?thesis . qed end diff --git a/thys/Green/Derivs.thy b/thys/Green/Derivs.thy --- a/thys/Green/Derivs.thy +++ b/thys/Green/Derivs.thy @@ -1,651 +1,649 @@ theory Derivs imports General_Utils begin lemma field_simp_has_vector_derivative [derivative_intros]: "(f has_field_derivative y) F \ (f has_vector_derivative y) F" by (simp add: has_real_derivative_iff_has_vector_derivative) lemma continuous_on_cases_empty [continuous_intros]: "\closed S; continuous_on S f; \x. \x \ S; \ P x\ \ f x = g x\ \ continuous_on S (\x. if P x then f x else g x)" using continuous_on_cases [of _ "{}"] by force lemma inj_on_cases: assumes "inj_on f (Collect P \ S)" "inj_on g (Collect (Not \ P) \ S)" "f ` (Collect P \ S) \ g ` (Collect (Not \ P) \ S) = {}" shows "inj_on (\x. if P x then f x else g x) S" using assms by (force simp: inj_on_def) lemma inj_on_arccos: "S \ {-1..1} \ inj_on arccos S" by (metis atLeastAtMost_iff cos_arccos inj_onI subsetCE) lemma has_vector_derivative_componentwise_within: "(f has_vector_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_vector_derivative (f' \ i)) (at a within S))" apply (simp add: has_vector_derivative_def) apply (subst has_derivative_componentwise_within) apply simp done lemma has_vector_derivative_pair_within: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "\u. u \ Basis \ ((\x. f x \ u) has_vector_derivative f' \ u) (at x within S)" "\u. u \ Basis \ ((\x. g x \ u) has_vector_derivative g' \ u) (at x within S)" shows "((\x. (f x, g x)) has_vector_derivative (f',g')) (at x within S)" apply (subst has_vector_derivative_componentwise_within) apply (auto simp: assms Basis_prod_def) done lemma piecewise_C1_differentiable_const: shows "(\x. c) piecewise_C1_differentiable_on s" using continuous_on_const by (auto simp add: piecewise_C1_differentiable_on_def) declare piecewise_C1_differentiable_const [simp, derivative_intros] declare piecewise_C1_differentiable_neg [simp, derivative_intros] declare piecewise_C1_differentiable_add [simp, derivative_intros] declare piecewise_C1_differentiable_diff [simp, derivative_intros] (*move to Derivative*) lemma piecewise_C1_differentiable_on_ident [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_vector" shows "(\x. x) piecewise_C1_differentiable_on S" unfolding piecewise_C1_differentiable_on_def using C1_differentiable_on_ident by (blast intro: continuous_on_id C1_differentiable_on_ident) lemma piecewise_C1_differentiable_on_mult [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_algebra" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. f x * g x) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def apply safe apply (blast intro: continuous_intros) apply (rename_tac A B) apply (rule_tac x="A \ B" in exI) apply (auto intro: C1_differentiable_on_mult C1_differentiable_on_subset) done lemma C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a :: real_normed_field" shows "f C1_differentiable_on S \ (\x. f x / c) C1_differentiable_on S" by (simp add: divide_inverse) lemma piecewise_C1_differentiable_on_cdiv [simp, derivative_intros]: fixes f :: "real \ 'a::real_normed_field" assumes "f piecewise_C1_differentiable_on S" shows "(\x. f x / c) piecewise_C1_differentiable_on S" by (simp add: divide_inverse piecewise_C1_differentiable_const piecewise_C1_differentiable_on_mult assms) lemma sqrt_C1_differentiable [simp, derivative_intros]: assumes f: "f C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) show ?thesis using assms unfolding C1_differentiable_on_def has_real_derivative_iff_has_vector_derivative [symmetric] by (fastforce intro!: contf continuous_intros derivative_intros) qed lemma sqrt_piecewise_C1_differentiable [simp, derivative_intros]: assumes f: "f piecewise_C1_differentiable_on S" and fim: "f ` S \ {0<..}" shows "(\x. sqrt (f x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (fastforce intro!: continuous_intros derivative_intros) lemma fixes f :: "real \ 'a::{banach,real_normed_field}" assumes f: "f C1_differentiable_on S" shows sin_C1_differentiable [simp, derivative_intros]: "(\x. sin (f x)) C1_differentiable_on S" and cos_C1_differentiable [simp, derivative_intros]: "(\x. cos (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df_sin = field_vector_diff_chain_at [where g=sin, unfolded o_def] note df_cos = field_vector_diff_chain_at [where g=cos, unfolded o_def] show "(\x. sin (f x)) C1_differentiable_on S" "(\x. cos (f x)) C1_differentiable_on S" using assms unfolding C1_differentiable_on_def has_real_derivative_iff_has_vector_derivative [symmetric] apply auto by (rule contf continuous_intros df_sin df_cos derivative_intros exI conjI ballI | force)+ qed lemma has_derivative_abs: fixes a::real assumes "a \ 0" shows "(abs has_derivative ((*) (sgn a))) (at a)" proof - have [simp]: "norm = abs" using real_norm_def by force show ?thesis using has_derivative_norm [where 'a=real, simplified] assms by (simp add: mult_commute_abs) qed lemma abs_C1_differentiable [simp, derivative_intros]: fixes f :: "real \ real" assumes f: "f C1_differentiable_on S" and "0 \ f ` S" shows "(\x. abs (f x)) C1_differentiable_on S" proof - have contf: "continuous_on S f" by (simp add: C1_differentiable_imp_continuous_on f) note df = DERIV_chain [where f=abs and g=f, unfolded o_def] show ?thesis using assms unfolding C1_differentiable_on_def has_real_derivative_iff_has_vector_derivative [symmetric] apply clarify apply (rule df exI conjI ballI)+ apply (force simp: has_field_derivative_def intro: has_derivative_abs continuous_intros contf)+ done qed lemma C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f C1_differentiable_on S" "g C1_differentiable_on S" shows "(\x. (f x, g x)) C1_differentiable_on S" using assms unfolding C1_differentiable_on_def apply safe apply (rename_tac A B) apply (intro exI ballI conjI) apply (rule_tac f'="A x" and g'="B x" in has_vector_derivative_pair_within) using has_vector_derivative_componentwise_within by (blast intro: continuous_on_Pair)+ lemma piecewise_C1_differentiable_on_pair [simp, derivative_intros]: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ 'b::euclidean_space" assumes "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on S" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on S" using assms unfolding piecewise_C1_differentiable_on_def by (blast intro!: continuous_intros C1_differentiable_on_pair del: continuous_on_discrete intro: C1_differentiable_on_subset) lemma test2: assumes s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" and "x \ {0..1}" "x \ (\t. (v-u) *\<^sub>R t + u) -` s" shows "vector_derivative (\x. g ((v-u) * x + u)) (at x within {0..1}) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u) within{0..1})" proof - have i:"(g has_vector_derivative vector_derivative g (at ((v - u) * x + u))) (at ((v-u) * x + u))" using assms s [of "(v - u) * x + u"] uv mult_left_le [of x "v-u"] by (auto simp: vector_derivative_works) have ii:"((\x. g ((v - u) * x + u)) has_vector_derivative (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))) (at x)" by (intro vector_diff_chain_at [simplified o_def] derivative_eq_intros | simp add: i)+ have 0: "0 \ (v - u) * x + u" using assms uv by auto have 1: "(v - u) * x + u \ 1" using assms uv by simp (metis add.commute atLeastAtMost_iff atLeastatMost_empty_iff diff_ge_0_iff_ge empty_iff le_diff_eq mult_left_le) have iii: "vector_derivative g (at ((v - u) * x + u) within {0..1}) = vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF i, of "0" "1", OF 0 1] by auto have iv: "vector_derivative (\x. g ((v - u) * x + u)) (at x within {0..1}) = (v - u) *\<^sub>R vector_derivative g (at ((v - u) * x + u))" using Derivative.vector_derivative_at_within_ivl[OF ii, of "0" "1"] assms by auto show ?thesis using iii iv by auto qed lemma C1_differentiable_on_components: assumes "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s" shows "f C1_differentiable_on s" proof (clarsimp simp add: C1_differentiable_on_def has_vector_derivative_def) have *:"\f i x. x *\<^sub>R (f \ i) = (x *\<^sub>R f) \ i" by auto have "\f'. \i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" using assms lambda_skolem_euclidean[of "\i D. (\x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D"] apply (simp only: C1_differentiable_on_def has_vector_derivative_def *) using continuous_on_componentwise[of "s"] by metis then obtain f' where f': "\i\Basis. \x\s. ((\x. f x \ i) has_derivative (\z. z *\<^sub>R f' x \ i)) (at x) \ continuous_on s f'" by auto then have 0: "(\x\s. (f has_derivative (\z. z *\<^sub>R f' x)) (at x)) \ continuous_on s f'" using f' has_derivative_componentwise_within[of "f", where S= UNIV] by auto then show "\D. (\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x)) \ continuous_on s D" by metis qed lemma piecewise_C1_differentiable_on_components: assumes "finite t" "\i. i \ Basis \ (\x. f x \ i) C1_differentiable_on s - t" "\i. i \ Basis \ continuous_on s (\x. f x \ i)" shows "f piecewise_C1_differentiable_on s" using C1_differentiable_on_components assms continuous_on_componentwise piecewise_C1_differentiable_on_def by blast lemma all_components_smooth_one_pw_smooth_is_pw_smooth: assumes "\i. i \ Basis - {j} \ (\x. f x \ i) C1_differentiable_on s" assumes "(\x. f x \ j) piecewise_C1_differentiable_on s" shows "f piecewise_C1_differentiable_on s" proof - have is_cont: "\i\Basis. continuous_on s (\x. f x \ i)" using assms C1_differentiable_imp_continuous_on piecewise_C1_differentiable_on_def by fastforce obtain t where t:"(finite t \ (\x. f x \ j) C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "f"] using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t is_cont by fastforce qed lemma derivative_component_fun_component: fixes i::"'a::euclidean_space" assumes "f differentiable (at x)" shows "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" proof - have "((\x. f x \ i) has_vector_derivative vector_derivative f (at x) \ i) (at x)" using assms and bounded_linear.has_vector_derivative[of "(\x. x \ i)" "f" "(vector_derivative f (at x))" "(at x)"] and bounded_linear_inner_left[of "i"] and vector_derivative_works[of "f" "(at x)"] by blast then show "((vector_derivative f (at x)) \ i) = ((vector_derivative (\x. (f x) \ i) (at x)) )" using vector_derivative_works[of "(\x. (f x) \ i)" "(at x)"] and differentiableI_vector[of "(\x. (f x) \ i)" "(vector_derivative f (at x) \ i)" "(at x)"] and Derivative.vector_derivative_at by force qed lemma gamma_deriv_at_within: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b}" and gamma_differentiable: "\x \ {a .. b}. \ differentiable at x" shows "vector_derivative \ (at x within {a..b}) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b by (auto simp add: vector_derivative_works) lemma islimpt_diff_finite: assumes "finite (t::'a::t1_space set)" shows "x islimpt s - t = x islimpt s" proof- have iii: "s - t = s - (t \ s)" by auto have "(t \ s) \ s" by auto have ii:"finite (t \ s)" using assms(1) by auto have i: "(t \ s) \ (s - (t \ s)) = ( s)" using assms by auto then have "x islimpt s - (t \ s) = x islimpt s" - using islimpt_Un_finite[OF ii,where ?t = "s - (t \ s)"] i assms(1) - by force + by (metis ii islimpt_Un_finite) then show ?thesis using iii by auto qed lemma ivl_limpt_diff: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x islimpt {a..b} - s" proof- have "x islimpt {a..b}" proof (cases "x \{a,b}") have i: "finite {a,b}" and ii: "{a, b} \ {a<..{a,b}" then show ?thesis - using islimpt_Un_finite[OF i, where ?t= "{a<..{a,b}" then show "x islimpt {a..b}" using assms by auto qed then show "x islimpt {a..b} - s" using islimpt_diff_finite[OF assms(1)] assms by fastforce qed lemma ivl_closure_diff_del: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "x \ closure (({a..b} - s) - {x})" using ivl_limpt_diff islimpt_in_closure assms by blast lemma ivl_not_trivial_limit_within: assumes "finite s" "a < b" "(x::real) \ {a..b} - s" shows "at x within {a..b} - s \ bot" using assms ivl_closure_diff_del not_trivial_limit_within by blast lemma vector_derivative_at_within_non_trivial_limit: "at x within s \ bot \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within s) = f'" using has_vector_derivative_at_within vector_derivative_within by fastforce lemma vector_derivative_at_within_ivl_diff: "finite s \ a < b \ (x::real) \ {a..b} - s \ (f has_vector_derivative f') (at x) \ vector_derivative f (at x within {a..b} - s) = f'" using vector_derivative_at_within_non_trivial_limit ivl_not_trivial_limit_within by fastforce lemma gamma_deriv_at_within_diff: assumes a_leq_b: "a < b" and x_within_bounds: "x \ {a..b} - s" and gamma_differentiable: "\x \ {a .. b} - s. \ differentiable at x" and s_subset: "s \ {a..b}" and finite_s: "finite s" shows "vector_derivative \ (at x within {a..b} - s) = vector_derivative \ (at x)" using vector_derivative_at_within_ivl_diff [of "s" "a" "b" "x" "\" "vector_derivative \ (at x)"] gamma_differentiable x_within_bounds a_leq_b s_subset finite_s by (auto simp add: vector_derivative_works) lemma gamma_deriv_at_within_gen: assumes a_leq_b: "a < b" and x_within_bounds: "x \ s" and s_subset: "s \ {a..b}" and gamma_differentiable: "\x \ s. \ differentiable at x" shows "vector_derivative \ (at x within ({a..b})) = vector_derivative \ (at x)" using Derivative.vector_derivative_at_within_ivl[of "\" "vector_derivative \ (at x)" "x" "a" "b"] gamma_differentiable x_within_bounds a_leq_b s_subset by (auto simp add: vector_derivative_works) lemma derivative_component_fun_component_at_within_gen: assumes gamma_differentiable: "\x \ s. \ differentiable at x" and s_subset: "s \ {0..1}" shows "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ s. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ s. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ s" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within_gen[of "0" "1"] x_within_bounds gamma_differentiable s_subset by (auto simp add: vector_derivative_works) then have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_deriv_at_within_gen[of "0" "1", where \ = "(\x. \ x \ i)"] x_within_bounds gamma_i_component_smooth s_subset by (auto simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma derivative_component_fun_component_at_within: assumes gamma_differentiable: "\x \ {0 .. 1}. \ differentiable at x" shows "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ (i::'a:: euclidean_space) = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof - have gamma_i_component_smooth: "\x \ {0 .. 1}. (\x. \ x \ i) differentiable at x" using gamma_differentiable by auto show "\x \ {0..1}. vector_derivative (\x. \ x) (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" proof fix x::real assume x_within_bounds: "x \ {0..1}" have gamma_deriv_at_within: "vector_derivative (\x. \ x) (at x within {0..1}) = vector_derivative (\x. \ x) (at x)" using gamma_deriv_at_within[of "0" "1"] x_within_bounds gamma_differentiable by (auto simp add: vector_derivative_works) have gamma_component_deriv_at_within: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x \ i) (at x within {0..1})" using Derivative.vector_derivative_at_within_ivl[of "(\x. (\ x) \ i)" "vector_derivative (\x. (\ x) \ i) (at x)" "x" "0" "1"] has_vector_derivative_at_within[of "(\x. \ x \ i)" "vector_derivative (\x. \ x \ i) (at x)" "x" "{0..1}"] gamma_i_component_smooth x_within_bounds by (simp add: vector_derivative_works) have gamma_component_deriv_eq_gamma_deriv_component: "vector_derivative (\x. \ x \ i) (at x) = vector_derivative (\x. \ x) (at x) \ i" using derivative_component_fun_component[of "\" "x" "i"] gamma_differentiable x_within_bounds by auto show "vector_derivative \ (at x within {0..1}) \ i = vector_derivative (\x. \ x \ i) (at x within {0..1})" using gamma_component_deriv_eq_gamma_deriv_component gamma_component_deriv_at_within gamma_deriv_at_within by auto qed qed lemma straight_path_diffrentiable_x: fixes b :: "real" and y1 ::"real" assumes gamma_def: "\ = (\x. (b, y2 + y1 * x))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma straight_path_diffrentiable_y: fixes b :: "real" and y1 y2 ::"real" assumes gamma_def: "\ = (\x. (y2 + y1 * x , b))" shows "\x. \ differentiable at x" unfolding gamma_def differentiable_def by (fast intro!: derivative_intros) lemma piecewise_C1_differentiable_on_imp_continuous_on: assumes "f piecewise_C1_differentiable_on s" shows "continuous_on s f" using assms by (auto simp add: piecewise_C1_differentiable_on_def) lemma boring_lemma1: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (f x, 0)) has_vector_derivative ((D,0::real))) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D *\<^sub>R (1,0))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(1,0)"] by auto have "((\x. (f x) *\<^sub>R (1,0)) has_vector_derivative (D,0)) (at x)" proof - have "(D, 0::'a) = D *\<^sub>R (1, 0)" by simp then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma boring_lemma2: fixes f :: "real\real" assumes "(f has_vector_derivative D) (at x)" shows "((\x. (0, f x)) has_vector_derivative (0, D)) (at x)" proof- have *: "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative (D *\<^sub>R (0,1))) (at x)" using bounded_linear.has_vector_derivative[OF Real_Vector_Spaces.bounded_linear_scaleR_left assms(1), of "(0,1)"] by auto then have "((\x. (f x) *\<^sub>R (0,1)) has_vector_derivative ((0,D))) (at x)" using scaleR_Pair Real_Vector_Spaces.real_scaleR_def proof - have "(0::'b, D) = D *\<^sub>R (0, 1)" by auto then show ?thesis by (metis (no_types) *) qed then show ?thesis by auto qed lemma pair_prod_smooth_pw_smooth: assumes "(f::real\real) C1_differentiable_on s" "(g::real\real) piecewise_C1_differentiable_on s" shows "(\x. (f x, g x)) piecewise_C1_differentiable_on s" proof - have f_cont: "continuous_on s f" using assms(1) C1_differentiable_imp_continuous_on by fastforce have g_cont: "continuous_on s g" using assms(2) by (auto simp add: piecewise_C1_differentiable_on_def) obtain t where t:"(finite t \ g C1_differentiable_on s - t)" using assms(2) piecewise_C1_differentiable_on_def by auto show ?thesis using piecewise_C1_differentiable_on_components[where ?f = "(\x. (f x, g x))"] apply (simp add: real_pair_basis) using assms(2) piecewise_C1_differentiable_on_def C1_differentiable_on_subset[OF assms(1) Diff_subset, where ?B1 ="t"] t f_cont g_cont by fastforce qed lemma scale_shift_smooth: shows "(\x. a + b * x) C1_differentiable_on s" proof - show "(\x. a + b * x) C1_differentiable_on s" using C1_differentiable_on_mult C1_differentiable_on_add C1_differentiable_on_const C1_differentiable_on_ident by auto qed lemma open_diff: assumes "finite (t::'a::t1_space set)" "open (s::'a set)" shows "open (s -t)" using assms proof(induction "t") show "open s \ open (s - {})" by auto next fix x::"'a::t1_space" fix F::"'a::t1_space set" assume step: "finite F " " x \ F" "open s" then have i: "(s - insert x F) = (s - F) - {x}" by auto assume ind_hyp: " (open s \ open (s - F))" show "open (s - insert x F)" apply (simp only: i) using open_delete[of "s -F"] ind_hyp[OF step(3)] by auto qed lemma has_derivative_transform_within: assumes "0 < d" and "x \ s" and "\x'\s. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s. f x = g x" and "x \ {a..b} -s" and "(f has_derivative f') (at x within {a..b} - s)" shows "(g has_derivative f') (at x within {a..b} - s)" using has_derivative_transform_within[of "b" "x" "{a..b} - s"] assms by auto lemma has_vector_derivative_transform_within_ivl: assumes "(0::real) < b" and "\x\{a..b} -s . f x = g x" and "x \ {a..b} - s" and "(f has_vector_derivative f') (at x within {a..b} - s)" shows "(g has_vector_derivative f') (at x within {a..b} - s)" using assms has_derivative_transform_within_ivl apply (auto simp add: has_vector_derivative_def) by blast lemma has_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_derivative f') (at x)" shows "(g has_derivative f') (at x)" using has_derivative_transform_within [of d x UNIV f g f'] assms by simp lemma has_vector_derivative_transform_at: assumes "0 < d" and "\x'. dist x' x < d \ f x' = g x'" and "(f has_vector_derivative f') (at x)" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_at) lemma C1_diff_components_2: assumes "b \ Basis" assumes "f C1_differentiable_on s" shows "(\x. f x \ b) C1_differentiable_on s" proof - obtain D where D:"(\x\s. (f has_derivative (\z. z *\<^sub>R D x)) (at x))" "continuous_on s D" using assms(2) by (fastforce simp add: C1_differentiable_on_def has_vector_derivative_def) show ?thesis proof (simp add: C1_differentiable_on_def has_vector_derivative_def, intro exI conjI) show "continuous_on s (\x. D x \ b)" using D continuous_on_componentwise assms(1) by fastforce show "(\x\s. ((\x. f x \ b) has_derivative (\y. y * (\x. D x \ b) x)) (at x))" using has_derivative_inner_left D(1) by fastforce qed qed lemma eq_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "f C1_differentiable_on s" shows "g C1_differentiable_on s" proof (simp add: C1_differentiable_on_def) obtain D where D: "(\x\s. (f has_vector_derivative D x) (at x)) \ continuous_on s D" using assms by (auto simp add: C1_differentiable_on_def) then have f: "(\x\s. (g has_vector_derivative D x) (at x))" using assms(1-2) by (metis dist_commute has_vector_derivative_transform_at) have "(\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" using D f by auto then show "\D. (\x\s. (g has_vector_derivative D x) (at x)) \ continuous_on s D" by metis qed lemma eq_pw_smooth: assumes "0 < d" "\x\s. \y. dist x y < d \ f y = g y" (*This crappy condition cannot be loosened :( *) "\x\s. f x = g x" "f piecewise_C1_differentiable_on s" shows "g piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def) have g_cont: "continuous_on s g" using assms piecewise_C1_differentiable_const by (simp add: piecewise_C1_differentiable_on_def) obtain t where t: "finite t \ f C1_differentiable_on s - t" using assms by (auto simp add: piecewise_C1_differentiable_on_def) then have "g C1_differentiable_on s - t" using assms eq_smooth by blast then show "continuous_on s g \ (\t. finite t \ g C1_differentiable_on s - t)" using t g_cont by metis qed lemma scale_piecewise_C1_differentiable_on: assumes "f piecewise_C1_differentiable_on s" shows "(\x. (c::real) * (f x)) piecewise_C1_differentiable_on s" proof (simp add: piecewise_C1_differentiable_on_def, intro conjI) show "continuous_on s (\x. c * f x)" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) show "\t. finite t \ (\x. c * f x) C1_differentiable_on s - t" using assms continuous_on_mult_left by (auto simp add: piecewise_C1_differentiable_on_def) qed lemma eq_smooth_gen: assumes "f C1_differentiable_on s" "\x. f x = g x" shows "g C1_differentiable_on s" using assms unfolding C1_differentiable_on_def by (metis (no_types, lifting) has_vector_derivative_weaken UNIV_I top_greatest) lemma subpath_compose: shows "(subpath a b \) = \ o (\x. (b - a) * x + a)" by (auto simp add: subpath_def) lemma subpath_smooth: assumes "\ C1_differentiable_on {0..1}" "0 \ a" "a < b" "b \ 1" shows "(subpath a b \) C1_differentiable_on {0..1}" proof- have "\ C1_differentiable_on {a..b}" apply (rule C1_differentiable_on_subset) using assms by auto then have "\ C1_differentiable_on (\x. (b - a) * x + a) ` {0..1}" using \a < b\ closed_segment_eq_real_ivl closed_segment_real_eq by auto moreover have "finite ({0..1} \ (\x. (b - a) * x + a) -` {x})" for x proof - have "((\x. (b - a) * x + a) -` {x}) = {(x -a) /(b-a)}" using assms by (auto simp add: divide_simps) then show ?thesis by auto qed ultimately show ?thesis by (force simp add: subpath_compose intro: C1_differentiable_compose derivative_intros) qed lemma has_vector_derivative_divide[derivative_intros]: fixes a :: "'a::real_normed_field" shows "(f has_vector_derivative x) F \ ((\x. f x / a) has_vector_derivative (x / a)) F" unfolding divide_inverse by (fact has_vector_derivative_mult_left) end diff --git a/thys/Markov_Models/ex/Zeroconf_Analysis.thy b/thys/Markov_Models/ex/Zeroconf_Analysis.thy --- a/thys/Markov_Models/ex/Zeroconf_Analysis.thy +++ b/thys/Markov_Models/ex/Zeroconf_Analysis.thy @@ -1,250 +1,248 @@ (* Author: Johannes Hölzl *) section \Formalizing the IPv4-address allocation in ZeroConf\ theory Zeroconf_Analysis imports "../Discrete_Time_Markov_Chain" begin declare UNIV_bool[simp] subsection \Definition of a ZeroConf allocation run\ datatype zc_state = start | probe nat | ok | error lemma inj_probe: "inj_on probe X" by (auto simp: inj_on_def) text \Countability of @{typ zc_state} simplifies measurability of functions on @{typ zc_state}.\ instance zc_state :: countable proof have "countable ({start, ok, error} \ probe`UNIV)" by auto also have "{start, ok, error} \ probe`UNIV = UNIV" using zc_state.nchotomy by auto finally show "\f::zc_state \ nat. inj f" using inj_on_to_nat_on[of "UNIV :: zc_state set"] by auto qed locale Zeroconf_Analysis = fixes N :: nat and p q r e :: real assumes p: "0 < p" "p < 1" and q: "0 < q" "q < 1" assumes r[simp]: "0 \ r" and e[simp]: "0 \ e" begin lemma p_bounds[simp]: "0 \ p" "p \ 1" using p by auto lemma q_bounds[simp]: "0 \ q" "q \ 1" using q by auto abbreviation states where "states \ probe ` {.. N} \ {start, ok, error}" primrec \ :: "zc_state \ zc_state pmf" where "\ start = map_pmf (\True \ probe 0 | False \ ok) (bernoulli_pmf q)" | "\ (probe n) = map_pmf (\True \ (if n < N then probe (Suc n) else error) | False \ start) (bernoulli_pmf p)" | "\ ok = return_pmf ok" | "\ error = return_pmf error" primrec \ :: "zc_state \ zc_state \ real" where "\ start = (\_. 0) (probe 0 := r, ok := r * (N + 1))" | "\ (probe n) = (if n < N then (\_. 0) (probe (Suc n) := r) else (\_. 0) (error := e))" | "\ ok = (\_. 0) (ok := 0)" | "\ error = (\_. 0) (error := 0)" lemma \_nonneg'[simp]: "0 \ \ s t" using r e by (cases s) auto sublocale MC_with_rewards \ \ "\s. 0" proof qed (simp_all add: pair_measure_countable) subsection \The allocation run is a rewarded DTMC\ abbreviation "E s \ set_pmf (\ s)" lemma enabled_ok: "enabled ok \ \ \ = sconst ok" by (simp add: enabled_iff_sconst) lemma finite_E[intro, simp]: "finite (E s)" by (cases s) auto lemma E_closed: "s \ states \ E s \ states" using p q by (cases s) (auto split: bool.splits) lemma enabled_error: "enabled error \ \ \ = sconst error" by (simp add: enabled_iff_sconst) lemma pos_neg_q_pn: "0 < 1 - q * (1 - p^Suc N)" proof - have "p ^ Suc N \ 1 ^ Suc N" using p by (intro power_mono) auto with p q have "q * (1 - p^Suc N) < 1 * 1" by (intro mult_strict_mono) (auto simp: field_simps simp del: power_Suc) then show ?thesis by simp qed lemma to_error: assumes "n \ N" shows "(probe n, error) \ acc" using \n \ N\ proof (induction rule: inc_induct) case (step n') with p show ?case by (intro rtrancl_trans[OF r_into_rtrancl step.IH]) auto qed (insert p, auto) subsection \Probability of a erroneous allocation\ definition "P_err s = \

(\ in T s. ev (HLD {error}) (s ## \))" lemma P_err: defines "p_start == (q * p ^ Suc N) / (1 - q * (1 - p ^ Suc N))" defines "p_probe == (\n. p ^ Suc (N - n) + (1 - p^Suc (N - n)) * p_start)" assumes s: "s \ states - {ok, error}" shows "P_err s = (case s of ok \ 0 | error \ 1 | probe n \ p_probe n | start \ p_start)" (is "\ = ?E s") using s proof (rule unique_les) have [arith]: "0 \ p * (q * p ^ N)" using p q by simp have p_eq: "p_start = p_probe 0 * q" "\n. n < N \ p_probe n = p_probe (Suc n) * p + p_start * (1 - p)" "p_probe N = p + p_start * (1 - p)" using p q by (auto simp: p_probe_def p_start_def power_Suc[symmetric] Suc_diff_Suc divide_simps simp del: power_Suc) (auto simp: field_simps) fix s assume s: "s \ states - {ok, error}" then show "?E s = (\t. ?E t \\ s) + 0" using p q by (auto intro: p_eq) show "\t\{ok, error}. (s, t) \ acc" using s q to_error by auto from s show "P_err s = integral\<^sup>L (measure_pmf (\ s)) P_err + 0" unfolding P_err_def[abs_def] by (subst prob_T) (auto simp: ev_Stream simp del: UNIV_bool) next fix s assume "s \ {ok, error}" then show "P_err s = ?E s" by (auto intro!: T.prob_eq_0_AE T.prob_Collect_eq_1[THEN iffD2] simp: P_err_def AE_sconst ev_sconst HLD_iff ev_Stream T.prob_space simp del: space_T sets_T ) qed (insert p q, auto intro!: integrable_measure_pmf_finite split: if_split_asm) lemma P_err_start: "P_err start = (q * p ^ Suc N) / (1 - q * (1 - p ^ Suc N))" by (simp add: P_err) subsection \An allocation run terminates almost surely\ lemma states_closed: assumes "s \ states" assumes "(s, t) \ acc_on (- {error, ok})" shows "t \ states" using assms(2,1) p q by induction (auto split: if_split_asm) lemma finite_reached: assumes s: "s \ states" shows "finite (acc_on (- {error, ok}) `` {s})" using states_closed[OF s] by (rule_tac finite_subset[of _ states]) auto lemma AE_reaches_error_or_ok: assumes s: "s \ states" shows "AE \ in T s. ev (HLD {error, ok}) \" proof (rule AE_T_ev_HLD) { fix t assume t: "(s, t) \ acc_on (- {error, ok})" with states_closed[OF s t] to_error p q show "\t'\{error, ok}. (t, t') \ acc" by auto } qed (rule finite_reached[OF s]) subsection \Expected runtime of an allocation run\ definition "R s = (\\<^sup>+ \. reward_until {error, ok} s \ \T s)" definition "R' s = enn2real (R s)" lemma R_iter: "s \ error \ s \ ok \ R s = (\\<^sup>+t. ennreal (\ s t) + R t \\ s)" unfolding R_def using T.emeasure_space_1 by (subst nn_integral_T) (auto simp del: \.simps \.simps simp add: AE_measure_pmf_iff nn_integral_add intro!: nn_integral_cong_AE) lemma R_finite: assumes s: "s \ states" shows "R s \ \" unfolding R_def proof (rule nn_integral_reward_until_finite) { fix t assume "(s, t) \ acc" from this s p q have "t \ states" by induction (auto split: if_split_asm) } then have "acc `` {s} \ states" by auto then show "finite (acc `` {s})" by (auto dest: finite_subset) qed (auto simp: AE_reaches_error_or_ok[OF s]) lemma R_less_top: "s \ states \ R s < top" using R_finite[of s] by (subst less_top[symmetric]) simp lemma R'_iter: assumes s: "s \ states" "s \ error" "s \ ok" shows "R' s = (\t. \ s t + R' t \\ s)" unfolding R'_def R_iter[OF s(2,3)] proof (rule enn2real_nn_integral_eq_integral) have "t \ E s \ R t < top" for t using \s\states\ E_closed[of s] by (intro R_less_top) auto then show "AE t in \ s. ennreal (\ s t) + R t = ennreal (\ s t + enn2real (R t))" by (auto simp: AE_measure_pmf_iff intro!: ennreal_enn2real[symmetric]) - show "(\\<^sup>+ t. ennreal (\ s t) + R t \\ s) < \" - unfolding R_iter[symmetric, OF s(2,3)] by (rule R_less_top) fact qed auto lemma cost_from_start: "R' start = (q * (r + p^Suc N * e + r * p * (1 - p^N) / (1 - p)) + (1 - q) * (r * Suc N)) / (1 - q + q * p^Suc N)" proof - have ok_error: "R' ok = 0 \ R' error = 0" unfolding R'_def R_def by (subst (1 2) reward_until_unfold[abs_def]) simp then have R_start: "R' start = q * (r + R' (probe 0)) + (1 - q) * (r * (N + 1))" using q r by (subst R'_iter) (simp_all add: field_simps) have R_probe: "\n. n < N \ R' (probe n) = p * R' (probe (Suc n)) + p * r + (1 - p) * R' start" using p r by (subst R'_iter) (simp_all add: field_simps distrib_right) have R_N: "R' (probe N) = p * e + (1 - p) * R' start" using p e ok_error by (subst R'_iter) (auto simp: mult.commute ) { fix n assume "n \ N" then have "R' (probe (N - n)) = p ^ Suc n * e + (1 - p^n) * r * p / (1 - p) + (1 - p^Suc n) * R' start" proof (induct n) case 0 with R_N show ?case by simp next case (Suc n) moreover then have "Suc (N - Suc n) = N - n" by simp ultimately show ?case using R_probe[of "N - Suc n"] p by (simp_all add: field_simps Suc) qed } from this[of N] have [simp]: "R' (probe 0) = p ^ Suc N * e + (1 - p^N) * r * p / (1 - p) + (1 - p^Suc N) * R' start" by simp have "R' start - q * (1 - p^Suc N) * R' start = q * (r + p^Suc N * e + (1 - p^N) * r * p / (1 - p)) + (1 - q) * (r * (N + 1))" by (subst R_start) (simp_all add: field_simps) then have "R' start = (q * (r + p^Suc N * e + (1 - p^N) * r * p / (1 - p)) + (1 - q) * (r * Suc N)) / (1 - q * (1 - p^Suc N))" using pos_neg_q_pn by (simp_all add: field_simps) then show ?thesis by (simp add: field_simps) qed end interpretation ZC: Zeroconf_Analysis 2 "16 / 65024 :: real" "0.01" "0.002" "3600" by standard auto lemma "ZC.P_err start \ 1 / 10^12" unfolding ZC.P_err_start by (simp add: power_divide power_one_over[symmetric]) lemma "ZC.R' start \ 0.007" unfolding ZC.cost_from_start by (simp add: power_divide power_one_over[symmetric]) end diff --git a/thys/Ordinary_Differential_Equations/IVP/Flow.thy b/thys/Ordinary_Differential_Equations/IVP/Flow.thy --- a/thys/Ordinary_Differential_Equations/IVP/Flow.thy +++ b/thys/Ordinary_Differential_Equations/IVP/Flow.thy @@ -1,3194 +1,3193 @@ section \Flow\ theory Flow imports Picard_Lindeloef_Qualitative "HOL-Library.Diagonal_Subsequence" "../Library/Bounded_Linear_Operator" "../Library/Multivariate_Taylor" "../Library/Interval_Integral_HK" begin text \TODO: extend theorems for dependence on initial time\ subsection \simp rules for integrability (TODO: move)\ lemma blinfun_ext: "x = y \ (\i. blinfun_apply x i = blinfun_apply y i)" by transfer auto notation id_blinfun ("1\<^sub>L") lemma blinfun_inverse_left: fixes f::"'a::euclidean_space \\<^sub>L 'a" and f' shows "f o\<^sub>L f' = 1\<^sub>L \ f' o\<^sub>L f = 1\<^sub>L" by transfer (auto dest!: bounded_linear.linear simp: id_def[symmetric] linear_inverse_left) lemma onorm_zero_blinfun[simp]: "onorm (blinfun_apply 0) = 0" by transfer (simp add: onorm_zero) lemma blinfun_compose_1_left[simp]: "x o\<^sub>L 1\<^sub>L = x" and blinfun_compose_1_right[simp]: "1\<^sub>L o\<^sub>L y = y" by (auto intro!: blinfun_eqI) named_theorems integrable_on_simps lemma integrable_on_refl_ivl[intro, simp]: "g integrable_on {b .. (b::'b::ordered_euclidean_space)}" and integrable_on_refl_closed_segment[intro, simp]: "h integrable_on closed_segment a a" using integrable_on_refl by auto lemma integrable_const_ivl_closed_segment[intro, simp]: "(\x. c) integrable_on closed_segment a (b::real)" by (auto simp: closed_segment_eq_real_ivl) lemma integrable_ident_ivl[intro, simp]: "(\x. x) integrable_on closed_segment a (b::real)" and integrable_ident_cbox[intro, simp]: "(\x. x) integrable_on cbox a (b::real)" by (auto simp: closed_segment_eq_real_ivl ident_integrable_on) lemma content_closed_segment_real: fixes a b::real shows "content (closed_segment a b) = abs (b - a)" by (auto simp: closed_segment_eq_real_ivl) lemma integral_const_closed_segment: fixes a b::real shows "integral (closed_segment a b) (\x. c) = abs (b - a) *\<^sub>R c" by (auto simp: closed_segment_eq_real_ivl content_closed_segment_real) lemmas [integrable_on_simps] = integrable_on_empty \ \empty\ integrable_on_refl integrable_on_refl_ivl integrable_on_refl_closed_segment \ \singleton\ integrable_const integrable_const_ivl integrable_const_ivl_closed_segment \ \constant\ ident_integrable_on integrable_ident_ivl integrable_ident_cbox \ \identity\ lemma integrable_cmul_real: fixes K::real shows "f integrable_on X \ (\x. K * f x) integrable_on X " unfolding real_scaleR_def[symmetric] by (rule integrable_cmul) lemmas [integrable_on_simps] = integrable_0 integrable_neg integrable_cmul integrable_cmul_real integrable_on_cmult_iff integrable_on_cmult_left integrable_on_cmult_right - integrable_on_cdivide integrable_on_cmult_iff integrable_on_cmult_left_iff integrable_on_cmult_right_iff integrable_on_cdivide_iff integrable_diff integrable_add integrable_sum lemma dist_cancel_add1: "dist (t0 + et) t0 = norm et" by (simp add: dist_norm) lemma double_nonneg_le: fixes a::real shows "a * 2 \ b \ a \ 0 \ a \ b" by arith subsection \Nonautonomous IVP on maximal existence interval\ context ll_on_open_it begin context fixes x0 assumes iv_defined: "t0 \ T" "x0 \ X" begin lemmas closed_segment_iv_subset_domain = closed_segment_subset_domainI[OF iv_defined(1)] lemma local_unique_solutions: obtains t u L where "0 < t" "0 < u" "cball t0 t \ existence_ivl t0 x0" "cball x0 (2 * u) \ X" "\t'. t' \ cball t0 t \ L-lipschitz_on (cball x0 (2 * u)) (f t')" "\x. x \ cball x0 u \ (flow t0 x usolves_ode f from t0) (cball t0 t) (cball x u)" "\x. x \ cball x0 u \ cball x u \ X" proof - from local_unique_solution[OF iv_defined] obtain et ex B L where "0 < et" "0 < ex" "cball t0 et \ T" "cball x0 ex \ X" "unique_on_cylinder t0 (cball t0 et) x0 ex f B L" by metis then interpret cyl: unique_on_cylinder t0 "cball t0 et" x0 ex "cball x0 ex" f B L by auto from cyl.solution_solves_ode order_refl \cball x0 ex \ X\ have "(cyl.solution solves_ode f) (cball t0 et) X" by (rule solves_ode_on_subset) then have "cball t0 et \ existence_ivl t0 x0" by (rule existence_ivl_maximal_interval) (insert \cball t0 et \ T\ \0 < et\, auto) have "cball t0 et = {t0 - et .. t0 + et}" using \et > 0\ by (auto simp: dist_real_def) then have cylbounds[simp]: "cyl.tmin = t0 - et" "cyl.tmax = t0 + et" unfolding cyl.tmin_def cyl.tmax_def using \0 < et\ by auto define et' where "et' \ et / 2" define ex' where "ex' \ ex / 2" have "et' > 0" "ex' > 0" using \0 < et\ \0 < ex\ by (auto simp: et'_def ex'_def) moreover from \cball t0 et \ existence_ivl t0 x0\ have "cball t0 et' \ existence_ivl t0 x0" by (force simp: et'_def dest!: double_nonneg_le) moreover from this have "cball t0 et' \ T" using existence_ivl_subset[of x0] by simp have "cball x0 (2 * ex') \ X" "\t'. t' \ cball t0 et' \ L-lipschitz_on (cball x0 (2 * ex')) (f t')" using cyl.lipschitz \0 < et\ \cball x0 ex \ X\ by (auto simp: ex'_def et'_def intro!:) moreover { fix x0'::'a assume x0': "x0' \ cball x0 ex'" { fix b assume d: "dist x0' b \ ex'" have "dist x0 b \ dist x0 x0' + dist x0' b" by (rule dist_triangle) also have "\ \ ex' + ex'" using x0' d by simp also have "\ \ ex" by (simp add: ex'_def) finally have "dist x0 b \ ex" . } note triangle = this have subs1: "cball t0 et' \ cball t0 et" and subs2: "cball x0' ex' \ cball x0 ex" and subs: "cball t0 et' \ cball x0' ex' \ cball t0 et \ cball x0 ex" using \0 < ex\ \0 < et\ x0' by (auto simp: ex'_def et'_def triangle dest!: double_nonneg_le) have subset_X: "cball x0' ex' \ X" using \cball x0 ex \ X\ subs2 \0 < ex'\ by force then have "x0' \ X" using \0 < ex'\ by force have x0': "t0 \ T" "x0' \ X" by fact+ have half_intros: "a \ ex' \ a \ ex" "a \ et' \ a \ et" and halfdiv_intro: "a * 2 \ ex / B \ a \ ex' / B" for a using \0 < ex\ \0 < et\ by (auto simp: ex'_def et'_def) interpret cyl': solution_in_cylinder t0 "cball t0 et'" x0' ex' f "cball x0' ex'" B using \0 < et'\ \0 < ex'\ \0 < et\ cyl.norm_f cyl.continuous subs1 \cball t0 et \ T\ apply unfold_locales apply (auto simp: split_beta' dist_cancel_add1 intro!: triangle continuous_intros cyl.norm_f order_trans[OF _ cyl.e_bounded] halfdiv_intro) by (simp add: ex'_def et'_def dist_commute) interpret cyl': unique_on_cylinder t0 "cball t0 et'" x0' ex' "cball x0' ex'" f B L using cyl.lipschitz[simplified] subs subs1 by (unfold_locales) (auto simp: triangle intro!: half_intros lipschitz_on_subset[OF _ subs2]) from cyl'.solution_usolves_ode have "(flow t0 x0' usolves_ode f from t0) (cball t0 et') (cball x0' ex')" apply (rule usolves_ode_solves_odeI) subgoal apply (rule cyl'.solves_ode_on_subset_domain[where Y=X]) subgoal apply (rule solves_ode_on_subset[where S="existence_ivl t0 x0'" and Y=X]) subgoal by (rule flow_solves_ode[OF x0']) subgoal using subs2 \cball x0 ex \ X\ \0 < et'\ \cball t0 et' \ T\ by (intro existence_ivl_maximal_interval[OF solves_ode_on_subset[OF cyl'.solution_solves_ode]]) auto subgoal by force done subgoal by (force simp: \x0' \ X\ iv_defined) subgoal using \0 < et'\ by force subgoal by force subgoal by force done subgoal by (force simp: \x0' \ X\ iv_defined cyl'.solution_iv) done note this subset_X } ultimately show thesis .. qed lemma Picard_iterate_mem_existence_ivlI: assumes "t \ T" assumes "compact C" "x0 \ C" "C \ X" assumes "\y s. s \ {t0 -- t} \ y t0 = x0 \ y \ {t0--s} \ C \ continuous_on {t0--s} y \ x0 + ivl_integral t0 s (\t. f t (y t)) \ C" shows "t \ existence_ivl t0 x0" "\s. s \ {t0 -- t} \ flow t0 x0 s \ C" proof - have "{t0 -- t} \ T" by (intro closed_segment_subset_domain iv_defined assms) from lipschitz_on_compact[OF compact_segment \{t0 -- t} \ T\ \compact C\ \C \ X\] obtain L where L: "\s. s \ {t0 -- t} \ L-lipschitz_on C (f s)" by metis interpret uc: unique_on_closed t0 "{t0 -- t}" x0 f C L using assms closed_segment_iv_subset_domain by unfold_locales (auto intro!: L compact_imp_closed \compact C\ continuous_on_f continuous_intros simp: split_beta) have "{t0 -- t} \ existence_ivl t0 x0" using assms closed_segment_iv_subset_domain by (intro maximal_existence_flow[OF solves_ode_on_subset[OF uc.solution_solves_ode]]) auto thus "t \ existence_ivl t0 x0" using assms by auto show "flow t0 x0 s \ C" if "s \ {t0 -- t}" for s proof - have "flow t0 x0 s = uc.solution s" "uc.solution s \ C" using solves_odeD[OF uc.solution_solves_ode] that assms by (auto simp: closed_segment_iv_subset_domain intro!: maximal_existence_flowI(2)[where K="{t0 -- t}"]) thus ?thesis by simp qed qed lemma flow_has_vderiv_on: "(flow t0 x0 has_vderiv_on (\t. f t (flow t0 x0 t))) (existence_ivl t0 x0)" by (rule solves_ode_vderivD[OF flow_solves_ode[OF iv_defined]]) lemmas flow_has_vderiv_on_compose[derivative_intros] = has_vderiv_on_compose2[OF flow_has_vderiv_on, THEN has_vderiv_on_eq_rhs] end lemma unique_on_intersection: assumes sols: "(x solves_ode f) U X" "(y solves_ode f) V X" assumes iv_mem: "t0 \ U" "t0 \ V" and subs: "U \ T" "V \ T" assumes ivls: "is_interval U" "is_interval V" assumes iv: "x t0 = y t0" assumes mem: "t \ U" "t \ V" shows "x t = y t" proof - from maximal_existence_flow(2)[OF sols(1) refl ivls(1) iv_mem(1) subs(1) mem(1)] maximal_existence_flow(2)[OF sols(2) iv[symmetric] ivls(2) iv_mem(2) subs(2) mem(2)] show ?thesis by simp qed lemma unique_solution: assumes sols: "(x solves_ode f) U X" "(y solves_ode f) U X" assumes iv_mem: "t0 \ U" and subs: "U \ T" assumes ivls: "is_interval U" assumes iv: "x t0 = y t0" assumes mem: "t \ U" shows "x t = y t" by (metis unique_on_intersection assms) lemma assumes s: "s \ existence_ivl t0 x0" assumes t: "t + s \ existence_ivl s (flow t0 x0 s)" shows flow_trans: "flow t0 x0 (s + t) = flow s (flow t0 x0 s) (s + t)" and existence_ivl_trans: "s + t \ existence_ivl t0 x0" proof - note ll_on_open_it_axioms moreover from ll_on_open_it_axioms have iv_defined: "t0 \ T" "x0 \ X" and iv_defined': "s \ T" "flow t0 x0 s \ X" using ll_on_open_it.mem_existence_ivl_iv_defined s t by blast+ have "{t0--s} \ existence_ivl t0 x0" by (simp add: s segment_subset_existence_ivl iv_defined) have "s \ existence_ivl s (flow t0 x0 s)" by (rule ll_on_open_it.existence_ivl_initial_time; fact) have "{s--t + s} \ existence_ivl s (flow t0 x0 s)" by (rule ll_on_open_it.segment_subset_existence_ivl; fact) have unique: "flow t0 x0 u = flow s (flow t0 x0 s) u" if "u \ {s--t + s}" "u \ {t0--s}" for u using ll_on_open_it_axioms ll_on_open_it.flow_solves_ode[OF ll_on_open_it_axioms iv_defined] ll_on_open_it.flow_solves_ode[OF ll_on_open_it_axioms iv_defined'] s apply (rule ll_on_open_it.unique_on_intersection) using \s \ existence_ivl s (flow t0 x0 s)\ existence_ivl_subset \flow t0 x0 s \ X\ \s \ T\ iv_defined s t ll_on_open_it.in_existence_between_zeroI that ll_on_open_it_axioms ll_on_open_it.mem_existence_ivl_subset by (auto simp: is_interval_existence_ivl) let ?un = "{t0 -- s} \ {s -- t + s}" let ?if = "\t. if t \ {t0 -- s} then flow t0 x0 t else flow s (flow t0 x0 s) t" have "(?if solves_ode (\t. if t \ {t0 -- s} then f t else f t)) ?un (X \ X)" apply (rule connection_solves_ode) subgoal by (rule solves_ode_on_subset[OF flow_solves_ode[OF iv_defined] \{t0--s} \ _\ order_refl]) subgoal by (rule solves_ode_on_subset[OF ll_on_open_it.flow_solves_ode[OF ll_on_open_it_axioms iv_defined'] \{s--t + s} \ _\ order_refl]) subgoal by simp subgoal by simp subgoal by (rule unique) auto subgoal by simp done then have ifsol: "(?if solves_ode f) ?un X" by simp moreover have "?un \ existence_ivl t0 x0" using existence_ivl_subset[of x0] ll_on_open_it.existence_ivl_subset[OF ll_on_open_it_axioms, of s "flow t0 x0 s"] \{t0 -- s} \ _\ \{s--t + s} \ _\ by (intro existence_ivl_maximal_interval[OF ifsol]) (auto intro!: is_real_interval_union) then show "s + t \ existence_ivl t0 x0" by (auto simp: ac_simps) have "(flow t0 x0 solves_ode f) ?un X" using \{t0--s} \ _\ \{s -- t + s} \ _\ by (intro solves_ode_on_subset[OF flow_solves_ode \?un \ _\ order_refl] iv_defined) moreover have "s \ ?un" by simp ultimately have "?if (s + t) = flow t0 x0 (s + t)" apply (rule ll_on_open_it.unique_solution) using existence_ivl_subset[of x0] ll_on_open_it.existence_ivl_subset[OF ll_on_open_it_axioms, of s "flow t0 x0 s"] \{t0 -- s} \ _\ \{s--t + s} \ _\ by (auto intro!: is_real_interval_union simp: ac_simps) with unique[of "s + t"] show "flow t0 x0 (s + t) = flow s (flow t0 x0 s) (s + t)" by (auto split: if_splits simp: ac_simps) qed lemma assumes t: "t \ existence_ivl t0 x0" shows flows_reverse: "flow t (flow t0 x0 t) t0 = x0" and existence_ivl_reverse: "t0 \ existence_ivl t (flow t0 x0 t)" proof - have iv_defined: "t0 \ T" "x0 \ X" using mem_existence_ivl_iv_defined t by blast+ show "t0 \ existence_ivl t (flow t0 x0 t)" using assms by (metis (no_types, opaque_lifting) closed_segment_commute closed_segment_subset_interval ends_in_segment(2) general.csol(2-4) general.existence_ivl_maximal_segment general.is_interval_existence_ivl is_interval_closed_segment_1 iv_defined ll_on_open_it.equals_flowI local.existence_ivl_initial_time local.flow_initial_time local.ll_on_open_it_axioms) then have "flow t (flow t0 x0 t) (t + (t0 - t)) = flow t0 x0 (t + (t0 - t))" by (intro flow_trans[symmetric]) (auto simp: t iv_defined) then show "flow t (flow t0 x0 t) t0 = x0" by (simp add: iv_defined) qed lemma flow_has_derivative: assumes "t \ existence_ivl t0 x0" shows "(flow t0 x0 has_derivative (\i. i *\<^sub>R f t (flow t0 x0 t))) (at t)" proof - have "(flow t0 x0 has_derivative (\i. i *\<^sub>R f t (flow t0 x0 t))) (at t within existence_ivl t0 x0)" using flow_has_vderiv_on by (auto simp: has_vderiv_on_def has_vector_derivative_def assms mem_existence_ivl_iv_defined[OF assms]) then show ?thesis by (simp add: at_within_open[OF assms open_existence_ivl]) qed lemma flow_has_vector_derivative: assumes "t \ existence_ivl t0 x0" shows "(flow t0 x0 has_vector_derivative f t (flow t0 x0 t)) (at t)" using flow_has_derivative[OF assms] by (simp add: has_vector_derivative_def) lemma flow_has_vector_derivative_at_0: assumes"t \ existence_ivl t0 x0" shows "((\h. flow t0 x0 (t + h)) has_vector_derivative f t (flow t0 x0 t)) (at 0)" proof - from flow_has_vector_derivative[OF assms] have "((+) t has_vector_derivative 1) (at 0)" "(flow t0 x0 has_vector_derivative f t (flow t0 x0 t)) (at (t + 0))" by (auto intro!: derivative_eq_intros) from vector_diff_chain_at[OF this] show ?thesis by (simp add: o_def) qed lemma assumes "t \ existence_ivl t0 x0" shows closed_segment_subset_existence_ivl: "closed_segment t0 t \ existence_ivl t0 x0" and ivl_subset_existence_ivl: "{t0 .. t} \ existence_ivl t0 x0" and ivl_subset_existence_ivl': "{t .. t0} \ existence_ivl t0 x0" using assms in_existence_between_zeroI by (auto simp: closed_segment_eq_real_ivl) lemma flow_fixed_point: assumes t: "t \ existence_ivl t0 x0" shows "flow t0 x0 t = x0 + ivl_integral t0 t (\t. f t (flow t0 x0 t))" proof - have "(flow t0 x0 has_vderiv_on (\s. f s (flow t0 x0 s))) {t0 -- t}" using closed_segment_subset_existence_ivl[OF t] by (auto intro!: has_vector_derivative_at_within flow_has_vector_derivative simp: has_vderiv_on_def) from fundamental_theorem_of_calculus_ivl_integral[OF this] have "((\t. f t (flow t0 x0 t)) has_ivl_integral flow t0 x0 t - x0) t0 t" by (simp add: mem_existence_ivl_iv_defined[OF assms]) from this[THEN ivl_integral_unique] show ?thesis by simp qed lemma flow_continuous: "t \ existence_ivl t0 x0 \ continuous (at t) (flow t0 x0)" by (metis has_derivative_continuous flow_has_derivative) lemma flow_tendsto: "t \ existence_ivl t0 x0 \ (ts \ t) F \ ((\s. flow t0 x0 (ts s)) \ flow t0 x0 t) F" by (rule isCont_tendsto_compose[OF flow_continuous]) lemma flow_continuous_on: "continuous_on (existence_ivl t0 x0) (flow t0 x0)" by (auto intro!: flow_continuous continuous_at_imp_continuous_on) lemma flow_continuous_on_intro: "continuous_on s g \ (\xa. xa \ s \ g xa \ existence_ivl t0 x0) \ continuous_on s (\xa. flow t0 x0 (g xa))" by (auto intro!: continuous_on_compose2[OF flow_continuous_on]) lemma f_flow_continuous: assumes "t \ existence_ivl t0 x0" shows "isCont (\t. f t (flow t0 x0 t)) t" by (rule continuous_on_interior) (insert existence_ivl_subset assms, auto intro!: flow_in_domain flow_continuous_on continuous_intros simp: interior_open open_existence_ivl) lemma exponential_initial_condition: assumes y0: "t \ existence_ivl t0 y0" assumes z0: "t \ existence_ivl t0 z0" assumes "Y \ X" assumes remain: "\s. s \ closed_segment t0 t \ flow t0 y0 s \ Y" "\s. s \ closed_segment t0 t \ flow t0 z0 s \ Y" assumes lipschitz: "\s. s \ closed_segment t0 t \ K-lipschitz_on Y (f s)" shows "norm (flow t0 y0 t - flow t0 z0 t) \ norm (y0 - z0) * exp ((K + 1) * abs (t - t0))" proof cases assume "y0 = z0" thus ?thesis by simp next assume ne: "y0 \ z0" define K' where "K' \ K + 1" from lipschitz have "K'-lipschitz_on Y (f s)" if "s \ {t0 -- t}" for s using that by (auto simp: lipschitz_on_def K'_def intro!: order_trans[OF _ mult_right_mono[of K "K + 1"]]) from mem_existence_ivl_iv_defined[OF y0] mem_existence_ivl_iv_defined[OF z0] have "t0 \ T" and inX: "y0 \ X" "z0 \ X" by auto from remain[of t0] inX \t0 \ T \ have "y0 \ Y" "z0 \ Y" by auto define v where "v \ \t. norm (flow t0 y0 t - flow t0 z0 t)" { fix s assume s: "s \ {t0 -- t}" with s closed_segment_subset_existence_ivl[OF y0] closed_segment_subset_existence_ivl[OF z0] have y0': "s \ existence_ivl t0 y0" and z0': "s \ existence_ivl t0 z0" by (auto simp: closed_segment_eq_real_ivl) have integrable: "(\t. f t (flow t0 y0 t)) integrable_on {t0--s}" "(\t. f t (flow t0 z0 t)) integrable_on {t0--s}" using closed_segment_subset_existence_ivl[OF y0'] closed_segment_subset_existence_ivl[OF z0'] \y0 \ X\ \z0 \ X\ \t0 \ T\ by (auto intro!: continuous_at_imp_continuous_on f_flow_continuous integrable_continuous_closed_segment) hence int: "flow t0 y0 s - flow t0 z0 s = y0 - z0 + ivl_integral t0 s (\t. f t (flow t0 y0 t) - f t (flow t0 z0 t))" unfolding v_def using flow_fixed_point[OF y0'] flow_fixed_point[OF z0'] s by (auto simp: algebra_simps ivl_integral_diff) have "v s \ v t0 + K' * integral {t0 -- s} (\t. v t)" using closed_segment_subset_existence_ivl[OF y0'] closed_segment_subset_existence_ivl[OF z0'] s using closed_segment_closed_segment_subset[OF _ _ s, of _ t0, simplified] by (subst integral_mult) (auto simp: integral_mult v_def int inX \t0 \ T\ simp del: Henstock_Kurzweil_Integration.integral_mult_right intro!: norm_triangle_le ivl_integral_norm_bound_integral integrable_continuous_closed_segment continuous_intros continuous_at_imp_continuous_on flow_continuous f_flow_continuous lipschitz_on_normD[OF \_ \ K'-lipschitz_on _ _\] remain) } note le = this have cont: "continuous_on {t0 -- t} v" using closed_segment_subset_existence_ivl[OF y0] closed_segment_subset_existence_ivl[OF z0] inX by (auto simp: v_def \t0 \ T\ intro!: continuous_at_imp_continuous_on continuous_intros flow_continuous) have nonneg: "\t. v t \ 0" by (auto simp: v_def) from ne have pos: "v t0 > 0" by (auto simp: v_def \t0 \ T\ inX) have lippos: "K' > 0" proof - have "0 \ dist (f t0 y0) (f t0 z0)" by simp also from lipschitz_onD[OF lipschitz \y0 \ Y\ \z0 \ Y\, of t0]ne have "\ \ K * dist y0 z0" by simp finally have "0 \ K" by (metis dist_le_zero_iff ne zero_le_mult_iff) thus ?thesis by (simp add: K'_def) qed from le cont nonneg pos \0 < K'\ have "v t \ v t0 * exp (K' * abs (t - t0))" by (rule gronwall_general_segment) simp_all thus ?thesis by (simp add: v_def K'_def \t0 \ T\ inX) qed lemma existence_ivl_cballs: assumes iv_defined: "t0 \ T" "x0 \ X" obtains t u L where "\y. y \ cball x0 u \ cball t0 t \ existence_ivl t0 y" "\s y. y \ cball x0 u \ s \ cball t0 t \ flow t0 y s \ cball y u" "L-lipschitz_on (cball t0 t\cball x0 u) (\(t, x). flow t0 x t)" "\y. y \ cball x0 u \ cball y u \ X" "0 < t" "0 < u" proof - note iv_defined from local_unique_solutions[OF this] obtain t u L where tu: "0 < t" "0 < u" and subsT: "cball t0 t \ existence_ivl t0 x0" and subs': "cball x0 (2 * u) \ X" and lipschitz: "\s. s \ cball t0 t \ L-lipschitz_on (cball x0 (2*u)) (f s)" and usol: "\y. y \ cball x0 u \ (flow t0 y usolves_ode f from t0) (cball t0 t) (cball y u)" and subs: "\y. y \ cball x0 u \ cball y u \ X" by metis { fix y assume y: "y \ cball x0 u" from subs[OF y] \0 < u\ have "y \ X" by auto note iv' = \t0 \ T\ \y \ X\ from usol[OF y, THEN usolves_odeD(1)] have sol1: "(flow t0 y solves_ode f) (cball t0 t) (cball y u)" . from sol1 order_refl subs[OF y] have sol: "(flow t0 y solves_ode f) (cball t0 t) X" by (rule solves_ode_on_subset) note * = maximal_existence_flow[OF sol flow_initial_time is_interval_cball_1 _ order_trans[OF subsT existence_ivl_subset], unfolded centre_in_cball, OF iv' less_imp_le[OF \0 < t\]] have eivl: "cball t0 t \ existence_ivl t0 y" by (rule *) have "flow t0 y s \ cball y u" if "s \ cball t0 t" for s by (rule solves_odeD(2)[OF sol1 that]) note eivl this } note * = this note * moreover have cont_on_f_flow: "\x1 S. S \ cball t0 t \ x1 \ cball x0 u \ continuous_on S (\t. f t (flow t0 x1 t))" using subs[of x0] \u > 0\ *(1) iv_defined by (auto intro!: continuous_at_imp_continuous_on f_flow_continuous) have "bounded ((\(t, x). f t x) ` (cball t0 t \ cball x0 (2 * u)))" using subs' subsT existence_ivl_subset[of x0] by (auto intro!: compact_imp_bounded compact_continuous_image compact_Times continuous_intros simp: split_beta') then obtain B where B: "\s y. s \ cball t0 t \ y \ cball x0 (2 * u) \ norm (f s y) \ B" "B > 0" by (auto simp: bounded_pos cball_def) have flow_in_cball: "flow t0 x1 s \ cball x0 (2 * u)" if s: "s \ cball t0 t" and x1: "x1 \ cball x0 u" for s::real and x1 proof - from *(2)[OF x1 s] have "flow t0 x1 s \ cball x1 u" . also have "\ \ cball x0 (2 * u)" using x1 by (auto intro!: dist_triangle_le[OF add_mono, of _ x1 u _ u, simplified] simp: dist_commute) finally show ?thesis . qed have "(B + exp ((L + 1) * \t\))-lipschitz_on (cball t0 t\cball x0 u) (\(t, x). flow t0 x t)" proof (rule lipschitz_onI, safe) fix t1 t2 :: real and x1 x2 assume t1: "t1 \ cball t0 t" and t2: "t2 \ cball t0 t" and x1: "x1 \ cball x0 u" and x2: "x2 \ cball x0 u" have t1_ex: "t1 \ existence_ivl t0 x1" and t2_ex: "t2 \ existence_ivl t0 x1" "t2 \ existence_ivl t0 x2" and "x1 \ cball x0 (2*u)" "x2 \ cball x0 (2*u)" using *(1)[OF x1] *(1)[OF x2] t1 t2 x1 x2 tu by auto have "dist (flow t0 x1 t1) (flow t0 x2 t2) \ dist (flow t0 x1 t1) (flow t0 x1 t2) + dist (flow t0 x1 t2) (flow t0 x2 t2)" by (rule dist_triangle) also have "dist (flow t0 x1 t2) (flow t0 x2 t2) \ dist x1 x2 * exp ((L + 1) * \t2 - t0\)" unfolding dist_norm proof (rule exponential_initial_condition[where Y = "cball x0 (2 * u)"]) fix s assume "s \ closed_segment t0 t2" hence s: "s \ cball t0 t" using t2 by (auto simp: dist_real_def closed_segment_eq_real_ivl split: if_split_asm) show "flow t0 x1 s \ cball x0 (2 * u)" by (rule flow_in_cball[OF s x1]) show "flow t0 x2 s \ cball x0 (2 * u)" by (rule flow_in_cball[OF s x2]) show "L-lipschitz_on (cball x0 (2 * u)) (f s)" if "s \ closed_segment t0 t2" for s using that centre_in_cball convex_contains_segment less_imp_le t2 tu(1) by (blast intro!: lipschitz) qed (fact)+ also have "\ \ dist x1 x2 * exp ((L + 1) * \t\)" using \u > 0\ t2 by (auto intro!: mult_left_mono add_nonneg_nonneg lipschitz[THEN lipschitz_on_nonneg] simp: cball_eq_empty cball_eq_sing' dist_real_def) also have "x1 \ X" using x1 subs[of x0] \u > 0\ by auto have *: "\t0 - t1\ \ t \ x \ {t0--t1} \ \t0 - x\ \ t" "\t0 - t2\ \ t \ x \ {t0--t2} \ \t0 - x\ \ t" "\t0 - t1\ \ t \ \t0 - t2\ \ t \ x \ {t1--t2} \ \t0 - x\ \ t" for x using t1 t2 t1_ex x1 flow_in_cball[OF _ x1] by (auto simp: closed_segment_eq_real_ivl split: if_splits) have integrable: "(\t. f t (flow t0 x1 t)) integrable_on {t0--t1}" "(\t. f t (flow t0 x1 t)) integrable_on {t0--t2}" "(\t. f t (flow t0 x1 t)) integrable_on {t1--t2}" using t1 t2 t1_ex x1 flow_in_cball[OF _ x1] by (auto intro!: order_trans[OF integral_bound[where B=B]] cont_on_f_flow B integrable_continuous_closed_segment intro: * simp: dist_real_def integral_minus_sets') have *: "\t0 - t1\ \ t \ \t0 - t2\ \ t \ s \ {t1--t2} \ \t0 - s\ \ t" for s by (auto simp: closed_segment_eq_real_ivl split: if_splits) note [simp] = t1_ex t2_ex \x1 \ X\ integrable have "dist (flow t0 x1 t1) (flow t0 x1 t2) \ dist t1 t2 * B" using t1 t2 x1 flow_in_cball[OF _ x1] \t0 \ T\ ivl_integral_combine[of "\t. f t (flow t0 x1 t)" t2 t0 t1] ivl_integral_combine[of "\t. f t (flow t0 x1 t)" t1 t0 t2] by (auto simp: flow_fixed_point dist_norm add.commute closed_segment_commute norm_minus_commute ivl_integral_minus_sets' ivl_integral_minus_sets intro!: order_trans[OF ivl_integral_bound[where B=B]] cont_on_f_flow B dest: *) finally have "dist (flow t0 x1 t1) (flow t0 x2 t2) \ dist t1 t2 * B + dist x1 x2 * exp ((L + 1) * \t\)" by arith also have "\ \ dist (t1, x1) (t2, x2) * B + dist (t1, x1) (t2, x2) * exp ((L + 1) * \t\)" using \B > 0\ by (auto intro!: add_mono mult_right_mono simp: dist_prod_def) finally show "dist (flow t0 x1 t1) (flow t0 x2 t2) \ (B + exp ((L + 1) * \t\)) * dist (t1, x1) (t2, x2)" by (simp add: algebra_simps) qed (simp add: \0 < B\ less_imp_le) ultimately show thesis using subs tu .. qed context fixes x0 assumes iv_defined: "t0 \ T" "x0 \ X" begin lemma existence_ivl_notempty: "existence_ivl t0 x0 \ {}" using existence_ivl_initial_time iv_defined by auto lemma initial_time_bounds: shows "bdd_above (existence_ivl t0 x0) \ t0 < Sup (existence_ivl t0 x0)" (is "?a \ _") and "bdd_below (existence_ivl t0 x0) \ Inf (existence_ivl t0 x0) < t0" (is "?b \ _") proof - from local_unique_solutions[OF iv_defined] obtain te where te: "te > 0" "cball t0 te \ existence_ivl t0 x0" by metis then show "t0 < Sup (existence_ivl t0 x0)" if bdd: "bdd_above (existence_ivl t0 x0)" using less_cSup_iff[OF existence_ivl_notempty bdd, of t0] iv_defined by (auto simp: dist_real_def intro!: bexI[where x="t0 + te"]) from te show "Inf (existence_ivl t0 x0) < t0" if bdd: "bdd_below (existence_ivl t0 x0)" unfolding cInf_less_iff[OF existence_ivl_notempty bdd, of t0] by (auto simp: dist_real_def iv_defined intro!: bexI[where x="t0 - te"]) qed lemma flow_leaves_compact_ivl_right: assumes bdd: "bdd_above (existence_ivl t0 x0)" defines "b \ Sup (existence_ivl t0 x0)" assumes "b \ T" assumes "compact K" assumes "K \ X" obtains t where "t \ t0" "t \ existence_ivl t0 x0" "flow t0 x0 t \ K" proof (atomize_elim, rule ccontr, auto) note iv_defined note ne = existence_ivl_notempty assume K[rule_format]: "\t. t \ existence_ivl t0 x0 \ t0 \ t \ flow t0 x0 t \ K" have b_upper: "t \ b" if "t \ existence_ivl t0 x0" for t unfolding b_def by (rule cSup_upper[OF that bdd]) have less_b_iff: "y < b \ (\x\existence_ivl t0 x0. y < x)" for y unfolding b_def less_cSup_iff[OF ne bdd] .. have "t0 \ b" by (simp add: iv_defined b_upper) then have geI: "t \ {t0-- t0 \ t" for t by (auto simp: half_open_segment_real) have subset: "{t0 --< b} \ existence_ivl t0 x0" using \t0 \ b\ in_existence_between_zeroI by (auto simp: half_open_segment_real iv_defined less_b_iff) have sol: "(flow t0 x0 solves_ode f) {t0 --< b} K" apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF flow_solves_ode] subset]) using subset iv_defined by (auto intro!: K geI) have cont: "continuous_on ({t0--b} \ K) (\(t, x). f t x)" using \K \ X\ closed_segment_subset_domainI[OF iv_defined(1) \b \ T\] by (auto simp: split_beta intro!: continuous_intros) from initial_time_bounds(1)[OF bdd] have "t0 \ b" by (simp add: b_def) from solves_ode_half_open_segment_continuation[OF sol cont \compact K\ \t0 \ b\] obtain l where lim: "(flow t0 x0 \ l) (at b within {t0--t. if t = b then l else flow t0 x0 t) solves_ode f) {t0--b} K" . have "b \ existence_ivl t0 x0" using \t0 \ b\ closed_segment_subset_domainI[OF \t0 \ T\ \b \ T\] by (intro existence_ivl_maximal_segment[OF solves_ode_on_subset[OF limsol order_refl \K \ X\]]) (auto simp: iv_defined) have "flow t0 x0 b \ X" by (simp add: \b \ existence_ivl t0 x0\ flow_in_domain iv_defined) from ll_on_open_it.local_unique_solutions[OF ll_on_open_it_axioms \b \ T\ \flow t0 x0 b \ X\] obtain e where "e > 0" "cball b e \ existence_ivl b (flow t0 x0 b)" by metis then have "e + b \ existence_ivl b (flow t0 x0 b)" by (auto simp: dist_real_def) from existence_ivl_trans[OF \b \ existence_ivl t0 x0\ \e + b \ existence_ivl _ _\] have "b + e \ existence_ivl t0 x0" . from b_upper[OF this] \e > 0\ show False by simp qed lemma flow_leaves_compact_ivl_left: assumes bdd: "bdd_below (existence_ivl t0 x0)" defines "b \ Inf (existence_ivl t0 x0)" assumes "b \ T" assumes "compact K" assumes "K \ X" obtains t where "t \ t0" "t \ existence_ivl t0 x0" "flow t0 x0 t \ K" proof - interpret rev: ll_on_open "(preflect t0 ` T)" "(\t. - f (preflect t0 t))" X .. from antimono_preflect bdd have bdd_rev: "bdd_above (rev.existence_ivl t0 x0)" unfolding rev_existence_ivl_eq by (rule bdd_above_image_antimono) note ne = existence_ivl_notempty have "Sup (rev.existence_ivl t0 x0) = preflect t0 b" using continuous_at_Inf_antimono[OF antimono_preflect _ ne bdd] by (simp add: continuous_preflect b_def rev_existence_ivl_eq) then have Sup_mem: "Sup (rev.existence_ivl t0 x0) \ preflect t0 ` T" using \b \ T\ by auto have rev_iv: "t0 \ preflect t0 ` T" "x0 \ X" using iv_defined by auto from rev.flow_leaves_compact_ivl_right[OF rev_iv bdd_rev Sup_mem \compact K\ \K \ X\] obtain t where "t0 \ t" "t \ rev.existence_ivl t0 x0" "rev.flow t0 x0 t \ K" . then have "preflect t0 t \ t0" "preflect t0 t \ existence_ivl t0 x0" "flow t0 x0 (preflect t0 t) \ K" by (auto simp: rev_existence_ivl_eq rev_flow_eq) thus ?thesis .. qed lemma sup_existence_maximal: assumes "\t. t0 \ t \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes "compact K" "K \ X" assumes "bdd_above (existence_ivl t0 x0)" shows "Sup (existence_ivl t0 x0) \ T" using flow_leaves_compact_ivl_right[of K] assms by force lemma inf_existence_minimal: assumes "\t. t \ t0 \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes "compact K" "K \ X" assumes "bdd_below (existence_ivl t0 x0)" shows "Inf (existence_ivl t0 x0) \ T" using flow_leaves_compact_ivl_left[of K] assms by force end lemma subset_mem_compact_implies_subset_existence_interval: assumes ivl: "t0 \ T'" "is_interval T'" "T' \ T" assumes iv_defined: "x0 \ X" assumes mem_compact: "\t. t \ T' \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes K: "compact K" "K \ X" shows "T' \ existence_ivl t0 x0" proof (rule ccontr) assume "\ T' \ existence_ivl t0 x0" then obtain t' where t': "t' \ existence_ivl t0 x0" "t' \ T'" by auto from assms have iv_defined: "t0 \ T" "x0 \ X" by auto show False proof (cases rule: not_in_connected_cases[OF connected_existence_ivl t'(1) existence_ivl_notempty[OF iv_defined]]) assume bdd: "bdd_below (existence_ivl t0 x0)" assume t'_lower: "t' \ y" if "y \ existence_ivl t0 x0" for y have i: "Inf (existence_ivl t0 x0) \ T'" using initial_time_bounds[OF iv_defined] iv_defined apply - by (rule mem_is_intervalI[of _ t' t0]) (auto simp: ivl t' bdd intro!: t'_lower cInf_greatest[OF existence_ivl_notempty[OF iv_defined]]) have *: "t \ T'" if "t \ t0" "t \ existence_ivl t0 x0" for t by (rule mem_is_intervalI[OF \is_interval T'\ i \t0 \ T'\]) (auto intro!: cInf_lower that bdd) from inf_existence_minimal[OF iv_defined mem_compact K bdd, OF *] show False using i ivl by auto next assume bdd: "bdd_above (existence_ivl t0 x0)" assume t'_upper: "y \ t'" if "y \ existence_ivl t0 x0" for y have s: "Sup (existence_ivl t0 x0) \ T'" using initial_time_bounds[OF iv_defined] apply - apply (rule mem_is_intervalI[of _ t0 t']) by (auto simp: ivl t' bdd intro!: t'_upper cSup_least[OF existence_ivl_notempty[OF iv_defined]]) have *: "t \ T'" if "t0 \ t" "t \ existence_ivl t0 x0" for t by (rule mem_is_intervalI[OF \is_interval T'\ \t0 \ T'\ s]) (auto intro!: cSup_upper that bdd) from sup_existence_maximal[OF iv_defined mem_compact K bdd, OF *] show False using s ivl by auto qed qed lemma mem_compact_implies_subset_existence_interval: assumes iv_defined: "t0 \ T" "x0 \ X" assumes mem_compact: "\t. t \ T \ t \ existence_ivl t0 x0 \ flow t0 x0 t \ K" assumes K: "compact K" "K \ X" shows "T \ existence_ivl t0 x0" by (rule subset_mem_compact_implies_subset_existence_interval; (fact | rule order_refl interval iv_defined)) lemma global_right_existence_ivl_explicit: assumes "b \ t0" assumes b: "b \ existence_ivl t0 x0" obtains d K where "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {t0 .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof - note iv_defined = mem_existence_ivl_iv_defined[OF b] define seg where "seg \ (\t. flow t0 x0 t) ` (closed_segment t0 b)" have [simp]: "x0 \ seg" by (auto simp: seg_def intro!: image_eqI[where x=t0] simp: closed_segment_eq_real_ivl iv_defined) have "seg \ {}" by (auto simp: seg_def closed_segment_eq_real_ivl) moreover have "compact seg" using iv_defined b by (auto simp: seg_def closed_segment_eq_real_ivl intro!: compact_continuous_image continuous_at_imp_continuous_on flow_continuous; metis (erased, opaque_lifting) atLeastAtMost_iff closed_segment_eq_real_ivl closed_segment_subset_existence_ivl contra_subsetD order.trans) moreover note open_domain(2) moreover have "seg \ X" using closed_segment_subset_existence_ivl b by (auto simp: seg_def intro!: flow_in_domain iv_defined) ultimately obtain e where e: "0 < e" "{x. infdist x seg \ e} \ X" thm compact_in_open_separated by (rule compact_in_open_separated) define A where "A \ {x. infdist x seg \ e}" have "A \ X" using e by (simp add: A_def) have mem_existence_ivlI: "\s. t0 \ s \ s \ b \ s \ existence_ivl t0 x0" by (rule in_existence_between_zeroI[OF b]) (auto simp: closed_segment_eq_real_ivl) have "compact A" unfolding A_def by (rule compact_infdist_le) fact+ have "compact {t0 .. b}" "{t0 .. b} \ T" subgoal by simp subgoal using mem_existence_ivlI mem_existence_ivl_subset[of _ x0] iv_defined b ivl_subset_existence_ivl by blast done from lipschitz_on_compact[OF this \compact A\ \A \ X\] obtain K' where K': "\t. t \ {t0 .. b} \ K'-lipschitz_on A (f t)" by metis define K where "K \ K' + 1" have "0 < K" "0 \ K" using assms lipschitz_on_nonneg[OF K', of t0] by (auto simp: K_def) have K: "\t. t \ {t0 .. b} \ K-lipschitz_on A (f t)" unfolding K_def using \_ \ lipschitz_on K' A _\ by (rule lipschitz_on_mono) auto have [simp]: "x0 \ A" using \0 < e\ by (auto simp: A_def) define d where "d \ min e (e * exp (-K * (b - t0)))" hence d: "0 < d" "d \ e" "d \ e * exp (-K * (b - t0))" using e by auto have d_times_exp_le: "d * exp (K * (t - t0)) \ e" if "t0 \ t" "t \ b" for t proof - from that have "d * exp (K * (t - t0)) \ d * exp (K * (b - t0))" using \0 \ K\ \0 < d\ by (auto intro!: mult_left_mono) also have "d * exp (K * (b - t0)) \ e" using d by (auto simp: exp_minus divide_simps) finally show ?thesis . qed have "ball x0 d \ X" using d \A \ X\ by (auto simp: A_def dist_commute intro!: infdist_le2[where a=x0]) note iv_defined { fix y assume y: "y \ ball x0 d" hence "y \ A" using d by (auto simp: A_def dist_commute intro!: infdist_le2[where a=x0]) hence "y \ X" using \A \ X\ by auto note y_iv = \t0 \ T\ \y \ X\ have in_A: "flow t0 y t \ A" if t: "t0 \ t" "t \ existence_ivl t0 y" "t \ b" for t proof (rule ccontr) assume flow_out: "flow t0 y t \ A" obtain t' where t': "t0 \ t'" "t' \ t" "\t. t \ {t0 .. t'} \ flow t0 x0 t \ A" "infdist (flow t0 y t') seg \ e" "\t. t \ {t0 .. t'} \ flow t0 y t \ A" proof - let ?out = "((\t. infdist (flow t0 y t) seg) -` {e..}) \ {t0..t}" have "compact ?out" unfolding compact_eq_bounded_closed proof safe show "bounded ?out" by (auto intro!: bounded_closed_interval) have "continuous_on {t0 .. t} ((\t. infdist (flow t0 y t) seg))" using closed_segment_subset_existence_ivl t iv_defined by (force intro!: continuous_at_imp_continuous_on continuous_intros flow_continuous simp: closed_segment_eq_real_ivl) thus "closed ?out" by (simp add: continuous_on_closed_vimage) qed moreover have "t \ (\t. infdist (flow t0 y t) seg) -` {e..} \ {t0..t}" using flow_out \t0 \ t\ by (auto simp: A_def) hence "?out \ {}" by blast ultimately have "\s\?out. \t\?out. s \ t" by (rule compact_attains_inf) then obtain t' where t': "\s. e \ infdist (flow t0 y s) seg \ t0 \ s \ s \ t \ t' \ s" "e \ infdist (flow t0 y t') seg" "t0 \ t'" "t' \ t" by (auto simp: vimage_def Ball_def) metis have flow_in: "flow t0 x0 s \ A" if s: "s \ {t0 .. t'}" for s proof - from s have "s \ closed_segment t0 b" using \t \ b\ t' by (auto simp: closed_segment_eq_real_ivl) then show ?thesis using s \e > 0\ by (auto simp: seg_def A_def) qed have "flow t0 y t' \ A" if "t' = t0" using y d iv_defined that by (auto simp: A_def \y \ X\ infdist_le2[where a=x0] dist_commute) moreover have "flow t0 y s \ A" if s: "s \ {t0 ..< t'}" for s proof - from s have "s \ closed_segment t0 b" using \t \ b\ t' by (auto simp: closed_segment_eq_real_ivl) from t'(1)[of s] have "t' > s \ t0 \ s \ s \ t \ e > infdist (flow t0 y s) seg" by force then show ?thesis using s t' \e > 0\ by (auto simp: seg_def A_def) qed moreover note left_of_in = this have "closed A" using \compact A\ by (auto simp: compact_eq_bounded_closed) have "((\s. flow t0 y s) \ flow t0 y t') (at_left t')" using closed_segment_subset_existence_ivl[OF t(2)] t' \y \ X\ iv_defined by (intro flow_tendsto) (auto intro!: tendsto_intros simp: closed_segment_eq_real_ivl) with \closed A\ _ _ have "t' \ t0 \ flow t0 y t' \ A" proof (rule Lim_in_closed_set) assume "t' \ t0" hence "t' > t0" using t' by auto hence "eventually (\x. x \ t0) (at_left t')" by (metis eventually_at_left less_imp_le) thus "eventually (\x. flow t0 y x \ A) (at_left t')" unfolding eventually_at_filter by eventually_elim (auto intro!: left_of_in) qed simp ultimately have flow_y_in: "s \ {t0 .. t'} \ flow t0 y s \ A" for s by (cases "s = t'"; fastforce) have "t0 \ t'" "t' \ t" "\t. t \ {t0 .. t'} \ flow t0 x0 t \ A" "infdist (flow t0 y t') seg \ e" "\t. t \ {t0 .. t'} \ flow t0 y t \ A" by (auto intro!: flow_in flow_y_in) fact+ thus ?thesis .. qed { fix s assume s: "s \ {t0 .. t'}" hence "t0 \ s" by simp have "s \ b" using t t' s b by auto hence sx0: "s \ existence_ivl t0 x0" by (simp add: \t0 \ s\ mem_existence_ivlI) have sy: "s \ existence_ivl t0 y" by (meson atLeastAtMost_iff contra_subsetD s t'(1) t'(2) that(2) ivl_subset_existence_ivl) have int: "flow t0 y s - flow t0 x0 s = y - x0 + (integral {t0 .. s} (\t. f t (flow t0 y t)) - integral {t0 .. s} (\t. f t (flow t0 x0 t)))" using iv_defined s unfolding flow_fixed_point[OF sx0] flow_fixed_point[OF sy] by (simp add: algebra_simps ivl_integral_def) have "norm (flow t0 y s - flow t0 x0 s) \ norm (y - x0) + norm (integral {t0 .. s} (\t. f t (flow t0 y t)) - integral {t0 .. s} (\t. f t (flow t0 x0 t)))" unfolding int by (rule norm_triangle_ineq) also have "norm (integral {t0 .. s} (\t. f t (flow t0 y t)) - integral {t0 .. s} (\t. f t (flow t0 x0 t))) = norm (integral {t0 .. s} (\t. f t (flow t0 y t) - f t (flow t0 x0 t)))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy by (subst Henstock_Kurzweil_Integration.integral_diff) (auto intro!: integrable_continuous_real continuous_at_imp_continuous_on f_flow_continuous simp: closed_segment_eq_real_ivl) also have "\ \ (integral {t0 .. s} (\t. norm (f t (flow t0 y t) - f t (flow t0 x0 t))))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy by (intro integral_norm_bound_integral) (auto intro!: integrable_continuous_real continuous_at_imp_continuous_on f_flow_continuous continuous_intros simp: closed_segment_eq_real_ivl) also have "\ \ (integral {t0 .. s} (\t. K * norm ((flow t0 y t) - (flow t0 x0 t))))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy iv_defined s t'(3,5) \s \ b\ by (auto simp del: Henstock_Kurzweil_Integration.integral_mult_right intro!: integral_le integrable_continuous_real continuous_at_imp_continuous_on lipschitz_on_normD[OF K] flow_continuous f_flow_continuous continuous_intros simp: closed_segment_eq_real_ivl) also have "\ = K * integral {t0 .. s} (\t. norm (flow t0 y t - flow t0 x0 t))" using closed_segment_subset_existence_ivl[of s x0] sx0 closed_segment_subset_existence_ivl[of s y] sy by (subst integral_mult) (auto intro!: integrable_continuous_real continuous_at_imp_continuous_on lipschitz_on_normD[OF K] flow_continuous f_flow_continuous continuous_intros simp: closed_segment_eq_real_ivl) finally have norm: "norm (flow t0 y s - flow t0 x0 s) \ norm (y - x0) + K * integral {t0 .. s} (\t. norm (flow t0 y t - flow t0 x0 t))" by arith note norm \s \ b\ sx0 sy } note norm_le = this from norm_le(2) t' have "t' \ closed_segment t0 b" by (auto simp: closed_segment_eq_real_ivl) hence "infdist (flow t0 y t') seg \ dist (flow t0 y t') (flow t0 x0 t')" by (auto simp: seg_def infdist_le) also have "\ \ norm (flow t0 y t' - flow t0 x0 t')" by (simp add: dist_norm) also have "\ \ norm (y - x0) * exp (K * \t' - t0\)" unfolding K_def apply (rule exponential_initial_condition[OF _ _ _ _ _ K']) subgoal by (metis atLeastAtMost_iff local.norm_le(4) order_refl t'(1)) subgoal by (metis atLeastAtMost_iff local.norm_le(3) order_refl t'(1)) subgoal using e by (simp add: A_def) subgoal by (metis closed_segment_eq_real_ivl t'(1,5)) subgoal by (metis closed_segment_eq_real_ivl t'(1,3)) subgoal by (simp add: closed_segment_eq_real_ivl local.norm_le(2) t'(1)) done also have "\ < d * exp (K * (t - t0))" using y d t' t by (intro mult_less_le_imp_less) (auto simp: dist_norm[symmetric] dist_commute intro!: mult_mono \0 \ K\) also have "\ \ e" by (rule d_times_exp_le; fact) finally have "infdist (flow t0 y t') seg < e" . with \infdist (flow t0 y t') seg \ e\ show False by (auto simp: frontier_def) qed have "{t0..b} \ existence_ivl t0 y" by (rule subset_mem_compact_implies_subset_existence_interval[OF _ is_interval_cc \{t0..b} \ T\ \y \ X\ in_A \compact A\ \A \ X\]) (auto simp: \t0 \ b\) with \t0 \ b\ have b_in: "b \ existence_ivl t0 y" by force { fix t assume t: "t \ {t0 .. b}" also have "{t0 .. b} = {t0 -- b}" by (auto simp: closed_segment_eq_real_ivl assms) also note closed_segment_subset_existence_ivl[OF b_in] finally have t_in: "t \ existence_ivl t0 y" . note t also note \{t0 .. b} = {t0 -- b}\ also note closed_segment_subset_existence_ivl[OF assms(2)] finally have t_in': "t \ existence_ivl t0 x0" . have "norm (flow t0 y t - flow t0 x0 t) \ norm (y - x0) * exp (K * \t - t0\)" unfolding K_def using t closed_segment_subset_existence_ivl[OF b_in] \0 < e\ by (intro in_A exponential_initial_condition[OF t_in t_in' \A \ X\ _ _ K']) (auto simp: closed_segment_eq_real_ivl A_def seg_def) hence "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)" by (auto simp: dist_norm[symmetric] dist_commute) } note b_in this } from \d > 0\ \K > 0\ \ball x0 d \ X\ this show ?thesis .. qed lemma global_left_existence_ivl_explicit: assumes "b \ t0" assumes b: "b \ existence_ivl t0 x0" assumes iv_defined: "t0 \ T" "x0 \ X" obtains d K where "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {b .. t0} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof - interpret rev: ll_on_open "(preflect t0 ` T)" "(\t. - f (preflect t0 t))" X .. have t0': "t0 \ preflect t0 ` T" "x0 \ X" by (auto intro!: iv_defined) from assms have "preflect t0 b \ t0" "preflect t0 b \ rev.existence_ivl t0 x0" by (auto simp: rev_existence_ivl_eq) from rev.global_right_existence_ivl_explicit[OF this] obtain d K where dK: "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ preflect t0 b \ rev.existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {t0 .. preflect t0 b} \ dist (rev.flow t0 x0 t) (rev.flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by (auto simp: rev_flow_eq \x0 \ X\) have ex_ivlI: "dist x0 y < d \ t \ existence_ivl t0 y" if "b \ t" "t \ t0" for t y using that dK(4)[of y] dK(3) iv_defined by (auto simp: subset_iff rev_existence_ivl_eq[of ] closed_segment_eq_real_ivl iv_defined in_existence_between_zeroI) have "b \ existence_ivl t0 y" if "dist x0 y < d" for y using that dK by (subst existence_ivl_eq_rev) (auto simp: iv_defined intro!: image_eqI[where x="preflect t0 b"]) with dK have "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {b .. t0} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by (auto simp: flow_eq_rev iv_defined ex_ivlI \x0 \ X\ subset_iff intro!: order_trans[OF dK(5)] image_eqI[where x="preflect t0 b"]) then show ?thesis .. qed lemma global_existence_ivl_explicit: assumes a: "a \ existence_ivl t0 x0" assumes b: "b \ existence_ivl t0 x0" assumes le: "a \ b" obtains d K where "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ a \ existence_ivl t0 y" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {a .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof - note iv_defined = mem_existence_ivl_iv_defined[OF a] define r where "r \ Max {t0, a, b}" define l where "l \ Min {t0, a, b}" have r: "r \ t0" "r \ existence_ivl t0 x0" using a b by (auto simp: max_def r_def iv_defined) obtain dr Kr where right: "0 < dr" "0 < Kr" "ball x0 dr \ X" "\y. y \ ball x0 dr \ r \ existence_ivl t0 y" "\y t. y \ ball x0 dr \ t \ {t0..r} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kr * \t - t0\)" by (rule global_right_existence_ivl_explicit[OF r]) blast have l: "l \ t0" "l \ existence_ivl t0 x0" using a b by (auto simp: min_def l_def iv_defined) obtain dl Kl where left: "0 < dl" "0 < Kl" "ball x0 dl \ X" "\y. y \ ball x0 dl \ l \ existence_ivl t0 y" "\y t. y \ ball x0 dl \ t \ {l .. t0} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kl * \t - t0\)" by (rule global_left_existence_ivl_explicit[OF l iv_defined]) blast define d where "d \ min dr dl" define K where "K \ max Kr Kl" note iv_defined have "0 < d" "0 < K" "ball x0 d \ X" using left right by (auto simp: d_def K_def) moreover { fix y assume y: "y \ ball x0 d" hence "y \ X" using \ball x0 d \ X\ by auto from y closed_segment_subset_existence_ivl[OF left(4), of y] closed_segment_subset_existence_ivl[OF right(4), of y] have "a \ existence_ivl t0 y" "b \ existence_ivl t0 y" by (auto simp: d_def l_def r_def min_def max_def closed_segment_eq_real_ivl split: if_split_asm) } moreover { fix t y assume y: "y \ ball x0 d" and t: "t \ {a .. b}" from y have "y \ X" using \ball x0 d \ X\ by auto have "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" proof cases assume "t \ t0" hence "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kr * abs (t - t0))" using y t by (intro right) (auto simp: d_def r_def) also have "exp (Kr * abs (t - t0)) \ exp (K * abs (t - t0))" by (auto simp: mult_left_mono K_def max_def mult_right_mono) finally show ?thesis by (simp add: mult_left_mono) next assume "\t \ t0" hence "dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (Kl * abs (t - t0))" using y t by (intro left) (auto simp: d_def l_def) also have "exp (Kl * abs (t - t0)) \ exp (K * abs (t - t0))" by (auto simp: mult_left_mono K_def max_def mult_right_mono) finally show ?thesis by (simp add: mult_left_mono) qed } ultimately show ?thesis .. qed lemma eventually_exponential_separation: assumes a: "a \ existence_ivl t0 x0" assumes b: "b \ existence_ivl t0 x0" assumes le: "a \ b" obtains K where "K > 0" "\\<^sub>F y in at x0. \t\{a..b}. dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)" proof - from global_existence_ivl_explicit[OF assms] obtain d K where *: "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ a \ existence_ivl t0 y" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {a .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by auto note \K > 0\ moreover have "eventually (\y. y \ ball x0 d) (at x0)" using \d > 0\[THEN eventually_at_ball] by eventually_elim simp hence "eventually (\y. \t\{a..b}. dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)) (at x0)" by eventually_elim (safe intro!: *) ultimately show ?thesis .. qed lemma eventually_mem_existence_ivl: assumes b: "b \ existence_ivl t0 x0" shows "\\<^sub>F x in at x0. b \ existence_ivl t0 x" proof - from mem_existence_ivl_iv_defined[OF b] have iv_defined: "t0 \ T" "x0 \ X" by simp_all note eiit = existence_ivl_initial_time[OF iv_defined] { fix a b assume assms: "a \ existence_ivl t0 x0" "b \ existence_ivl t0 x0" "a \ b" from global_existence_ivl_explicit[OF assms] obtain d K where *: "d > 0" "K > 0" "ball x0 d \ X" "\y. y \ ball x0 d \ a \ existence_ivl t0 y" "\y. y \ ball x0 d \ b \ existence_ivl t0 y" "\t y. y \ ball x0 d \ t \ {a .. b} \ dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * abs (t - t0))" by auto have "eventually (\y. y \ ball x0 d) (at x0)" using \d > 0\[THEN eventually_at_ball] by eventually_elim simp then have "\\<^sub>F x in at x0. a \ existence_ivl t0 x \ b \ existence_ivl t0 x" by (eventually_elim) (auto intro!: *) } from this[OF b eiit] this[OF eiit b] show ?thesis by (cases "t0 \ b") (auto simp: eventually_mono) qed lemma uniform_limit_flow: assumes a: "a \ existence_ivl t0 x0" assumes b: "b \ existence_ivl t0 x0" assumes le: "a \ b" shows "uniform_limit {a .. b} (flow t0) (flow t0 x0) (at x0)" proof (rule uniform_limitI) fix e::real assume "0 < e" from eventually_exponential_separation[OF assms] obtain K where "0 < K" "\\<^sub>F y in at x0. \t\{a..b}. dist (flow t0 x0 t) (flow t0 y t) \ dist x0 y * exp (K * \t - t0\)" by auto note this(2) moreover let ?m = "max (abs (b - t0)) (abs (a - t0))" have "eventually (\y. \t\{a..b}. dist x0 y * exp (K * \t - t0\) \ dist x0 y * exp (K * ?m)) (at x0)" using \a \ b\ \0 < K\ by (auto intro!: mult_left_mono always_eventually) moreover have "eventually (\y. dist x0 y * exp (K * ?m) < e) (at x0)" using \0 < e\ by (auto intro!: order_tendstoD tendsto_eq_intros) ultimately show "eventually (\y. \t\{a..b}. dist (flow t0 y t) (flow t0 x0 t) < e) (at x0)" by eventually_elim (force simp: dist_commute) qed lemma eventually_at_fst: assumes "eventually P (at (fst x))" assumes "P (fst x)" shows "eventually (\h. P (fst h)) (at x)" using assms unfolding eventually_at_topological by (metis open_vimage_fst rangeI range_fst vimageE vimageI) lemma eventually_at_snd: assumes "eventually P (at (snd x))" assumes "P (snd x)" shows "eventually (\h. P (snd h)) (at x)" using assms unfolding eventually_at_topological by (metis open_vimage_snd rangeI range_snd vimageE vimageI) lemma shows open_state_space: "open (Sigma X (existence_ivl t0))" and flow_continuous_on_state_space: "continuous_on (Sigma X (existence_ivl t0)) (\(x, t). flow t0 x t)" proof (safe intro!: topological_space_class.openI continuous_at_imp_continuous_on) fix t x assume "x \ X" and t: "t \ existence_ivl t0 x" have iv_defined: "t0 \ T" "x \ X" using mem_existence_ivl_iv_defined[OF t] by auto from \x \ X\ t open_existence_ivl obtain e where e: "e > 0" "cball t e \ existence_ivl t0 x" by (metis open_contains_cball) hence ivl: "t - e \ existence_ivl t0 x" "t + e \ existence_ivl t0 x" "t - e \ t + e" by (auto simp: cball_def dist_real_def) obtain d K where dK: "0 < d" "0 < K" "ball x d \ X" "\y. y \ ball x d \ t - e \ existence_ivl t0 y" "\y. y \ ball x d \ t + e \ existence_ivl t0 y" "\y s. y \ ball x d \ s \ {t - e..t + e} \ dist (flow t0 x s) (flow t0 y s) \ dist x y * exp (K * \s - t0\)" by (rule global_existence_ivl_explicit[OF ivl]) blast let ?T = "ball x d \ ball t e" have "open ?T" by (auto intro!: open_Times) moreover have "(x, t) \ ?T" by (auto simp: dK \0 < e\) moreover have "?T \ Sigma X (existence_ivl t0)" proof safe fix s y assume y: "y \ ball x d" and s: "s \ ball t e" with \ball x d \ X\ show "y \ X" by auto have "ball t e \ closed_segment t0 (t - e) \ closed_segment t0 (t + e)" by (auto simp: closed_segment_eq_real_ivl dist_real_def) with \y \ X\ s closed_segment_subset_existence_ivl[OF dK(4)[OF y]] closed_segment_subset_existence_ivl[OF dK(5)[OF y]] show "s \ existence_ivl t0 y" by auto qed ultimately show "\T. open T \ (x, t) \ T \ T \ Sigma X (existence_ivl t0)" by blast have **: "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x t) < 2 * eps" if "eps > 0" for eps :: real proof - have "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x t) = norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s) + (flow t0 x (t + snd s) - flow t0 x t))" by auto moreover have "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s) + (flow t0 x (t + snd s) - flow t0 x t)) \ norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s)) + norm (flow t0 x (t + snd s) - flow t0 x t)" by eventually_elim (rule norm_triangle_ineq) moreover have "\\<^sub>F s in at 0. t + snd s \ ball t e" by (auto simp: dist_real_def intro!: order_tendstoD[OF _ \0 < e\] intro!: tendsto_eq_intros) moreover from uniform_limit_flow[OF ivl, THEN uniform_limitD, OF \eps > 0\] have "\\<^sub>F (h::(_ )) in at (fst (0::'a*real)). \t\{t - e..t + e}. dist (flow t0 x t) (flow t0 (x + h) t) < eps" by (subst (asm) at_to_0) (auto simp: eventually_filtermap dist_commute ac_simps) hence "\\<^sub>F (h::(_ * real)) in at 0. \t\{t - e..t + e}. dist (flow t0 x t) (flow t0 (x + fst h) t) < eps" by (rule eventually_at_fst) (simp add: \eps > 0\) moreover have "\\<^sub>F h in at (snd (0::'a * _)). norm (flow t0 x (t + h) - flow t0 x t) < eps" using flow_continuous[OF t, unfolded isCont_def, THEN tendstoD, OF \eps > 0\] by (subst (asm) at_to_0) (auto simp: eventually_filtermap dist_norm ac_simps) hence "\\<^sub>F h::('a * _) in at 0. norm (flow t0 x (t + snd h) - flow t0 x t) < eps" by (rule eventually_at_snd) (simp add: \eps > 0\) ultimately show ?thesis proof eventually_elim case (elim s) note elim(1) also note elim(2) also note elim(5) also from elim(3) have "t + snd s \ {t - e..t + e}" by (auto simp: dist_real_def algebra_simps) from elim(4)[rule_format, OF this] have "norm (flow t0 (x + fst s) (t + snd s) - flow t0 x (t + snd s)) < eps" by (auto simp: dist_commute dist_norm[symmetric]) finally show ?case by simp qed qed have *: "\\<^sub>F s in at 0. norm (flow t0 (x + fst s) (t + snd s) - flow t0 x t) < eps" if "eps > 0" for eps::real proof - from that have "eps / 2 > 0" by simp from **[OF this] show ?thesis by auto qed show "isCont (\(x, y). flow t0 x y) (x, t)" unfolding isCont_iff by (rule LIM_zero_cancel) (auto simp: split_beta' norm_conv_dist[symmetric] intro!: tendstoI *) qed lemmas flow_continuous_on_compose[continuous_intros] = continuous_on_compose_Pair[OF flow_continuous_on_state_space] lemma flow_isCont_state_space: "t \ existence_ivl t0 x0 \ isCont (\(x, t). flow t0 x t) (x0, t)" using flow_continuous_on_state_space[of] mem_existence_ivl_iv_defined[of t x0] using continuous_on_eq_continuous_at open_state_space by fastforce lemma flow_absolutely_integrable_on[integrable_on_simps]: assumes "s \ existence_ivl t0 x0" shows "(\x. norm (flow t0 x0 x)) integrable_on closed_segment t0 s" using assms by (auto simp: closed_segment_eq_real_ivl intro!: integrable_continuous_real continuous_intros flow_continuous_on_intro intro: in_existence_between_zeroI) lemma existence_ivl_eq_domain: assumes iv_defined: "t0 \ T" "x0 \ X" assumes bnd: "\tm tM t x. tm \ T \ tM \ T \ \M. \L. \t \ {tm .. tM}. \x \ X. norm (f t x) \ M + L * norm x" assumes "is_interval T" "X = UNIV" shows "existence_ivl t0 x0 = T" proof - from assms have XI: "x \ X" for x by auto { fix tm tM assume tm: "tm \ T" and tM: "tM \ T" and tmtM: "tm \ t0" "t0 \ tM" from bnd[OF tm tM] obtain M' L' where bnd': "\x t. x \ X \ tm \ t \ t \ tM \ norm (f t x) \ M' + L' * norm x" by force define M where "M \ norm M' + 1" define L where "L \ norm L' + 1" have bnd: "\x t. x \ X \ tm \ t \ t \ tM \ norm (f t x) \ M + L * norm x" by (auto simp: M_def L_def intro!: bnd'[THEN order_trans] add_mono mult_mono) have "M > 0" "L > 0" by (auto simp: L_def M_def) let ?r = "(norm x0 + \tm - tM\ * M + 1) * exp (L * \tm - tM\)" define K where "K \ cball (0::'a) ?r" have K: "compact K" "K \ X" by (auto simp: K_def \X = UNIV\) { fix t assume t: "t \ existence_ivl t0 x0" and le: "tm \ t" "t \ tM" { fix s assume sc: "s \ closed_segment t0 t" then have s: "s \ existence_ivl t0 x0" and le: "tm \ s" "s \ tM" using t le sc using closed_segment_subset_existence_ivl apply - subgoal by force subgoal by (metis (full_types) atLeastAtMost_iff closed_segment_eq_real_ivl order_trans tmtM(1)) subgoal by (metis (full_types) atLeastAtMost_iff closed_segment_eq_real_ivl order_trans tmtM(2)) done from sc have nle: "norm (t0 - s) \ norm (t0 - t)" by (auto simp: closed_segment_eq_real_ivl split: if_split_asm) from flow_fixed_point[OF s] have "norm (flow t0 x0 s) \ norm x0 + integral (closed_segment t0 s) (\t. M + L * norm (flow t0 x0 t))" using tmtM using closed_segment_subset_existence_ivl[OF s] le by (auto simp: intro!: norm_triangle_le norm_triangle_ineq4[THEN order.trans] ivl_integral_norm_bound_integral bnd integrable_continuous_closed_segment integrable_continuous_real continuous_intros continuous_on_subset[OF flow_continuous_on] flow_in_domain mem_existence_ivl_subset) (auto simp: closed_segment_eq_real_ivl split: if_splits) also have "\ = norm x0 + norm (t0 - s) * M + L * integral (closed_segment t0 s) (\t. norm (flow t0 x0 t))" by (simp add: integral_add integrable_on_simps \s \ existence_ivl _ _\ integral_const_closed_segment abs_minus_commute) also have "norm (t0 - s) * M \ norm (t0 - t) * M " using nle \M > 0\ by auto also have "\ \ \ + 1" by simp finally have "norm (flow t0 x0 s) \ norm x0 + norm (t0 - t) * M + 1 + L * integral (closed_segment t0 s) (\t. norm (flow t0 x0 t))" by simp } then have "norm (flow t0 x0 t) \ (norm x0 + norm (t0 - t) * M + 1) * exp (L * \t - t0\)" using closed_segment_subset_existence_ivl[OF t] by (intro gronwall_more_general_segment[where a=t0 and b = t and t = t]) (auto simp: \0 < L\ \0 < M\ less_imp_le intro!: add_nonneg_pos mult_nonneg_nonneg add_nonneg_nonneg continuous_intros flow_continuous_on_intro) also have "\ \ ?r" using le tmtM by (auto simp: less_imp_le \0 < M\ \0 < L\ abs_minus_commute intro!: mult_mono) finally have "flow t0 x0 t \ K" by (simp add: dist_norm K_def) } note flow_compact = this have "{tm..tM} \ existence_ivl t0 x0" using tmtM tm \x0 \ X\ \compact K\ \K \ X\ mem_is_intervalI[OF \is_interval T\ \tm \ T\ \tM \ T\] by (intro subset_mem_compact_implies_subset_existence_interval[OF _ _ _ _flow_compact]) (auto simp: tmtM is_interval_cc) } note bnds = this show "existence_ivl t0 x0 = T" proof safe fix x assume x: "x \ T" from bnds[OF x iv_defined(1)] bnds[OF iv_defined(1) x] show "x \ existence_ivl t0 x0" by (cases "x \ t0") auto qed (insert existence_ivl_subset, fastforce) qed lemma flow_unique: assumes "t \ existence_ivl t0 x0" assumes "phi t0 = x0" assumes "\t. t \ existence_ivl t0 x0 \ (phi has_vector_derivative f t (phi t)) (at t)" assumes "\t. t \ existence_ivl t0 x0 \ phi t \ X" shows "flow t0 x0 t = phi t" apply (rule maximal_existence_flow[where K="existence_ivl t0 x0"]) subgoal by (auto intro!: solves_odeI simp: has_vderiv_on_def assms at_within_open[OF _ open_existence_ivl]) subgoal by fact subgoal by simp subgoal using mem_existence_ivl_iv_defined[OF \t \ existence_ivl t0 x0\] by simp subgoal by (simp add: existence_ivl_subset) subgoal by fact done lemma flow_unique_on: assumes "t \ existence_ivl t0 x0" assumes "phi t0 = x0" assumes "(phi has_vderiv_on (\t. f t (phi t))) (existence_ivl t0 x0)" assumes "\t. t \ existence_ivl t0 x0 \ phi t \ X" shows "flow t0 x0 t = phi t" using flow_unique[where phi=phi, OF assms(1,2) _ assms(4)] assms(3) by (auto simp: has_vderiv_on_open) end \ \@{thm local_lipschitz}\ locale two_ll_on_open = F: ll_on_open T1 F X + G: ll_on_open T2 G X for F T1 G T2 X J x0 + fixes e::real and K assumes t0_in_J: "0 \ J" assumes J_subset: "J \ F.existence_ivl 0 x0" assumes J_ivl: "is_interval J" assumes F_lipschitz: "\t. t \ J \ K-lipschitz_on X (F t)" assumes K_pos: "0 < K" assumes F_G_norm_ineq: "\t x. t \ J \ x \ X \ norm (F t x - G t x) < e" begin context begin lemma F_iv_defined: "0 \ T1" "x0 \ X" subgoal using F.existence_ivl_initial_time_iff J_subset t0_in_J by blast subgoal using F.mem_existence_ivl_iv_defined(2) J_subset t0_in_J by blast done lemma e_pos: "0 < e" using le_less_trans[OF norm_ge_zero F_G_norm_ineq[OF t0_in_J F_iv_defined(2)]] by assumption qualified definition "flow0 t = F.flow 0 x0 t" qualified definition "Y t = G.flow 0 x0 t" lemma norm_X_Y_bound: shows "\t \ J \ G.existence_ivl 0 x0. norm (flow0 t - Y t) \ e / K * (exp(K * \t\) - 1)" proof(safe) fix t assume "t \ J" assume tG: "t \ G.existence_ivl 0 x0" have "0 \ J" by (simp add: t0_in_J) let ?u="\t. norm (flow0 t - Y t)" show "norm (flow0 t - Y t) \ e / K * (exp (K * \t\) - 1)" proof(cases "0 \ t") assume "0 \ t" hence [simp]: "\t\ = t" by simp have t0_t_in_J: "{0..t} \ J" using \t \ J\ \0 \ J\ J_ivl using mem_is_interval_1_I atLeastAtMost_iff subsetI by blast note F_G_flow_cont[continuous_intros] = continuous_on_subset[OF F.flow_continuous_on] continuous_on_subset[OF G.flow_continuous_on] have "?u t + e/K \ e/K * exp(K * t)" proof(rule gronwall[where g="\t. ?u t + e/K", OF _ _ _ _ K_pos \0 \ t\ order.refl]) fix s assume "0 \ s" "s \ t" hence "{0..s} \ J" using t0_t_in_J by auto hence t0_s_in_existence: "{0..s} \ F.existence_ivl 0 x0" "{0..s} \ G.existence_ivl 0 x0" using J_subset tG \0 \ s\ \s \ t\ G.ivl_subset_existence_ivl[OF tG] by auto hence s_in_existence: "s \ F.existence_ivl 0 x0" "s \ G.existence_ivl 0 x0" using \0 \ s\ by auto note cont_statements[continuous_intros] = F_iv_defined (* G.iv_defined *) F.flow_in_domain G.flow_in_domain F.mem_existence_ivl_subset G.mem_existence_ivl_subset have [integrable_on_simps]: "continuous_on {0..s} (\s. F s (F.flow 0 x0 s))" "continuous_on {0..s} (\s. G s (G.flow 0 x0 s))" "continuous_on {0..s} (\s. F s (G.flow 0 x0 s))" "continuous_on {0..s} (\s. G s (F.flow 0 x0 s))" using t0_s_in_existence by (auto intro!: continuous_intros integrable_continuous_real) have "flow0 s - Y s = integral {0..s} (\s. F s (flow0 s) - G s (Y s))" using \0 \ s\ by (simp add: flow0_def Y_def Henstock_Kurzweil_Integration.integral_diff integrable_on_simps ivl_integral_def F.flow_fixed_point[OF s_in_existence(1)] G.flow_fixed_point[OF s_in_existence(2)]) also have "... = integral {0..s} (\s. (F s (flow0 s) - F s (Y s)) + (F s (Y s) - G s (Y s)))" by simp also have "... = integral {0..s} (\s. F s (flow0 s) - F s (Y s)) + integral {0..s} (\s. F s (Y s) - G s (Y s))" by (simp add: Henstock_Kurzweil_Integration.integral_diff Henstock_Kurzweil_Integration.integral_add flow0_def Y_def integrable_on_simps) finally have "?u s \ norm (integral {0..s} (\s. F s (flow0 s) - F s (Y s))) + norm (integral {0..s} (\s. F s (Y s) - G s (Y s)))" by (simp add: norm_triangle_ineq) also have "... \ integral {0..s} (\s. norm (F s (flow0 s) - F s (Y s))) + integral {0..s} (\s. norm (F s (Y s) - G s (Y s)))" using t0_s_in_existence by (auto simp add: flow0_def Y_def intro!: add_mono continuous_intros continuous_on_imp_absolutely_integrable_on) also have "... \ integral {0..s} (\s. K * ?u s) + integral {0..s} (\s. e)" proof (rule add_mono[OF integral_le integral_le]) show "norm (F x (flow0 x) - F x (Y x)) \ K * norm (flow0 x - Y x)" if "x \ {0..s}" for x using F_lipschitz[unfolded lipschitz_on_def, THEN conjunct2] that cont_statements(1,2,4) t0_s_in_existence F_iv_defined (* G.iv_defined *) by (metis F_lipschitz flow0_def Y_def \{0..s} \ J\ lipschitz_on_normD F.flow_in_domain G.flow_in_domain subsetCE) show "\x. x \ {0..s} \ norm (F x (Y x) - G x (Y x)) \ e" using F_G_norm_ineq cont_statements(2,3) t0_s_in_existence using Y_def \{0..s} \ J\ cont_statements(5) subset_iff G.flow_in_domain by (metis eucl_less_le_not_le) qed (simp_all add: t0_s_in_existence continuous_intros integrable_on_simps flow0_def Y_def) also have "... = K * integral {0..s} (\s. ?u s + e / K)" using K_pos t0_s_in_existence by (simp_all add: algebra_simps Henstock_Kurzweil_Integration.integral_add flow0_def Y_def continuous_intros continuous_on_imp_absolutely_integrable_on) finally show "?u s + e / K \ e / K + K * integral {0..s} (\s. ?u s + e / K)" by simp next show "continuous_on {0..t} (\t. norm (flow0 t - Y t) + e / K)" using t0_t_in_J J_subset G.ivl_subset_existence_ivl[OF tG] by (auto simp add: flow0_def Y_def intro!: continuous_intros) next fix s assume "0 \ s" "s \ t" show "0 \ norm (flow0 s - Y s) + e / K" using e_pos K_pos by simp next show "0 < e / K" using e_pos K_pos by simp qed thus ?thesis by (simp add: algebra_simps) next assume "\0 \ t" hence "t \ 0" by simp hence [simp]: "\t\ = -t" by simp have t0_t_in_J: "{t..0} \ J" using \t \ J\ \0 \ J\ J_ivl \\ 0 \ t\ atMostAtLeast_subset_convex is_interval_convex_1 by auto note F_G_flow_cont[continuous_intros] = continuous_on_subset[OF F.flow_continuous_on] continuous_on_subset[OF G.flow_continuous_on] have "?u t + e/K \ e/K * exp(- K * t)" proof(rule gronwall_left[where g="\t. ?u t + e/K", OF _ _ _ _ K_pos order.refl \t \ 0\]) fix s assume "t \ s" "s \ 0" hence "{s..0} \ J" using t0_t_in_J by auto hence t0_s_in_existence: "{s..0} \ F.existence_ivl 0 x0" "{s..0} \ G.existence_ivl 0 x0" using J_subset G.ivl_subset_existence_ivl'[OF tG] \s \ 0\ \t \ s\ by auto hence s_in_existence: "s \ F.existence_ivl 0 x0" "s \ G.existence_ivl 0 x0" using \s \ 0\ by auto note cont_statements[continuous_intros] = F_iv_defined F.flow_in_domain G.flow_in_domain F.mem_existence_ivl_subset G.mem_existence_ivl_subset then have [continuous_intros]: "{s..0} \ T1" "{s..0} \ T2" "F.flow 0 x0 ` {s..0} \ X" "G.flow 0 x0 ` {s..0} \ X" "s \ x \ x \ 0 \ x \ F.existence_ivl 0 x0" "s \ x \ x \ 0 \ x \ G.existence_ivl 0 x0" for x using t0_s_in_existence by auto have "flow0 s - Y s = - integral {s..0} (\s. F s (flow0 s) - G s (Y s))" using t0_s_in_existence \s \ 0\ by (simp add: flow0_def Y_def ivl_integral_def F.flow_fixed_point[OF s_in_existence(1)] G.flow_fixed_point[OF s_in_existence(2)] continuous_intros integrable_on_simps Henstock_Kurzweil_Integration.integral_diff) also have "... = - integral {s..0} (\s. (F s (flow0 s) - F s (Y s)) + (F s (Y s) - G s (Y s)))" by simp also have "... = - (integral {s..0} (\s. F s (flow0 s) - F s (Y s)) + integral {s..0} (\s. F s (Y s) - G s (Y s)))" using t0_s_in_existence by (subst Henstock_Kurzweil_Integration.integral_add) (simp_all add: integral_add flow0_def Y_def continuous_intros integrable_on_simps) finally have "?u s \ norm (integral {s..0} (\s. F s (flow0 s) - F s (Y s))) + norm (integral {s..0} (\s. F s (Y s) - G s (Y s)))" by (metis (no_types, lifting) norm_minus_cancel norm_triangle_ineq) also have "... \ integral {s..0} (\s. norm (F s (flow0 s) - F s (Y s))) + integral {s..0} (\s. norm (F s (Y s) - G s (Y s)))" using t0_s_in_existence by (auto simp add: flow0_def Y_def intro!: continuous_intros continuous_on_imp_absolutely_integrable_on add_mono) also have "... \ integral {s..0} (\s. K * ?u s) + integral {s..0} (\s. e)" proof (rule add_mono[OF integral_le integral_le]) show "norm (F x (flow0 x) - F x (Y x)) \ K * norm (flow0 x - Y x)" if "x\{s..0}" for x using F_lipschitz[unfolded lipschitz_on_def, THEN conjunct2] cont_statements(1,2,4) that t0_s_in_existence F_iv_defined (* G.iv_defined *) by (metis F_lipschitz flow0_def Y_def \{s..0} \ J\ lipschitz_on_normD F.flow_in_domain G.flow_in_domain subsetCE) show "\x. x \ {s..0} \ norm (F x (Y x) - G x (Y x)) \ e" using F_G_norm_ineq Y_def \{s..0} \ J\ cont_statements(5) subset_iff t0_s_in_existence(2) using Y_def \{s..0} \ J\ cont_statements(5) subset_iff G.flow_in_domain by (metis eucl_less_le_not_le) qed (simp_all add: t0_s_in_existence continuous_intros integrable_on_simps flow0_def Y_def) also have "... = K * integral {s..0} (\s. ?u s + e / K)" using K_pos t0_s_in_existence by (simp_all add: algebra_simps Henstock_Kurzweil_Integration.integral_add t0_s_in_existence continuous_intros integrable_on_simps flow0_def Y_def) finally show "?u s + e / K \ e / K + K * integral {s..0} (\s. ?u s + e / K)" by simp next show "continuous_on {t..0} (\t. norm (flow0 t - Y t) + e / K)" using t0_t_in_J J_subset G.ivl_subset_existence_ivl'[OF tG] F_iv_defined by (auto simp add: flow0_def Y_def intro!: continuous_intros) next fix s assume "t \ s" "s \ 0" show "0 \ norm (flow0 s - Y s) + e / K" using e_pos K_pos by simp next show "0 < e / K" using e_pos K_pos by simp qed thus ?thesis by (simp add: algebra_simps) qed qed end end locale auto_ll_on_open = fixes f::"'a::{banach, heine_borel} \ 'a" and X assumes auto_local_lipschitz: "local_lipschitz UNIV X (\_::real. f)" assumes auto_open_domain[intro!, simp]: "open X" begin text \autonomous flow and existence interval \ definition "flow0 x0 t = ll_on_open.flow UNIV (\_. f) X 0 x0 t" definition "existence_ivl0 x0 = ll_on_open.existence_ivl UNIV (\_. f) X 0 x0" sublocale ll_on_open_it UNIV "\_. f" X 0 rewrites "flow = (\t0 x0 t. flow0 x0 (t - t0))" and "existence_ivl = (\t0 x0. (+) t0 ` existence_ivl0 x0)" and "(+) 0 = (\x::real. x)" and "s - 0 = s" and "(\x. x) ` S = S" and "s \ (+) t ` S \ s - t \ (S::real set)" and "P (s + t - s) = P (t::real)"\ \TODO: why does just the equation not work?\ and "P (t + s - s) = P t"\ \TODO: why does just the equation not work?\ proof - interpret ll_on_open UNIV "\_. f" X by unfold_locales (auto intro!: continuous_on_const auto_local_lipschitz) show "ll_on_open_it UNIV (\_. f) X" .. show "(+) 0 = (\x::real. x)" "(\x. x) ` S = S" "s - 0 = s" "P (t + s - s) = P t" "P (s + t - s) = P (t::real)" by auto show "flow = (\t0 x0 t. flow0 x0 (t - t0))" unfolding flow0_def by (metis flow_def flow_shift_autonomous1 flow_shift_autonomous2 mem_existence_ivl_iv_defined(2)) show "existence_ivl = (\t0 x0. (+) t0 ` existence_ivl0 x0)" unfolding existence_ivl0_def apply (safe intro!: ext) subgoal using image_iff mem_existence_ivl_shift_autonomous1 by fastforce subgoal premises prems for t0 x0 x s proof - have f2: "\x1 x2. (x2::real) - x1 = - 1 * x1 + x2" by auto have "- 1 * t0 + (t0 + s) = s" by auto then show ?thesis using f2 prems mem_existence_ivl_iv_defined(2) mem_existence_ivl_shift_autonomous2 by presburger qed done show "(s \ (+) t ` S) = (s - t \ S)" by force qed \ \at this point, there should be no theorems about \existence_ivl\, only \existence_ivl0\. Moreover, \(+) _ ` _\ and \_ + _ - _\ etc should have been removed\ lemma existence_ivl_zero: "x0 \ X \ 0 \ existence_ivl0 x0" by simp lemmas [continuous_intros del] = continuous_on_f lemmas continuous_on_f_comp[continuous_intros] = continuous_on_f[OF continuous_on_const _ subset_UNIV] lemma flow_in_compact_right_existence: assumes "\t. 0 \ t \ t \ existence_ivl0 x \ flow0 x t \ K" assumes "compact K" "K \ X" assumes "x \ X" "t \ 0" shows "t \ existence_ivl0 x" proof (rule ccontr) assume "t \ existence_ivl0 x" have "bdd_above (existence_ivl0 x)" by (rule bdd_above_is_intervalI[OF is_interval_existence_ivl \0 \ t\ existence_ivl_zero]) fact+ from sup_existence_maximal[OF UNIV_I \x \ X\ assms(1-3) this] show False by auto qed lemma flow_in_compact_left_existence: assumes "\t. t \ 0 \ t \ existence_ivl0 x \ flow0 x t \ K" assumes "compact K" "K \ X" assumes "x \ X" "t \ 0" shows "t \ existence_ivl0 x" proof (rule ccontr) assume "t \ existence_ivl0 x" have "bdd_below (existence_ivl0 x)" by (rule bdd_below_is_intervalI[OF is_interval_existence_ivl \t \ 0\ _ existence_ivl_zero]) fact+ from inf_existence_minimal[OF UNIV_I \x \ X\ assms(1-3) this] show False by auto qed end locale compact_continuously_diff = derivative_on_prod T X f "\(t, x). f' x o\<^sub>L snd_blinfun" for T X and f::"real \ 'a::{banach,perfect_space,heine_borel} \ 'a" and f'::"'a \ ('a, 'a) blinfun" + assumes compact_domain: "compact X" assumes convex: "convex X" assumes nonempty_domains: "T \ {}" "X \ {}" assumes continuous_derivative: "continuous_on X f'" begin lemma ex_onorm_bound: "\B. \x \ X. norm (f' x) \ B" proof - from _ compact_domain have "compact (f' ` X)" by (intro compact_continuous_image continuous_derivative) hence "bounded (f' ` X)" by (rule compact_imp_bounded) thus ?thesis by (auto simp add: bounded_iff cball_def norm_blinfun.rep_eq) qed definition "onorm_bound = (SOME B. \x \ X. norm (f' x) \ B)" lemma onorm_bound: assumes "x \ X" shows "norm (f' x) \ onorm_bound" unfolding onorm_bound_def using someI_ex[OF ex_onorm_bound] assms by blast sublocale closed_domain X using compact_domain by unfold_locales (rule compact_imp_closed) sublocale global_lipschitz T X f onorm_bound proof (unfold_locales, rule lipschitz_onI) fix t z y assume "t \ T" "y \ X" "z \ X" then have "norm (f t y - f t z) \ onorm_bound * norm (y - z)" using onorm_bound by (intro differentiable_bound[where f'=f', OF convex]) (auto intro!: derivative_eq_intros simp: norm_blinfun.rep_eq) thus "dist (f t y) (f t z) \ onorm_bound * dist y z" by (auto simp: dist_norm norm_Pair) next from nonempty_domains obtain x where x: "x \ X" by auto show "0 \ onorm_bound" using dual_order.trans local.onorm_bound norm_ge_zero x by blast qed end \ \@{thm compact_domain}\ locale unique_on_compact_continuously_diff = self_mapping + compact_interval T + compact_continuously_diff T X f begin sublocale unique_on_closed t0 T x0 f X onorm_bound by unfold_locales (auto intro!: f' has_derivative_continuous_on) end locale c1_on_open = fixes f::"'a::{banach, perfect_space, heine_borel} \ 'a" and f' X assumes open_dom[simp]: "open X" assumes derivative_rhs: "\x. x \ X \ (f has_derivative blinfun_apply (f' x)) (at x)" assumes continuous_derivative: "continuous_on X f'" begin lemmas continuous_derivative_comp[continuous_intros] = continuous_on_compose2[OF continuous_derivative] lemma derivative_tendsto[tendsto_intros]: assumes [tendsto_intros]: "(g \ l) F" and "l \ X" shows "((\x. f' (g x)) \ f' l) F" using continuous_derivative[simplified continuous_on] assms by (auto simp: at_within_open[OF _ open_dom] intro!: tendsto_eq_intros intro: tendsto_compose) lemma c1_on_open_rev[intro, simp]: "c1_on_open (-f) (-f') X" using derivative_rhs continuous_derivative by unfold_locales (auto intro!: continuous_intros derivative_eq_intros simp: fun_Compl_def blinfun.bilinear_simps) lemma derivative_rhs_compose[derivative_intros]: "((g has_derivative g') (at x within s)) \ g x \ X \ ((\x. f (g x)) has_derivative (\xa. blinfun_apply (f' (g x)) (g' xa))) (at x within s)" by (metis has_derivative_compose[of g g' x s f "f' (g x)"] derivative_rhs) sublocale auto_ll_on_open proof (standard, rule local_lipschitzI) fix x and t::real assume "x \ X" with open_contains_cball[of "UNIV::real set"] open_UNIV open_contains_cball[of X] open_dom obtain u v where uv: "cball t u \ UNIV" "cball x v \ X" "u > 0" "v > 0" by blast let ?T = "cball t u" and ?X = "cball x v" have "bounded ?X" by simp have "compact (cball x v)" by simp interpret compact_continuously_diff ?T ?X "\_. f" f' using uv by unfold_locales (auto simp: convex_cball cball_eq_empty split_beta' intro!: derivative_eq_intros continuous_on_compose2[OF continuous_derivative] continuous_intros) have "onorm_bound-lipschitz_on ?X f" using lipschitz[of t] uv by auto thus "\u>0. \L. \t \ cball t u \ UNIV. L-lipschitz_on (cball x u \ X) f" by (intro exI[where x=v]) (auto intro!: exI[where x=onorm_bound] \0 < v\ simp: Int_absorb2 uv) qed (auto intro!: continuous_intros) end \ \@{thm derivative_rhs}\ locale c1_on_open_euclidean = c1_on_open f f' X for f::"'a::euclidean_space \ _" and f' X begin lemma c1_on_open_euclidean_anchor: True .. definition "vareq x0 t = f' (flow0 x0 t)" interpretation var: ll_on_open "existence_ivl0 x0" "vareq x0" UNIV apply standard apply (auto intro!: c1_implies_local_lipschitz[where f' = "\(t, x). vareq x0 t"] continuous_intros derivative_eq_intros simp: split_beta' blinfun.bilinear_simps vareq_def) using local.mem_existence_ivl_iv_defined(2) apply blast using local.existence_ivl_reverse local.mem_existence_ivl_iv_defined(2) apply blast using local.mem_existence_ivl_iv_defined(2) apply blast using local.existence_ivl_reverse local.mem_existence_ivl_iv_defined(2) apply blast done context begin lemma continuous_on_A[continuous_intros]: assumes "continuous_on S a" assumes "continuous_on S b" assumes "\s. s \ S \ a s \ X" assumes "\s. s \ S \ b s \ existence_ivl0 (a s)" shows "continuous_on S (\s. vareq (a s) (b s))" proof - have "continuous_on S (\x. f' (flow0 (a x) (b x)))" by (auto intro!: continuous_intros assms flow_in_domain) then show ?thesis by (rule continuous_on_eq) (auto simp: assms vareq_def) qed lemmas [intro] = mem_existence_ivl_iv_defined context fixes x0::'a begin lemma flow0_defined: "xa \ existence_ivl0 x0 \ flow0 x0 xa \ X" by (auto simp: flow_in_domain) lemma continuous_on_flow0: "continuous_on (existence_ivl0 x0) (flow0 x0)" by (auto simp: intro!: continuous_intros) lemmas continuous_on_flow0_comp[continuous_intros] = continuous_on_compose2[OF continuous_on_flow0] lemma varexivl_eq_exivl: assumes "t \ existence_ivl0 x0" shows "var.existence_ivl x0 t a = existence_ivl0 x0" proof (rule var.existence_ivl_eq_domain) fix s t x assume s: "s \ existence_ivl0 x0" and t: "t \ existence_ivl0 x0" then have "{s .. t} \ existence_ivl0 x0" by (metis atLeastatMost_empty_iff2 empty_subsetI real_Icc_closed_segment var.closed_segment_subset_domain) then have "continuous_on {s .. t} (vareq x0)" by (auto simp: closed_segment_eq_real_ivl intro!: continuous_intros flow0_defined) then have "compact ((vareq x0) ` {s .. t})" using compact_Icc by (rule compact_continuous_image) then obtain B where B: "\u. u \ {s .. t} \ norm (vareq x0 u) \ B" by (force dest!: compact_imp_bounded simp: bounded_iff) show "\M L. \t\{s..t}. \x\UNIV. norm (blinfun_apply (vareq x0 t) x) \ M + L * norm x" by (rule exI[where x=0], rule exI[where x=B]) (auto intro!: order_trans[OF norm_blinfun] mult_right_mono B simp:) qed (auto intro: assms) definition "vector_Dflow u0 t \ var.flow x0 0 u0 t" qualified abbreviation "Y z t \ flow0 (x0 + z) t" text \Linearity of the solution to the variational equation. TODO: generalize this and some other things for arbitrary linear ODEs\ lemma vector_Dflow_linear: assumes "t \ existence_ivl0 x0" shows "vector_Dflow (\ *\<^sub>R a + \ *\<^sub>R b) t = \ *\<^sub>R vector_Dflow a t + \ *\<^sub>R vector_Dflow b t" proof - note mem_existence_ivl_iv_defined[OF assms, intro, simp] have "((\c. \ *\<^sub>R var.flow x0 0 a c + \ *\<^sub>R var.flow x0 0 b c) solves_ode (\x. vareq x0 x)) (existence_ivl0 x0) UNIV" by (auto intro!: derivative_intros var.flow_has_vector_derivative solves_odeI simp: blinfun.bilinear_simps varexivl_eq_exivl vareq_def[symmetric]) moreover have "\ *\<^sub>R var.flow x0 0 a 0 + \ *\<^sub>R var.flow x0 0 b 0 = \ *\<^sub>R a + \ *\<^sub>R b" by simp moreover note is_interval_existence_ivl[of x0] ultimately show ?thesis unfolding vareq_def[symmetric] vector_Dflow_def by (rule var.maximal_existence_flow) (auto simp: assms) qed lemma linear_vector_Dflow: assumes "t \ existence_ivl0 x0" shows "linear (\z. vector_Dflow z t)" using vector_Dflow_linear[OF assms, of 1 _ 1] vector_Dflow_linear[OF assms, of _ _ 0] by (auto intro!: linearI) lemma bounded_linear_vector_Dflow: assumes "t \ existence_ivl0 x0" shows "bounded_linear (\z. vector_Dflow z t)" by (simp add: linear_linear linear_vector_Dflow assms) lemma vector_Dflow_continuous_on_time: "x0 \ X \ continuous_on (existence_ivl0 x0) (\t. vector_Dflow z t)" using var.flow_continuous_on[of x0 0 z] varexivl_eq_exivl unfolding vector_Dflow_def by (auto simp: ) proposition proposition_17_6_weak: \ \from "Differential Equations, Dynamical Systems, and an Introduction to Chaos", Hirsch/Smale/Devaney\ assumes "t \ existence_ivl0 x0" shows "(\y. (Y (y - x0) t - flow0 x0 t - vector_Dflow (y - x0) t) /\<^sub>R norm (y - x0)) \ x0 \ 0" proof- note x0_def = mem_existence_ivl_iv_defined[OF assms] have "0 \ existence_ivl0 x0" by (simp add: x0_def) text \Find some \J \ existence_ivl0 x0\ with \0 \ J\ and \t \ J\.\ define t0 where "t0 \ min 0 t" define t1 where "t1 \ max 0 t" define J where "J \ {t0..t1}" have "t0 \ 0" "0 \ t1" "0 \ J" "J \ {}" "t \ J" "compact J" and J_in_existence: "J \ existence_ivl0 x0" using ivl_subset_existence_ivl ivl_subset_existence_ivl' x0_def assms by (auto simp add: J_def t0_def t1_def min_def max_def) { fix z S assume assms: "x0 + z \ X" "S \ existence_ivl0 (x0 + z)" have "continuous_on S (Y z)" using flow_continuous_on assms(1) by (intro continuous_on_subset[OF _ assms(2)]) (simp add:) } note [continuous_intros] = this integrable_continuous_real blinfun.continuous_on have U_continuous[continuous_intros]: "\z. continuous_on J (vector_Dflow z)" by(rule continuous_on_subset[OF vector_Dflow_continuous_on_time[OF \x0 \ X\] J_in_existence]) from \t \ J\ have "t0 \ t" and "t \ t1" and "t0 \ t1" and "t0 \ existence_ivl0 x0" and "t \ existence_ivl0 x0" and "t1 \ existence_ivl0 x0" using J_def J_in_existence by auto from global_existence_ivl_explicit[OF \t0 \ existence_ivl0 x0\ \t1 \ existence_ivl0 x0\ \t0 \ t1\] obtain u K where uK_def: "0 < u" "0 < K" "ball x0 u \ X" "\y. y \ ball x0 u \ t0 \ existence_ivl0 y" "\y. y \ ball x0 u \ t1 \ existence_ivl0 y" "\t y. y \ ball x0 u \ t \ J \ dist (flow0 x0 t) (Y (y - x0) t) \ dist x0 y * exp (K * \t\)" by (auto simp add: J_def) have J_in_existence_ivl: "\y. y \ ball x0 u \ J \ existence_ivl0 y" unfolding J_def using uK_def by (simp add: real_Icc_closed_segment segment_subset_existence_ivl t0_def t1_def) have ball_in_X: "\z. z \ ball 0 u \ x0 + z \ X" using uK_def(3) by (auto simp: dist_norm) have flow0_J_props: "flow0 x0 ` J \ {}" "compact (flow0 x0 ` J)" "flow0 x0` J \ X" using \t0 \ t1\ using J_def(1) J_in_existence by (auto simp add: J_def intro!: compact_continuous_image continuous_intros flow_in_domain) have [continuous_intros]: "continuous_on J (\s. f' (flow0 x0 s))" using J_in_existence by (auto intro!: continuous_intros flow_in_domain simp:) text \ Show the thesis via cases \t = 0\, \0 < t\ and \t < 0\. \ show ?thesis proof(cases "t = 0") assume "t = 0" show ?thesis unfolding \t = 0\ Lim_at proof(simp add: dist_norm[of _ 0] del: zero_less_dist_iff, safe, rule exI, rule conjI[OF \0 < u\], safe) fix e::real and x assume "0 < e" "dist x x0 < u" hence "x \ X" using uK_def(3) by (auto simp: dist_commute) hence "inverse (norm (x - x0)) * norm (Y (x - x0) 0 - flow0 x0 0 - vector_Dflow (x - x0) 0) = 0" using x0_def by (simp add: vector_Dflow_def) thus "inverse (norm (x - x0)) * norm (flow0 x 0 - flow0 x0 0 - vector_Dflow (x - x0) 0) < e" using \0 < e\ by auto qed next assume "t \ 0" show ?thesis proof(unfold Lim_at, safe) fix e::real assume "0 < e" then obtain e' where "0 < e'" "e' < e" using dense by auto obtain N where N_ge_SupS: "Sup { norm (f' (flow0 x0 s)) |s. s \ J } \ N" (is "Sup ?S \ N") and N_gr_0: "0 < N" \ \We need N to be an upper bound of @{term ?S}, but also larger than zero.\ by (meson le_cases less_le_trans linordered_field_no_ub) have N_ineq: "\s. s \ J \ norm (f' (flow0 x0 s)) \ N" proof- fix s assume "s \ J" have "?S = (norm o f' o flow0 x0) ` J" by auto moreover have "continuous_on J (norm o f' o flow0 x0)" using J_in_existence by (auto intro!: continuous_intros) ultimately have "\a b. ?S = {a..b} \ a \ b" using continuous_image_closed_interval[OF \t0 \ t1\] by (simp add: J_def) then obtain a b where "?S = {a..b}" and "a \ b" by auto hence "bdd_above ?S" by simp from \s \ J\ cSup_upper[OF _ this] have "norm (f' (flow0 x0 s)) \ Sup ?S" by auto thus "norm (f' (flow0 x0 s)) \ N" using N_ge_SupS by simp qed text \ Define a small region around \flow0 ` J\, that is a subset of the domain \X\. \ from compact_in_open_separated[OF flow0_J_props(1,2) auto_open_domain flow0_J_props(3)] obtain e_domain where e_domain_def: "0 < e_domain" "{x. infdist x (flow0 x0 ` J) \ e_domain} \ X" by auto define G where "G \ {x\X. infdist x (flow0 x0 ` J) < e_domain}" have G_vimage: "G = ((\x. infdist x (flow0 x0 ` J)) -` {.. X" by (auto simp: G_def) have "open G" "G \ X" unfolding G_vimage by (auto intro!: open_Int open_vimage continuous_intros continuous_at_imp_continuous_on) text \Define a compact subset H of G. Inside H, we can guarantee an upper bound on the Taylor remainder.\ define e_domain2 where "e_domain2 \ e_domain / 2" have "e_domain2 > 0" "e_domain2 < e_domain" using \e_domain > 0\ by (simp_all add: e_domain2_def) define H where "H \ {x. infdist x (flow0 x0 ` J) \ e_domain2}" have H_props: "H \ {}" "compact H" "H \ G" proof- have "x0 \ flow0 x0 ` J" unfolding image_iff using \0 \ J\ x0_def by force hence "x0 \ H" using \0 < e_domain2\ by (simp add: H_def x0_def) thus "H \ {}" by auto next show "compact H" unfolding H_def using \0 < e_domain2\ flow0_J_props by (intro compact_infdist_le) simp_all next show "H \ G" proof fix x assume "x \ H" then have *: "infdist x (flow0 x0 ` J) < e_domain" using \0 < e_domain\ by (simp add: H_def e_domain2_def) then have "x \ X" using e_domain_def(2) by auto with * show "x \ G" unfolding G_def by auto qed qed have f'_cont_on_G: "(\x. x \ G \ isCont f' x)" using continuous_on_interior[OF continuous_on_subset[OF continuous_derivative \G \ X\]] by (simp add: interior_open[OF \open G\]) define e1 where "e1 \ e' / (\t\ * exp (K * \t\) * exp (N * \t\))" \ \@{term e1} is the bounding term for the Taylor remainder.\ have "0 < \t\" using \t \ 0\ by simp hence "0 < e1" using \0 < e'\ by (simp add: e1_def) text \ Taylor expansion of f on set G. \ from uniform_explicit_remainder_Taylor_1[where f=f and f'=f', OF derivative_rhs[OF subsetD[OF \G \ X\]] f'_cont_on_G \open G\ H_props \0 < e1\] obtain d_Taylor R where Taylor_expansion: "0 < d_Taylor" "\x z. f z = f x + (f' x) (z - x) + R x z" "\x y. x \ H \ y \ H \ dist x y < d_Taylor \ norm (R x y) \ e1 * dist x y" "continuous_on (G \ G) (\(a, b). R a b)" by auto text \ Find d, such that solutions are always at least \min (e_domain/2) d_Taylor\ apart, i.e. always in H. This later gives us the bound on the remainder. \ have "0 < min (e_domain/2) d_Taylor" using \0 < d_Taylor\ \0 < e_domain\ by auto from uniform_limit_flow[OF \t0 \ existence_ivl0 x0\ \t1 \ existence_ivl0 x0\ \t0 \ t1\, THEN uniform_limitD, OF this, unfolded eventually_at] obtain d_ivl where d_ivl_def: "0 < d_ivl" "\x. 0 < dist x x0 \ dist x x0 < d_ivl \ (\t\J. dist (flow0 x0 t) (Y (x - x0) t) < min (e_domain / 2) d_Taylor)" by (auto simp: dist_commute J_def) define d where "d \ min u d_ivl" have "0 < d" using \0 < u\ \0 < d_ivl\ by (simp add: d_def) hence "d \ u" "d \ d_ivl" by (auto simp: d_def) text \ Therefore, any flow0 starting in \ball x0 d\ will be in G. \ have Y_in_G: "\y. y \ ball x0 d \ (\s. Y (y - x0) s) ` J \ G" proof fix x y assume assms: "y \ ball x0 d" "x \ (\s. Y (y - x0) s) ` J" show "x \ G" proof(cases) assume "y = x0" from assms(2) have "x \ flow0 x0 ` J" by (simp add: \y = x0\) thus "x \ G" using \0 < e_domain\ \flow0 x0 ` J \ X\ by (auto simp: G_def) next assume "y \ x0" hence "0 < dist y x0" by (simp add: dist_norm) from d_ivl_def(2)[OF this] \d \ d_ivl\ \0 < e_domain\ assms(1) have dist_flow0_Y: "\t. t \ J \ dist (flow0 x0 t) (Y (y - x0) t) < e_domain" by (auto simp: dist_commute) from assms(2) obtain t where t_def: "t \ J" "x = Y (y - x0) t" by auto have "x \ X" unfolding t_def(2) using uK_def(3) assms(1) \d \ u\ subsetD[OF J_in_existence_ivl t_def(1)] by (auto simp: intro!: flow_in_domain) have "flow0 x0 t \ flow0 x0 ` J" using t_def by auto from dist_flow0_Y[OF t_def(1)] have "dist x (flow0 x0 t) < e_domain" by (simp add: t_def(2) dist_commute) from le_less_trans[OF infdist_le[OF \flow0 x0 t \ flow0 x0 ` J\] this] \x \ X\ show "x \ G" by (auto simp: G_def) qed qed from this[of x0] \0 < d\ have X_in_G: "flow0 x0 ` J \ G" by simp show "\d>0. \x. 0 < dist x x0 \ dist x x0 < d \ dist ((Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) /\<^sub>R norm (x - x0)) 0 < e" proof(rule exI, rule conjI[OF \0 < d\], safe, unfold norm_conv_dist[symmetric]) fix x assume x_x0_dist: "0 < dist x x0" "dist x x0 < d" hence x_in_ball': "x \ ball x0 d" by (simp add: dist_commute) hence x_in_ball: "x \ ball x0 u" using \d \ u\ by simp text \ First, some prerequisites. \ from x_in_ball have z_in_ball: "x - x0 \ ball 0 u" using \0 < u\ by (simp add: dist_norm) hence [continuous_intros]: "dist x0 x < u" by (auto simp: dist_norm) from J_in_existence_ivl[OF x_in_ball] have J_in_existence_ivl_x: "J \ existence_ivl0 x" . from ball_in_X[OF z_in_ball] have x_in_X[continuous_intros]: "x \ X" by simp text \ On all of \J\, we can find upper bounds for the distance of \flow0\ and \Y\. \ have dist_flow0_Y: "\s. s \ J \ dist (flow0 x0 s) (Y (x - x0) s) \ dist x0 x * exp (K * \t\)" using t0_def t1_def uK_def(2) by (intro order_trans[OF uK_def(6)[OF x_in_ball] mult_left_mono]) (auto simp add: J_def intro!: mult_mono) from d_ivl_def x_x0_dist \d \ d_ivl\ have dist_flow0_Y2: "\t. t \ J \ dist (flow0 x0 t) (Y (x - x0) t) < min (e_domain2) d_Taylor" by (auto simp: e_domain2_def) let ?g = "\t. norm (Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t)" let ?C = "\t\ * dist x0 x * exp (K * \t\) * e1" text \ Find an upper bound to \?g\, i.e. show that \?g s \ ?C + N * integral {a..b} ?g\ for \{a..b} = {0..s}\ or \{a..b} = {s..0}\ for some \s \ J\. We can then apply Grönwall's inequality to obtain a true bound for \?g\. \ have g_bound: "?g s \ ?C + N * integral {a..b} ?g" if s_def: "s \ {a..b}" and J'_def: "{a..b} \ J" and ab_cases: "(a = 0 \ b = s) \ (a = s \ b = 0)" for s a b proof - from that have "s \ J" by auto have s_in_existence_ivl_x0: "s \ existence_ivl0 x0" using J_in_existence \s \ J\ by auto have s_in_existence_ivl: "\y. y \ ball x0 u \ s \ existence_ivl0 y" using J_in_existence_ivl \s \ J\ by auto have s_in_existence_ivl2: "\z. z \ ball 0 u \ s \ existence_ivl0 (x0 + z)" using s_in_existence_ivl by (simp add: dist_norm) text \Prove continuities beforehand.\ note continuous_on_0_s[continuous_intros] = continuous_on_subset[OF _ \{a..b} \ J\] have[continuous_intros]: "continuous_on J (flow0 x0)" using J_in_existence by (auto intro!: continuous_intros simp:) { fix z S assume assms: "x0 + z \ X" "S \ existence_ivl0 (x0 + z)" have "continuous_on S (\s. f (Y z s))" proof(rule continuous_on_subset[OF _ assms(2)]) show "continuous_on (existence_ivl0 (x0 + z)) (\s. f (Y z s))" using assms by (auto intro!: continuous_intros flow_in_domain flow_continuous_on simp:) qed } note [continuous_intros] = this have [continuous_intros]: "continuous_on J (\s. f (flow0 x0 s))" by(rule continuous_on_subset[OF _ J_in_existence]) (auto intro!: continuous_intros flow_continuous_on flow_in_domain simp: x0_def) have [continuous_intros]: "\z. continuous_on J (\s. f' (flow0 x0 s) (vector_Dflow z s))" proof- fix z have a1: "continuous_on J (flow0 x0)" by (auto intro!: continuous_intros) have a2: "(\s. (flow0 x0 s, vector_Dflow z s)) ` J \ (flow0 x0 ` J) \ ((\s. vector_Dflow z s) ` J)" by auto have a3: "continuous_on ((\s. (flow0 x0 s, vector_Dflow z s)) ` J) (\(x, u). f' x u)" using assms flow0_J_props by (auto intro!: continuous_intros simp: split_beta') from continuous_on_compose[OF continuous_on_Pair[OF a1 U_continuous] a3] show "continuous_on J (\s. f' (flow0 x0 s) (vector_Dflow z s))" by simp qed have [continuous_intros]: "continuous_on J (\s. R (flow0 x0 s) (Y (x - x0) s))" using J_in_existence J_in_existence_ivl[OF x_in_ball] X_in_G \{a..b} \ J\ Y_in_G x_x0_dist by (auto intro!: continuous_intros continuous_on_compose_Pair[OF Taylor_expansion(4)] simp: dist_commute subset_iff) hence [continuous_intros]: "(\s. R (flow0 x0 s) (Y (x - x0) s)) integrable_on J" unfolding J_def by (rule integrable_continuous_real) have i1: "integral {a..b} (\s. f (flow0 x s)) - integral {a..b} (\s. f (flow0 x0 s)) = integral {a..b} (\s. f (flow0 x s) - f (flow0 x0 s))" using J_in_existence_ivl[OF x_in_ball] apply (intro Henstock_Kurzweil_Integration.integral_diff[symmetric]) by (auto intro!: continuous_intros existence_ivl_reverse) have i2: "integral {a..b} (\s. f (flow0 x s) - f (flow0 x0 s) - (f' (flow0 x0 s)) (vector_Dflow (x - x0) s)) = integral {a..b} (\s. f (flow0 x s) - f (flow0 x0 s)) - integral {a..b} (\s. f' (flow0 x0 s) (vector_Dflow (x - x0) s))" using J_in_existence_ivl[OF x_in_ball] by (intro Henstock_Kurzweil_Integration.integral_diff[OF Henstock_Kurzweil_Integration.integrable_diff]) (auto intro!: continuous_intros existence_ivl_reverse) from ab_cases have "?g s = norm (integral {a..b} (\s'. f (Y (x - x0) s')) - integral {a..b} (\s'. f (flow0 x0 s')) - integral {a..b} (\s'. (f' (flow0 x0 s')) (vector_Dflow (x - x0) s')))" proof(safe) assume "a = 0" "b = s" hence "0 \ s" using \s \ {a..b}\ by simp text\ Integral equations for flow0, Y and U. \ have flow0_integral_eq: "flow0 x0 s = x0 + ivl_integral 0 s (\s. f (flow0 x0 s))" by (rule flow_fixed_point[OF s_in_existence_ivl_x0]) have Y_integral_eq: "flow0 x s = x0 + (x - x0) + ivl_integral 0 s (\s. f (Y (x - x0) s))" using flow_fixed_point \0 \ s\ s_in_existence_ivl2[OF z_in_ball] ball_in_X[OF z_in_ball] by (simp add:) have U_integral_eq: "vector_Dflow (x - x0) s = (x - x0) + ivl_integral 0 s (\s. vareq x0 s (vector_Dflow (x - x0) s))" unfolding vector_Dflow_def by (rule var.flow_fixed_point) (auto simp: \0 \ s\ x0_def varexivl_eq_exivl s_in_existence_ivl_x0) show "?g s = norm (integral {0..s} (\s'. f (Y (x - x0) s')) - integral {0..s} (\s'. f (flow0 x0 s')) - integral {0..s} (\s'. blinfun_apply (f' (flow0 x0 s')) (vector_Dflow (x - x0) s')))" using \0 \ s\ unfolding vareq_def[symmetric] by (simp add: flow0_integral_eq Y_integral_eq U_integral_eq ivl_integral_def) next assume "a = s" "b = 0" hence "s \ 0" using \s \ {a..b}\ by simp have flow0_integral_eq_left: "flow0 x0 s = x0 + ivl_integral 0 s (\s. f (flow0 x0 s))" by (rule flow_fixed_point[OF s_in_existence_ivl_x0]) have Y_integral_eq_left: "Y (x - x0) s = x0 + (x - x0) + ivl_integral 0 s (\s. f (Y (x - x0) s))" using flow_fixed_point \s \ 0\ s_in_existence_ivl2[OF z_in_ball] ball_in_X[OF z_in_ball] by simp have U_integral_eq_left: "vector_Dflow (x - x0) s = (x - x0) + ivl_integral 0 s (\s. vareq x0 s (vector_Dflow (x - x0) s))" unfolding vector_Dflow_def by (rule var.flow_fixed_point) (auto simp: \s \ 0\ x0_def varexivl_eq_exivl s_in_existence_ivl_x0) have "?g s = norm (- integral {s..0} (\s'. f (Y (x - x0) s')) + integral {s..0} (\s'. f (flow0 x0 s')) + integral {s..0} (\s'. vareq x0 s' (vector_Dflow (x - x0) s')))" unfolding flow0_integral_eq_left Y_integral_eq_left U_integral_eq_left using \s \ 0\ by (simp add: ivl_integral_def) also have "... = norm (integral {s..0} (\s'. f (Y (x - x0) s')) - integral {s..0} (\s'. f (flow0 x0 s')) - integral {s..0} (\s'. vareq x0 s' (vector_Dflow (x - x0) s')))" by (subst norm_minus_cancel[symmetric], simp) finally show "?g s = norm (integral {s..0} (\s'. f (Y (x - x0) s')) - integral {s..0} (\s'. f (flow0 x0 s')) - integral {s..0} (\s'. blinfun_apply (f' (flow0 x0 s')) (vector_Dflow (x - x0) s')))" unfolding vareq_def . qed also have "... = norm (integral {a..b} (\s. f (Y (x - x0) s) - f (flow0 x0 s) - (f' (flow0 x0 s)) (vector_Dflow (x - x0) s)))" by (simp add: i1 i2) also have "... \ integral {a..b} (\s. norm (f (Y (x - x0) s) - f (flow0 x0 s) - f' (flow0 x0 s) (vector_Dflow (x - x0) s)))" using x_in_X J_in_existence_ivl_x J_in_existence \{a..b} \ J\ by (auto intro!: continuous_intros continuous_on_imp_absolutely_integrable_on existence_ivl_reverse) also have "... = integral {a..b} (\s. norm (f' (flow0 x0 s) (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s) + R (flow0 x0 s) (Y (x - x0) s)))" proof (safe intro!: integral_spike[OF negligible_empty, simplified] arg_cong[where f=norm]) fix s' assume "s' \ {a..b}" show "f' (flow0 x0 s') (Y (x - x0) s' - flow0 x0 s' - vector_Dflow (x - x0) s') + R (flow0 x0 s') (Y (x - x0) s') = f (Y (x - x0) s') - f (flow0 x0 s') - f' (flow0 x0 s') (vector_Dflow (x - x0) s')" by (simp add: blinfun.diff_right Taylor_expansion(2)[of "flow0 x s'" "flow0 x0 s'"]) qed also have "... \ integral {a..b} (\s. norm (f' (flow0 x0 s) (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s)) + norm (R (flow0 x0 s) (Y (x - x0) s)))" using J_in_existence_ivl[OF x_in_ball] norm_triangle_ineq using \continuous_on J (\s. R (flow0 x0 s) (Y (x - x0) s))\ by (auto intro!: continuous_intros integral_le) also have "... = integral {a..b} (\s. norm (f' (flow0 x0 s) (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s))) + integral {a..b} (\s. norm (R (flow0 x0 s) (Y (x - x0) s)))" using J_in_existence_ivl[OF x_in_ball] using \continuous_on J (\s. R (flow0 x0 s) (Y (x - x0) s))\ by (auto intro!: continuous_intros Henstock_Kurzweil_Integration.integral_add) also have "... \ N * integral {a..b} ?g + ?C" (is "?l1 + ?r1 \ _") proof(rule add_mono) have "?l1 \ integral {a..b} (\s. norm (f' (flow0 x0 s)) * norm (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s))" using norm_blinfun J_in_existence_ivl[OF x_in_ball] by (auto intro!: continuous_intros integral_le) also have "... \ integral {a..b} (\s. N * norm (Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s))" using J_in_existence_ivl[OF x_in_ball] N_ineq[OF \{a..b} \ J\[THEN subsetD]] by (intro integral_le) (auto intro!: continuous_intros mult_right_mono) also have "... = N * integral {a..b} (\s. norm ((Y (x - x0) s - flow0 x0 s - vector_Dflow (x - x0) s)))" unfolding real_scaleR_def[symmetric] by(rule integral_cmul) finally show "?l1 \ N * integral {a..b} ?g" . next have "?r1 \ integral {a..b} (\s. e1 * dist (flow0 x0 s) (Y (x - x0) s))" using J_in_existence_ivl[OF x_in_ball] \0 < e_domain\ dist_flow0_Y2 \0 < e_domain2\ by (intro integral_le) (force intro!: continuous_intros Taylor_expansion(3) order_trans[OF infdist_le] dest!: \{a..b} \ J\[THEN subsetD] intro: less_imp_le simp: dist_commute H_def)+ also have "... \ integral {a..b} (\s. e1 * (dist x0 x * exp (K * \t\)))" apply(rule integral_le) subgoal using J_in_existence_ivl[OF x_in_ball] by (force intro!: continuous_intros) subgoal by force subgoal by (force dest!: \{a..b} \ J\[THEN subsetD] intro!: less_imp_le[OF \0 < e1\] mult_left_mono[OF dist_flow0_Y]) done also have "... \ ?C" using \s \ J\ x_x0_dist \0 < e1\ \{a..b} \ J\ \0 < \t\\ t0_def t1_def by (auto simp: integral_const_real J_def(1)) finally show "?r1 \ ?C" . qed finally show ?thesis by simp qed have g_continuous: "continuous_on J ?g" using J_in_existence_ivl[OF x_in_ball] J_in_existence using J_def(1) U_continuous by (auto simp: J_def intro!: continuous_intros) note [continuous_intros] = continuous_on_subset[OF g_continuous] have C_gr_zero: "0 < ?C" using \0 < \t\\ \0 < e1\ x_x0_dist(1) by (simp add: dist_commute) have "0 \ t \ t \ 0" by auto then have "?g t \ ?C * exp (N * \t\)" proof assume "0 \ t" moreover have "continuous_on {0..t} (vector_Dflow (x - x0))" using U_continuous by (rule continuous_on_subset) (auto simp: J_def t0_def t1_def) then have "norm (Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) \ \t\ * dist x0 x * exp (K * \t\) * e1 * exp (N * t)" using \t \ J\ J_def \t0 \ 0\ J_in_existence J_in_existence_ivl_x by (intro gronwall[OF g_bound _ _ C_gr_zero \0 < N\ \0 \ t\ order.refl]) (auto intro!: continuous_intros simp: ) ultimately show ?thesis by simp next assume "t \ 0" moreover have "continuous_on {t .. 0} (vector_Dflow (x - x0))" using U_continuous by (rule continuous_on_subset) (auto simp: J_def t0_def t1_def) then have "norm (Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) \ \t\ * dist x0 x * exp (K * \t\) * e1 * exp (- N * t)" using \t \ J\ J_def \0 \ t1\ J_in_existence J_in_existence_ivl_x by (intro gronwall_left[OF g_bound _ _ C_gr_zero \0 < N\ order.refl \t \ 0\]) (auto intro!: continuous_intros) ultimately show ?thesis by simp qed also have "... = dist x x0 * (\t\ * exp (K * \t\) * e1 * exp (N * \t\))" by (auto simp: dist_commute) also have "... < norm (x - x0) * e" unfolding e1_def using \e' < e\ \0 < \t\\ \0 < e1\ x_x0_dist(1) by (simp add: dist_norm) finally show "norm ((Y (x - x0) t - flow0 x0 t - vector_Dflow (x - x0) t) /\<^sub>R norm (x - x0)) < e" by (simp, metis x_x0_dist(1) dist_norm divide_inverse mult.commute pos_divide_less_eq) qed qed qed qed lemma local_lipschitz_A: "OT \ existence_ivl0 x0 \ local_lipschitz OT (OS::('a \\<^sub>L 'a) set) (\t. (o\<^sub>L) (vareq x0 t))" by (rule local_lipschitz_subset[OF _ _ subset_UNIV, where T="existence_ivl0 x0"]) (auto simp: split_beta' vareq_def intro!: c1_implies_local_lipschitz[where f'="\(t, x). comp3 (f' (flow0 x0 t))"] derivative_eq_intros blinfun_eqI ext continuous_intros flow_in_domain) lemma total_derivative_ll_on_open: "ll_on_open (existence_ivl0 x0) (\t. blinfun_compose (vareq x0 t)) (UNIV::('a \\<^sub>L 'a) set)" by standard (auto intro!: continuous_intros local_lipschitz_A[OF order_refl]) end end sublocale mvar: ll_on_open "existence_ivl0 x0" "\t. blinfun_compose (vareq x0 t)" "UNIV::('a \\<^sub>L 'a) set" for x0 by (rule total_derivative_ll_on_open) lemma mvar_existence_ivl_eq_existence_ivl[simp]:\ \TODO: unify with @{thm varexivl_eq_exivl}\ assumes "t \ existence_ivl0 x0" shows "mvar.existence_ivl x0 t = (\_. existence_ivl0 x0)" proof (rule ext, rule mvar.existence_ivl_eq_domain) fix s t x assume s: "s \ existence_ivl0 x0" and t: "t \ existence_ivl0 x0" then have "{s .. t} \ existence_ivl0 x0" by (meson atLeastAtMost_iff is_interval_1 is_interval_existence_ivl subsetI) then have "continuous_on {s .. t} (vareq x0)" by (auto intro!: continuous_intros) then have "compact (vareq x0 ` {s .. t})" using compact_Icc by (rule compact_continuous_image) then obtain B where B: "\u. u \ {s .. t} \ norm (vareq x0 u) \ B" by (force dest!: compact_imp_bounded simp: bounded_iff) show "\M L. \t\{s .. t}. \x\UNIV. norm (vareq x0 t o\<^sub>L x) \ M + L * norm x" unfolding o_def by (rule exI[where x=0], rule exI[where x=B]) (auto intro!: order_trans[OF norm_blinfun_compose] mult_right_mono B) qed (auto intro: assms) lemma assumes "t \ existence_ivl0 x0" shows "continuous_on (UNIV \ existence_ivl0 x0) (\(x, ta). mvar.flow x0 t x ta)" proof - from mvar.flow_continuous_on_state_space[of x0 t, unfolded mvar_existence_ivl_eq_existence_ivl[OF assms]] show "continuous_on (UNIV \ existence_ivl0 x0) (\(x, ta). mvar.flow x0 t x ta)" . qed definition "Dflow x0 = mvar.flow x0 0 id_blinfun" lemma var_eq_mvar: assumes "t0 \ existence_ivl0 x0" assumes "t \ existence_ivl0 x0" shows "var.flow x0 t0 i t = mvar.flow x0 t0 id_blinfun t i" by (rule var.flow_unique) (auto intro!: assms derivative_eq_intros mvar.flow_has_derivative simp: varexivl_eq_exivl assms has_vector_derivative_def blinfun.bilinear_simps) lemma Dflow_zero[simp]: "x \ X \ Dflow x 0 = 1\<^sub>L" unfolding Dflow_def by (subst mvar.flow_initial_time) auto subsection \Differentiability of the flow0\ text \ \U t\, i.e. the solution of the variational equation, is the space derivative at the initial value \x0\. \ lemma flow_dx_derivative: assumes "t \ existence_ivl0 x0" shows "((\x0. flow0 x0 t) has_derivative (\z. vector_Dflow x0 z t)) (at x0)" unfolding has_derivative_at2 using assms by (intro iffD1[OF LIM_equal proposition_17_6_weak[OF assms]] conjI[OF bounded_linear_vector_Dflow[OF assms]]) (simp add: diff_diff_add inverse_eq_divide) lemma flow_dx_derivative_blinfun: assumes "t \ existence_ivl0 x0" shows "((\x. flow0 x t) has_derivative Blinfun (\z. vector_Dflow x0 z t)) (at x0)" by (rule has_derivative_Blinfun[OF flow_dx_derivative[OF assms]]) definition "flowderiv x0 t = comp12 (Dflow x0 t) (blinfun_scaleR_left (f (flow0 x0 t)))" lemma flowderiv_eq: "flowderiv x0 t (\\<^sub>1, \\<^sub>2) = (Dflow x0 t) \\<^sub>1 + \\<^sub>2 *\<^sub>R f (flow0 x0 t)" by (auto simp: flowderiv_def) lemma W_continuous_on: "continuous_on (Sigma X existence_ivl0) (\(x0, t). Dflow x0 t)" \ \TODO: somewhere here is hidden continuity wrt rhs of ODE, extract it!\ unfolding continuous_on split_beta' proof (safe intro!: tendstoI) fix e'::real and t x assume x: "x \ X" and tx: "t \ existence_ivl0 x" and e': "e' > 0" let ?S = "Sigma X existence_ivl0" have "(x, t) \ ?S" using x tx by auto from open_prod_elim[OF open_state_space this] obtain OX OT where OXOT: "open OX" "open OT" "(x, t) \ OX \ OT" "OX \ OT \ ?S" by blast then obtain dx dt where dx: "dx > 0" "cball x dx \ OX" and dt: "dt > 0" "cball t dt \ OT" by (force simp: open_contains_cball) from OXOT dt dx have "cball t dt \ existence_ivl0 x" "cball x dx \ X" apply (auto simp: subset_iff) subgoal for ta apply (drule spec[where x=ta]) apply (drule spec[where x=t])+ apply auto done done have one_exivl: "mvar.existence_ivl x 0 = (\_. existence_ivl0 x)" by (rule mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \x \ X\]]) have *: "closed ({t .. 0} \ {0 .. t})" "{t .. 0} \ {0 .. t} \ {}" by auto let ?T = "{t .. 0} \ {0 .. t} \ cball t dt" have "compact ?T" by (auto intro!: compact_Un) have "?T \ existence_ivl0 x" by (intro Un_least ivl_subset_existence_ivl' ivl_subset_existence_ivl \x \ X\ \t \ existence_ivl0 x\ \cball t dt \ existence_ivl0 x\) have "compact (mvar.flow x 0 id_blinfun ` ?T)" using \?T \ _\ \x \ X\ mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \x \ X\]] by (auto intro!: \0 < dx\ compact_continuous_image \compact ?T\ continuous_on_subset[OF mvar.flow_continuous_on]) let ?line = "mvar.flow x 0 id_blinfun ` ?T" let ?X = "{x. infdist x ?line \ dx}" have "compact ?X" using \?T \ _\ \x \ X\ mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \x \ X\]] by (auto intro!: compact_infdist_le \0 < dx\ compact_continuous_image compact_Un continuous_on_subset[OF mvar.flow_continuous_on ]) from mvar.local_lipschitz \?T \ _\ have llc: "local_lipschitz ?T ?X (\t. (o\<^sub>L) (vareq x t))" by (rule local_lipschitz_subset) auto have cont: "\xa. xa \ ?X \ continuous_on ?T (\t. vareq x t o\<^sub>L xa)" using \?T \ _\ by (auto intro!: continuous_intros \x \ X\) from local_lipschitz_compact_implies_lipschitz[OF llc \compact ?X\ \compact ?T\ cont] obtain K' where K': "\ta. ta \ ?T \ K'-lipschitz_on ?X ((o\<^sub>L) (vareq x ta))" by blast define K where "K \ abs K' + 1" have "K > 0" by (simp add: K_def) have K: "\ta. ta \ ?T \ K-lipschitz_on ?X ((o\<^sub>L) (vareq x ta))" by (auto intro!: lipschitz_onI mult_right_mono order_trans[OF lipschitz_onD[OF K']] simp: K_def) have ex_ivlI: "\y. y \ cball x dx \ ?T \ existence_ivl0 y" using dx dt OXOT by (intro Un_least ivl_subset_existence_ivl' ivl_subset_existence_ivl; force) have cont: "continuous_on ((?T \ ?X) \ cball x dx) (\((ta, xa), y). (vareq y ta o\<^sub>L xa))" using \cball x dx \ X\ ex_ivlI by (force intro!: continuous_intros simp: split_beta' ) have "mvar.flow x 0 id_blinfun t \ mvar.flow x 0 id_blinfun ` ({t..0} \ {0..t} \ cball t dt)" by auto then have mem: "(t, mvar.flow x 0 id_blinfun t, x) \ ?T \ ?X \ cball x dx" by (auto simp: \0 < dx\ less_imp_le) define e where "e \ min e' (dx / 2) / 2" have "e > 0" using \e' > 0\ by (auto simp: e_def \0 < dx\) define d where "d \ e * K / (exp (K * (abs t + abs dt + 1)) - 1)" have "d > 0" by (auto simp: d_def intro!: mult_pos_pos divide_pos_pos \0 < e\ \K > 0\) have cmpct: "compact ((?T \ ?X) \ cball x dx)" "compact (?T \ ?X)" using \compact ?T\ \compact ?X\ by (auto intro!: compact_cball compact_Times) have compact_line: "compact ?line" using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 x\ one_exivl by (force intro!: compact_continuous_image \compact ?T\ continuous_on_subset[OF mvar.flow_continuous_on] simp: \x \ X\) from compact_uniformly_continuous[OF cont cmpct(1), unfolded uniformly_continuous_on_def, rule_format, OF \0 < d\] obtain d' where d': "d' > 0" "\ta xa xa' y. ta \ ?T \ xa \ ?X \ xa'\cball x dx \ y\cball x dx \ dist xa' y < d' \ dist (vareq xa' ta o\<^sub>L xa) (vareq y ta o\<^sub>L xa) < d" by (auto simp: dist_prod_def) { fix y assume dxy: "dist x y < d'" assume "y \ cball x dx" then have "y \ X" using dx dt OXOT by force+ have two_exivl: "mvar.existence_ivl y 0 = (\_. existence_ivl0 y)" by (rule mvar_existence_ivl_eq_existence_ivl[OF existence_ivl_zero[OF \y \ X\]]) let ?X' = "\x \ ?line. ball x dx" have "open ?X'" by auto have "?X' \ ?X" by (auto intro!: infdist_le2 simp: dist_commute) interpret oneR: ll_on_open "existence_ivl0 x" "(\t. (o\<^sub>L) (vareq x t))" ?X' by standard (auto intro!: \x \ X\ continuous_intros local_lipschitz_A[OF order_refl]) interpret twoR: ll_on_open "existence_ivl0 y" "(\t. (o\<^sub>L) (vareq y t))" ?X' by standard (auto intro!: \y \ X\ continuous_intros local_lipschitz_A[OF order_refl]) interpret both: two_ll_on_open "(\t. (o\<^sub>L) (vareq x t))" "existence_ivl0 x" "(\t. (o\<^sub>L) (vareq y t))" "existence_ivl0 y" ?X' ?T "id_blinfun" d K proof unfold_locales show "0 < K" by (simp add: \0 < K\) show iv_defined: "0 \ {t..0} \ {0..t} \ cball t dt" by auto show "is_interval ({t..0} \ {0..t} \ cball t dt)" by (auto simp: is_interval_def dist_real_def) show "{t..0} \ {0..t} \ cball t dt \ oneR.existence_ivl 0 id_blinfun" apply (rule oneR.maximal_existence_flow[where x="mvar.flow x 0 id_blinfun"]) subgoal apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF mvar.flow_solves_ode[of 0 x id_blinfun]]]) subgoal using \x \ X\ \?T \ _\ \0 < dx\ by simp subgoal by simp subgoal by (simp add: \cball t dt \ existence_ivl0 x\ ivl_subset_existence_ivl ivl_subset_existence_ivl' one_exivl tx) subgoal using dx by (auto; force) done subgoal by (simp add: \x \ X\) subgoal by fact subgoal using iv_defined by blast subgoal using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 x\ by blast done fix s assume s: "s \ ?T" then show "K-lipschitz_on ?X' ((o\<^sub>L) (vareq x s))" by (intro lipschitz_on_subset[OF K \?X' \ ?X\]) auto fix j assume j: "j \ ?X'" show "norm ((vareq x s o\<^sub>L j) - (vareq y s o\<^sub>L j)) < d" unfolding dist_norm[symmetric] apply (rule d') subgoal by (rule s) subgoal using \?X' \ ?X\ j .. subgoal using \dx > 0\ by simp subgoal using \y \ cball x dx\ by simp subgoal using dxy by simp done qed have less_e: "norm (Dflow x s - both.Y s) < e" if s: "s \ ?T \ twoR.existence_ivl 0 id_blinfun" for s proof - from s have s_less: "\s\ < \t\ + \dt\ + 1" by (auto simp: dist_real_def) note both.norm_X_Y_bound[rule_format, OF s] also have "d / K * (exp (K * \s\) - 1) = e * ((exp (K * \s\) - 1) / (exp (K * (\t\ + \dt\ + 1)) - 1))" by (simp add: d_def) also have "\ < e * 1" by (rule mult_strict_left_mono[OF _ \0 < e\]) (simp add: add_nonneg_pos \0 < K\ \0 < e\ s_less) also have "\ = e" by simp also from s have s: "s \ ?T" by simp have "both.flow0 s = Dflow x s" unfolding both.flow0_def Dflow_def apply (rule oneR.maximal_existence_flow[where K="?T"]) subgoal apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF mvar.flow_solves_ode[of 0 x id_blinfun]]]) subgoal using \x \ X\ \0 < dx\ by simp subgoal by simp subgoal by (simp add: \cball t dt \ existence_ivl0 x\ ivl_subset_existence_ivl ivl_subset_existence_ivl' one_exivl tx) subgoal using dx by (auto; force) done subgoal by (simp add: \x \ X\) subgoal by (rule both.J_ivl) subgoal using both.t0_in_J by blast subgoal using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 x\ by blast subgoal using s by blast done finally show ?thesis . qed have "e < dx" using \dx > 0\ by (auto simp: e_def) let ?i = "{y. infdist y (mvar.flow x 0 id_blinfun ` ?T) \ e}" have 1: "?i \ (\x\mvar.flow x 0 id_blinfun ` ?T. ball x dx)" proof - have cl: "closed ?line" "?line \ {}" using compact_line by (auto simp: compact_imp_closed) have "?i \ (\y\mvar.flow x 0 id_blinfun ` ?T. cball y e)" proof safe fix x assume H: "infdist x ?line \ e" from infdist_attains_inf[OF cl, of x] obtain y where "y \ ?line" "infdist x ?line = dist x y" by auto then show "x \ (\x\?line. cball x e)" using H by (auto simp: dist_commute) qed also have "\ \ (\x\?line. ball x dx)" using \e < dx\ by auto finally show ?thesis . qed have 2: "twoR.flow 0 id_blinfun s \ ?i" if "s \ ?T" "s \ twoR.existence_ivl 0 id_blinfun" for s proof - from that have sT: "s \ ?T \ twoR.existence_ivl 0 id_blinfun" by force from less_e[OF this] have "dist (twoR.flow 0 id_blinfun s) (mvar.flow x 0 id_blinfun s) \ e" unfolding Dflow_def both.Y_def dist_commute dist_norm by simp then show ?thesis using sT by (force intro: infdist_le2) qed have T_subset: "?T \ twoR.existence_ivl 0 id_blinfun" apply (rule twoR.subset_mem_compact_implies_subset_existence_interval[ where K="{x. infdist x ?line \ e}"]) subgoal using \0 < dt\ by force subgoal by (rule both.J_ivl) subgoal using \y \ cball x dx\ ex_ivlI by blast subgoal using both.F_iv_defined(2) by blast subgoal by (rule 2) subgoal using \dt > 0\ by (intro compact_infdist_le) (auto intro!: compact_line \0 < e\) subgoal by (rule 1) done also have "twoR.existence_ivl 0 id_blinfun \ existence_ivl0 y" by (rule twoR.existence_ivl_subset) finally have "?T \ existence_ivl0 y" . have "norm (Dflow x s - Dflow y s) < e" if s: "s \ ?T" for s proof - from s have "s \ ?T \ twoR.existence_ivl 0 id_blinfun" using T_subset by force from less_e[OF this] have "norm (Dflow x s - both.Y s) < e" . also have "mvar.flow y 0 id_blinfun s = twoR.flow 0 id_blinfun s" apply (rule mvar.maximal_existence_flow[where K="?T"]) subgoal apply (rule solves_odeI) apply (rule has_vderiv_on_subset[OF solves_odeD(1)[OF twoR.flow_solves_ode[of 0 id_blinfun]]]) subgoal using \y \ X\ by simp subgoal using both.F_iv_defined(2) by blast subgoal using T_subset by blast subgoal by simp done subgoal using \y \ X\ auto_ll_on_open.existence_ivl_zero auto_ll_on_open_axioms both.F_iv_defined(2) twoR.flow_initial_time by blast subgoal by (rule both.J_ivl) subgoal using both.t0_in_J by blast subgoal using \{t..0} \ {0..t} \ cball t dt \ existence_ivl0 y\ by blast subgoal using s by blast done then have "both.Y s = Dflow y s" unfolding both.Y_def Dflow_def by simp finally show ?thesis . qed } note cont_data = this have "\\<^sub>F (y, s) in at (x, t) within ?S. dist x y < d'" unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] using \d' > 0\ by (intro eventually_at_Pair_within_TimesI1) (auto simp: eventually_at less_imp_le dist_commute) moreover have "\\<^sub>F (y, s) in at (x, t) within ?S. y \ cball x dx" unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] using \dx > 0\ by (intro eventually_at_Pair_within_TimesI1) (auto simp: eventually_at less_imp_le dist_commute) moreover have "\\<^sub>F (y, s) in at (x, t) within ?S. s \ ?T" unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] using \dt > 0\ by (intro eventually_at_Pair_within_TimesI2) (auto simp: eventually_at less_imp_le dist_commute) moreover have "0 \ existence_ivl0 x" by (simp add: \x \ X\) have "\\<^sub>F y in at t within existence_ivl0 x. dist (mvar.flow x 0 id_blinfun y) (mvar.flow x 0 id_blinfun t) < e" using mvar.flow_continuous_on[of x 0 id_blinfun] using \0 < e\ tx by (auto simp add: continuous_on one_exivl dest!: tendstoD) then have "\\<^sub>F (y, s) in at (x, t) within ?S. dist (Dflow x s) (Dflow x t) < e" using \0 < e\ unfolding at_within_open[OF \(x, t) \ ?S\ open_state_space] UNIV_Times_UNIV[symmetric] Dflow_def by (intro eventually_at_Pair_within_TimesI2) (auto simp: at_within_open[OF tx open_existence_ivl]) ultimately have "\\<^sub>F (y, s) in at (x, t) within ?S. dist (Dflow y s) (Dflow x t) < e'" apply eventually_elim proof (safe del: UnE, goal_cases) case (1 y s) have "dist (Dflow y s) (Dflow x t) \ dist (Dflow y s) (Dflow x s) + dist (Dflow x s) (Dflow x t)" by (rule dist_triangle) also have "dist (Dflow x s) (Dflow x t) < e" by (rule 1) also have "dist (Dflow y s) (Dflow x s) < e" unfolding dist_norm norm_minus_commute using 1 by (intro cont_data) also have "e + e \ e'" by (simp add: e_def) finally show "dist (Dflow y s) (Dflow x t) < e'" by arith qed then show "\\<^sub>F ys in at (x, t) within ?S. dist (Dflow (fst ys) (snd ys)) (Dflow (fst (x, t)) (snd (x, t))) < e'" by (simp add: split_beta') qed lemma W_continuous_on_comp[continuous_intros]: assumes h: "continuous_on S h" and g: "continuous_on S g" shows "(\s. s \ S \ h s \ X) \ (\s. s \ S \ g s \ existence_ivl0 (h s)) \ continuous_on S (\s. Dflow (h s) (g s))" using continuous_on_compose[OF continuous_on_Pair[OF h g] continuous_on_subset[OF W_continuous_on]] by auto lemma f_flow_continuous_on: "continuous_on (Sigma X existence_ivl0) (\(x0, t). f (flow0 x0 t))" using flow_continuous_on_state_space by (auto intro!: continuous_on_f flow_in_domain simp: split_beta') lemma flow_has_space_derivative: assumes "t \ existence_ivl0 x0" shows "((\x0. flow0 x0 t) has_derivative Dflow x0 t) (at x0)" by (rule flow_dx_derivative_blinfun[THEN has_derivative_eq_rhs]) (simp_all add: var_eq_mvar assms blinfun.blinfun_apply_inverse Dflow_def vector_Dflow_def mem_existence_ivl_iv_defined[OF assms]) lemma flow_has_flowderiv: assumes "t \ existence_ivl0 x0" shows "((\(x0, t). flow0 x0 t) has_derivative flowderiv x0 t) (at (x0, t) within S)" proof - have Sigma: "(x0, t) \ Sigma X existence_ivl0" using assms by auto from open_state_space assms obtain e' where e': "e' > 0" "ball (x0, t) e' \ Sigma X existence_ivl0" by (force simp: open_contains_ball) define e where "e = e' / sqrt 2" have "0 < e" using e' by (auto simp: e_def) have "ball x0 e \ ball t e \ ball (x0, t) e'" by (auto simp: dist_prod_def real_sqrt_sum_squares_less e_def) also note e'(2) finally have subs: "ball x0 e \ ball t e \ Sigma X existence_ivl0" . have d1: "((\x0. flow0 x0 s) has_derivative blinfun_apply (Dflow y s)) (at y within ball x0 e)" if "y \ ball x0 e" "s \ ball t e" for y s using subs that by (subst at_within_open; force intro!: flow_has_space_derivative) have d2: "(flow0 y has_derivative blinfun_apply (blinfun_scaleR_left (f (flow0 y s)))) (at s within ball t e)" if "y \ ball x0 e" "s \ ball t e" for y s using subs that unfolding has_vector_derivative_eq_has_derivative_blinfun[symmetric] by (subst at_within_open; force intro!: flow_has_vector_derivative) have "((\(x0, t). flow0 x0 t) has_derivative flowderiv x0 t) (at (x0, t) within ball x0 e \ ball t e)" using subs unfolding UNIV_Times_UNIV[symmetric] by (intro has_derivative_partialsI[OF d1 d2, THEN has_derivative_eq_rhs]) (auto intro!: \0 < e\ continuous_intros flow_in_domain continuous_on_imp_continuous_within[where s="Sigma X existence_ivl0"] assms simp: flowderiv_def split_beta' flow0_defined assms mem_ball) then have "((\(x0, t). flow0 x0 t) has_derivative flowderiv x0 t) (at (x0, t) within Sigma X existence_ivl0)" by (auto simp: at_within_open[OF _ open_state_space] at_within_open[OF _ open_Times] assms \0 < e\ mem_existence_ivl_iv_defined[OF assms]) then show ?thesis unfolding at_within_open[OF Sigma open_state_space] by (rule has_derivative_at_withinI) qed lemma flow0_comp_has_derivative: assumes h: "h s \ existence_ivl0 (g s)" assumes [derivative_intros]: "(g has_derivative g') (at s within S)" assumes [derivative_intros]: "(h has_derivative h') (at s within S)" shows "((\x. flow0 (g x) (h x)) has_derivative (\x. blinfun_apply (flowderiv (g s) (h s)) (g' x, h' x))) (at s within S)" by (rule has_derivative_compose[where f="\x. (g x, h x)" and s=S, OF _ flow_has_flowderiv[OF h], simplified]) (auto intro!: derivative_eq_intros) lemma flowderiv_continuous_on: "continuous_on (Sigma X existence_ivl0) (\(x0, t). flowderiv x0 t)" unfolding flowderiv_def split_beta' by (subst blinfun_of_matrix_works[where f="comp12 (Dflow (fst x) (snd x)) (blinfun_scaleR_left (f (flow0 (fst x) (snd x))))" for x, symmetric]) (auto intro!: continuous_intros flow_in_domain) lemma flowderiv_continuous_on_comp[continuous_intros]: assumes "continuous_on S x" assumes "continuous_on S t" assumes "\s. s \ S \ x s \ X" "\s. s \ S \ t s \ existence_ivl0 (x s)" shows "continuous_on S (\xa. flowderiv (x xa) (t xa))" by (rule continuous_on_compose2[OF flowderiv_continuous_on, where f="\s. (x s, t s)", unfolded split_beta' fst_conv snd_conv]) (auto intro!: continuous_intros assms) lemmas [intro] = flow_in_domain lemma vareq_trans: "t0 \ existence_ivl0 x0 \ t \ existence_ivl0 (flow0 x0 t0) \ vareq (flow0 x0 t0) t = vareq x0 (t0 + t)" by (auto simp: vareq_def flow_trans) lemma diff_existence_ivl_trans: "t0 \ existence_ivl0 x0 \ t \ existence_ivl0 x0 \ t - t0 \ existence_ivl0 (flow0 x0 t0)" for t by (metis (no_types, opaque_lifting) add.left_neutral diff_add_eq local.existence_ivl_reverse local.existence_ivl_trans local.flows_reverse) lemma has_vderiv_on_blinfun_compose_right[derivative_intros]: assumes "(g has_vderiv_on g') T" assumes "\x. x \ T \ gd' x = g' x o\<^sub>L d" shows "((\x. g x o\<^sub>L d) has_vderiv_on gd') T" using assms by (auto simp: has_vderiv_on_def has_vector_derivative_def blinfun_ext blinfun.bilinear_simps intro!: derivative_eq_intros ext) lemma has_vderiv_on_blinfun_compose_left[derivative_intros]: assumes "(g has_vderiv_on g') T" assumes "\x. x \ T \ gd' x = d o\<^sub>L g' x" shows "((\x. d o\<^sub>L g x) has_vderiv_on gd') T" using assms by (auto simp: has_vderiv_on_def has_vector_derivative_def blinfun_ext blinfun.bilinear_simps intro!: derivative_eq_intros ext) lemma mvar_flow_shift: assumes "t0 \ existence_ivl0 x0" "t1 \ existence_ivl0 x0" shows "mvar.flow x0 t0 d t1 = Dflow (flow0 x0 t0) (t1 - t0) o\<^sub>L d" proof - have "mvar.flow x0 t0 d t1 = mvar.flow x0 t0 d (t0 + (t1 - t0))" by simp also have "\ = mvar.flow x0 t0 (mvar.flow x0 t0 d t0) t1" by (subst mvar.flow_trans) (auto simp add: assms) also have "\ = Dflow (flow0 x0 t0) (t1 - t0) o\<^sub>L d" apply (rule mvar.flow_unique_on) apply (auto simp add: assms mvar.flow_initial_time_if blinfun_ext Dflow_def intro!: derivative_intros derivative_eq_intros) apply (auto simp: assms has_vderiv_on_open has_vector_derivative_def intro!: derivative_eq_intros blinfun_eqI) apply (subst mvar_existence_ivl_eq_existence_ivl) by (auto simp add: vareq_trans assms diff_existence_ivl_trans) finally show ?thesis . qed lemma Dflow_trans: assumes "h \ existence_ivl0 x0" assumes "i \ existence_ivl0 (flow0 x0 h)" shows "Dflow x0 (h + i) = Dflow (flow0 x0 h) i o\<^sub>L (Dflow x0 h)" proof - have [intro, simp]: "h + i \ existence_ivl0 x0" "i + h \ existence_ivl0 x0" "x0 \ X" using assms by (auto simp add: add.commute existence_ivl_trans) show ?thesis unfolding Dflow_def apply (subst mvar.flow_trans[where s=h and t=i]) subgoal by (auto simp: assms) subgoal by (auto simp: assms) by (subst mvar_flow_shift) (auto simp: assms Dflow_def ) qed lemma Dflow_trans_apply: assumes "h \ existence_ivl0 x0" assumes "i \ existence_ivl0 (flow0 x0 h)" shows "Dflow x0 (h + i) d0 = Dflow (flow0 x0 h) i (Dflow x0 h d0)" proof - have [intro, simp]: "h + i \ existence_ivl0 x0" "i + h \ existence_ivl0 x0" "x0 \ X" using assms by (auto simp add: add.commute existence_ivl_trans) show ?thesis unfolding Dflow_def apply (subst mvar.flow_trans[where s=h and t=i]) subgoal by (auto simp: assms) subgoal by (auto simp: assms) by (subst mvar_flow_shift) (auto simp: assms Dflow_def ) qed end \ \@{thm c1_on_open_euclidean_anchor}\ end diff --git a/thys/Poincare_Bendixson/Limit_Set.thy b/thys/Poincare_Bendixson/Limit_Set.thy --- a/thys/Poincare_Bendixson/Limit_Set.thy +++ b/thys/Poincare_Bendixson/Limit_Set.thy @@ -1,473 +1,474 @@ section \Limit Sets\ theory Limit_Set imports Invariance begin context auto_ll_on_open begin text \Positive limit point, assuming \{0..} \ existence_ivl0 x\\ (* TODO: Perhaps a more intrinsic definition would be better here *) definition "\_limit_point x p \ {0..} \ existence_ivl0 x \ (\s. s \\<^bsub>\<^esub> \ \ (flow0 x \ s) \ p)" text \ Also called the \\\-limit set of x \ definition "\_limit_set x = {p. \_limit_point x p}" definition "\_limit_point x p \ {..0} \ existence_ivl0 x \ (\s. s \\<^bsub>\<^esub> -\ \ (flow0 x \ s) \ p)" text \ Also called the \\\-limit set of x \ definition "\_limit_set x = {p. \_limit_point x p}" end context auto_ll_on_open begin lemma \_limit_point_eq_rev: "\_limit_point x p = rev.\_limit_point x p" unfolding \_limit_point_def rev.\_limit_point_def apply (auto simp: rev_eq_flow[abs_def] o_def filterlim_uminus_at_bot rev_existence_ivl_eq0 subset_iff intro: exI[where x="uminus o s" for s]) using neg_0_le_iff_le by fastforce lemma \_limit_set_eq_rev: "\_limit_set x = rev.\_limit_set x" unfolding \_limit_set_def rev.\_limit_set_def \_limit_point_eq_rev .. lemma \_limit_pointE: assumes "\_limit_point x p" obtains s where "filterlim s at_top sequentially" "(flow0 x \ s) \ p" "\n. b \ s n" using assms filterlim_at_top_choose_lower \_limit_point_def by blast lemma \_limit_set_eq: assumes "{0..} \ existence_ivl0 x" shows "\_limit_set x = (INF \ \ {0..}. closure (flow0 x ` {\..}))" unfolding \_limit_set_def proof safe fix p t assume pt: "0 \ (t::real)" "\_limit_point x p" from \_limit_pointE[OF pt(2)] obtain s where "filterlim s at_top sequentially" "(flow0 x \ s) \ p" "\n. t \ s n" by blast thus "p \ closure (flow0 x ` {t..})" unfolding closure_sequential by (metis atLeast_iff comp_apply imageI) next fix p assume "p \ (\\\{0..}. closure (flow0 x ` {\..}))" then have "\t. t \0 \ p \ closure (flow0 x ` {t..})" by blast then have "\t e. t \0 \ e > 0 \ (\tt \ t. dist (flow0 x tt) p < e)" unfolding closure_approachable by fastforce from approachable_sequenceE[OF this] obtain s where "filterlim s at_top sequentially" "(flow0 x \ s) \ p" by auto thus "\_limit_point x p" unfolding \_limit_point_def using assms by auto qed lemma \_limit_set_empty: assumes "\ ({0..} \ existence_ivl0 x)" shows "\_limit_set x = {}" unfolding \_limit_set_def \_limit_point_def by (simp add: assms) lemma \_limit_set_closed: "closed (\_limit_set x)" using \_limit_set_eq by (metis \_limit_set_empty closed_INT closed_closure closed_empty) (* TODO: Walter gives a more direct proof *) lemma \_limit_set_positively_invariant: shows "positively_invariant (\_limit_set x)" unfolding positively_invariant_def trapped_forward_def proof safe fix dummy p t assume xa: "p \ \_limit_set x" "t \ existence_ivl0 p" "0 \ t" have "p \ X" using mem_existence_ivl_iv_defined(2) xa(2) by blast have exist: "{0..} \ existence_ivl0 x" using xa(1) unfolding \_limit_set_def \_limit_point_def by auto from xa(1) obtain s where s: "filterlim s at_top sequentially" "(flow0 x \ s) \ p" "\n. 0 \ s n" unfolding \_limit_set_def by (auto elim!:\_limit_pointE) define r where "r = (\n. t + s n)" have rlim: "filterlim r at_top sequentially" unfolding r_def by (auto intro: filterlim_tendsto_add_at_top[OF _ s(1)]) define dom where "dom = image (flow0 x) {0..} \ {p} " have domin: "\n. (flow0 x \ s) n \ dom" "p \ dom" unfolding dom_def o_def using exist by(auto simp add: s(3)) have xt: "\x. x \ dom \ t \ existence_ivl0 x" unfolding dom_def using xa(2) apply auto apply (rule existence_ivl_trans') using exist xa(3) apply auto[1] using exist by auto have cont: "continuous_on dom (\x. flow0 x t)" apply (rule flow_continuous_on_compose) apply auto using \p \ X\ exist local.dom_def local.flow_in_domain apply auto[1] using xt . then have f1: "((\x. flow0 x t) \ (flow0 x \ s)) \ flow0 p t" using domin s(2) unfolding continuous_on_sequentially by blast have ff:"\n. (flow0 x \ r) n = ((\x. flow0 x t) \ (flow0 x \ s)) n" unfolding o_def r_def proof - fix n have s:"s n \ existence_ivl0 x" using s(3) exist by auto then have t:"t \ existence_ivl0 (flow0 x (s n))" using domin(1) xt by auto from flow_trans[OF s t] show "flow0 x (t + s n) = flow0 (flow0 x (s n)) t" by (simp add: add.commute) qed have f2: "(flow0 x \ r) \ flow0 p t" using f1 unfolding ff . show "flow0 p t \ \_limit_set x" using exist f2 rlim unfolding \_limit_set_def \_limit_point_def using flow_in_domain r_def s(3) xa(2) xa(3) by auto qed lemma \_limit_set_invariant: shows "invariant (\_limit_set x)" unfolding invariant_iff_pos_invariant_and_compl_pos_invariant proof safe show "positively_invariant (\_limit_set x)" using \_limit_set_positively_invariant . next show "positively_invariant (X - \_limit_set x)" unfolding positively_invariant_def trapped_forward_def apply safe using local.flow_in_domain apply blast proof - fix dummy p t assume xa: "p \ X" "p \ \_limit_set x" "t \ existence_ivl0 p" "0 \ t" and f: "flow0 p t \ \_limit_set x" have exist: "{0..} \ existence_ivl0 x" using f unfolding \_limit_set_def \_limit_point_def by auto from f obtain s where s: "filterlim s at_top sequentially" "(flow0 x \ s) \ flow0 p t" "\n. t \ s n" unfolding \_limit_set_def by (auto elim!:\_limit_pointE) define r where "r = (\n. (-t) + s n)" have "(\x. -t) \ -t" by simp from filterlim_tendsto_add_at_top[OF this s(1)] have rlim: "filterlim r at_top sequentially" unfolding r_def by simp define dom where "dom = image (flow0 x) {t..} \ {flow0 p t} " have domin: "\n. (flow0 x \ s) n \ dom" "flow0 p t \ dom" unfolding dom_def o_def using exist by(auto simp add: s(3)) have xt: "\x. x \ dom \ -t \ existence_ivl0 x" unfolding dom_def using xa(2) apply auto using local.existence_ivl_reverse xa(3) apply auto[1] by (metis exist atLeast_iff diff_conv_add_uminus diff_ge_0_iff_ge linordered_ab_group_add_class.zero_le_double_add_iff_zero_le_single_add local.existence_ivl_trans' order_trans subset_iff xa(4)) have cont: "continuous_on dom (\x. flow0 x (-t))" apply (rule flow_continuous_on_compose) apply auto using local.mem_existence_ivl_iv_defined(2) xt apply blast by (simp add: xt) then have f1: "((\x. flow0 x (-t)) \ (flow0 x \ s)) \ flow0 (flow0 p t) (-t)" using domin s(2) unfolding continuous_on_sequentially by blast have ff:"\n. (flow0 x \ r) n = ((\x. flow0 x (-t)) \ (flow0 x \ s)) n" unfolding o_def r_def proof - fix n have s:"s n \ existence_ivl0 x" using s(3) exist \0\ t\ by (meson atLeast_iff order_trans subset_eq) then have t:"-t \ existence_ivl0 (flow0 x (s n))" using domin(1) xt by auto from flow_trans[OF s t] show "flow0 x (-t + s n) = flow0 (flow0 x (s n)) (-t)" by (simp add: add.commute) qed have "(flow0 x \ r) \ flow0 (flow0 p t) (-t)" using f1 unfolding ff . then have f2: "(flow0 x \ r) \ p" using flows_reverse xa(3) by auto then have "p \ \_limit_set x" unfolding \_limit_set_def \_limit_point_def using rlim exist by auto thus False using xa(2) by auto qed qed end context auto_ll_on_open begin lemma \_limit_set_eq: assumes "{..0} \ existence_ivl0 x" shows "\_limit_set x = (INF \ \ {..0}. closure (flow0 x ` {..\}))" using rev.\_limit_set_eq[of x, OF assms[folded infinite_rev_existence_ivl0_rewrites]] unfolding \_limit_set_eq_rev rev_flow_image_eq image_uminus_atLeast by (smt INT_extend_simps(10) Sup.SUP_cong image_uminus_atMost) lemma \_limit_set_closed: shows "closed (\_limit_set x)" unfolding \_limit_set_eq_rev by (rule rev.\_limit_set_closed) lemma \_limit_set_positively_invariant: shows "negatively_invariant (\_limit_set x)" unfolding negatively_invariant_eq_rev \_limit_set_eq_rev by (simp add: rev.\_limit_set_positively_invariant) lemma \_limit_set_invariant: shows "invariant (\_limit_set x)" unfolding invariant_eq_rev \_limit_set_eq_rev by (simp add: rev.\_limit_set_invariant) text \ Fundamental properties of the positive limit set \ context fixes x K assumes K: "compact K" "K \ X" assumes x: "x \ X" "trapped_forward x K" begin text \Bunch of facts for what's to come\ private lemma props: shows "{0..} \ existence_ivl0 x" "seq_compact K" apply (rule trapped_sol_right) using x K by (auto simp add: compact_imp_seq_compact) private lemma flowimg: shows "flow0 x ` (existence_ivl0 x \ {0..}) = flow0 x ` {0..}" using props(1) by auto lemma \_limit_set_in_compact_subset: shows "\_limit_set x \ K" unfolding \_limit_set_def proof safe fix p s assume "\_limit_point x p" from \_limit_pointE[OF this] obtain s where s: "filterlim s at_top sequentially" "(flow0 x \ s) \ p" "\n. 0 \ s n" by blast then have fin: "\n. (flow0 x \ s) n \ K" using s(3) x K props(1) unfolding trapped_forward_def by (simp add: subset_eq) from seq_compactE[OF props(2) fin] show "p \ K" using s(2) by (metis LIMSEQ_subseq_LIMSEQ LIMSEQ_unique) qed lemma \_limit_set_in_compact_compact: shows "compact (\_limit_set x)" proof - from \_limit_set_in_compact_subset have "bounded (\_limit_set x)" using bounded_subset compact_imp_bounded using K(1) by auto thus ?thesis using \_limit_set_closed by (simp add: compact_eq_bounded_closed) qed lemma \_limit_set_in_compact_nonempty: shows "\_limit_set x \ {}" proof - have fin: "\n. (flow0 x \ real) n \ K" using x K props(1) by (simp add: flowimg image_subset_iff trapped_forward_def) from seq_compactE[OF props(2) this] obtain r l where "l \ K" "strict_mono r" "(flow0 x \ real \ r) \ l" by blast then have "\_limit_point x l" unfolding \_limit_point_def using props(1) by (smt comp_def filterlim_sequentially_iff_filterlim_real filterlim_subseq tendsto_at_top_eq_left) thus ?thesis unfolding \_limit_set_def by auto qed lemma \_limit_set_in_compact_existence: shows "\y. y \ \_limit_set x \ existence_ivl0 y = UNIV" proof - fix y assume y: "y \ \_limit_set x" then have "y \ X" using \_limit_set_in_compact_subset K by blast from \_limit_set_invariant have "\t. t \ existence_ivl0 y \ flow0 y t \ \_limit_set x" unfolding invariant_def trapped_iff_on_existence_ivl0 using y by blast then have t: "\t. t \ existence_ivl0 y \ flow0 y t \ K" using \_limit_set_in_compact_subset by blast thus "existence_ivl0 y = UNIV" by (meson \y \ X\ existence_ivl_zero existence_ivl_initial_time_iff existence_ivl_subset mem_compact_implies_subset_existence_interval subset_antisym K) qed lemma \_limit_set_in_compact_tendsto: shows "((\t. infdist (flow0 x t) (\_limit_set x)) \ 0) at_top" proof (rule ccontr) assume "\ ((\t. infdist (flow0 x t) (\_limit_set x)) \ 0) at_top" from not_tendsto_frequentlyE[OF this] obtain S where S: "open S" "0 \ S" "\\<^sub>F t in at_top. infdist (flow0 x t) (\_limit_set x) \ S" . then obtain e where "e > 0" "ball 0 e \ S" using openE by blast then have "\x. x \0 \ x \ S \ x \ e" by force then have "\xa. infdist (flow0 x xa) (\_limit_set x) \ S \ infdist (flow0 x xa) (\_limit_set x) \ e" using infdist_nonneg by blast from frequently_mono[OF this S(3)] have "\\<^sub>F t in at_top. infdist (flow0 x t) (\_limit_set x) \ e" by blast then have "\n. \\<^sub>F t in at_top. infdist (flow0 x t) (\_limit_set x) \ e \ real n \ t" by (auto intro!: eventually_frequently_conj) from frequently_at_topE[OF this] obtain s where s: "\i. e \ infdist (flow0 x (s i)) (\_limit_set x)" "\i. real i \ s i" "strict_mono s" by force then have sf: "filterlim s at_top sequentially" using filterlim_at_top_mono filterlim_real_sequentially not_eventuallyD by blast have fin: "\n. (flow0 x \ s) n \ K" using x K props(1) s unfolding flowimg trapped_forward_def by (metis atLeast_iff comp_apply image_subset_iff of_nat_0_le_iff order_trans) from seq_compactE[OF props(2) this] obtain r l where r:"strict_mono r" and l: "l \ K" "(flow0 x \ s \ r) \ l" by blast moreover from filterlim_at_top_strict_mono[OF s(3) r(1) sf] have "filterlim (s \ r) at_top sequentially" . moreover have "\_limit_point x l" unfolding \_limit_point_def using props(1) calculation by (metis comp_assoc) ultimately have "infdist l (\_limit_set x) = 0" by (simp add: \_limit_set_def) then have c1:"((\y. infdist y (\_limit_set x)) \ (flow0 x \ s \ r)) \ 0" by (auto intro!: tendsto_compose_at[OF l(2)] tendsto_eq_intros) have c2: "\i. e \ infdist (flow0 x ((s \ r) i)) (\_limit_set x)" using s(1) by simp show False using c1 c2 \e > 0\ unfolding o_def using Lim_bounded2 by (metis (no_types, lifting) ball_eq_empty centre_in_ball empty_iff) qed lemma \_limit_set_in_compact_connected: shows "connected (\_limit_set x)" unfolding connected_closed_set[OF \_limit_set_closed] proof clarsimp fix Apre Bpre assume pre: "closed Apre" "Apre \ Bpre = \_limit_set x" "closed Bpre" "Apre \ {}" "Bpre \ {}" "Apre \ Bpre = {}" (* TODO: this move is very strange *) then obtain A B where "Apre \ A" "Bpre \ B" "open A" "open B" and disj:"A \ B = {}" by (meson t4_space) then have "\_limit_set x \ A \ B" "\_limit_set x \ A \ {}" "\_limit_set x \ B \ {}" using pre by auto then obtain p q where p: "\_limit_point x p" "p \ A" and q: "\_limit_point x q" "q \ B" using \_limit_set_def by auto from \_limit_pointE[OF p(1)] obtain ps where ps: "filterlim ps at_top sequentially" "(flow0 x \ ps) \ p" "\n. 0 \ ps n" by blast from \_limit_pointE[OF q(1)] obtain qs where qs: "filterlim qs at_top sequentially" "(flow0 x \ qs) \ q" "\n. 0 \ qs n" by blast have "\n. \\<^sub>F t in at_top. flow0 x t \ A \ flow0 x t \ B" unfolding frequently_at_top proof safe fix dummy mpre obtain m where "m \ (0::real)" "m > mpre" by (meson approximation_preproc_push_neg(1) gt_ex le_cases order_trans) from ps obtain a where a:"a > m" "(flow0 x a) \ A" using \open A\ p unfolding tendsto_def filterlim_at_top eventually_sequentially by (metis approximation_preproc_push_neg(1) comp_apply gt_ex le_cases order_trans) from qs obtain b where b:"b > a" "(flow0 x b) \ B" using \open B\ q unfolding tendsto_def filterlim_at_top eventually_sequentially by (metis approximation_preproc_push_neg(1) comp_apply gt_ex le_cases order_trans) have "continuous_on {a..b} (flow0 x)" by (metis Icc_subset_Ici_iff \0 \ m\ \m < a\ approximation_preproc_push_neg(2) atMost_iff atMost_subset_iff continuous_on_subset le_cases local.flow_continuous_on props(1) subset_eq) from connected_continuous_image[OF this connected_Icc] have c:"connected (flow0 x ` {a..b})" . have "\t\ {a..b}. flow0 x t \ A \ flow0 x t \ B" proof (rule ccontr) assume "\ (\t\{a..b}. flow0 x t \ A \ flow0 x t \ B)" then have "flow0 x ` {a..b} \ A \ B" by blast from topological_space_class.connectedD[OF c \open A\ \open B\ _ this] show False using a b disj by force qed thus "\n>mpre. flow0 x n \ A \ flow0 x n \ B" by (smt \mpre < m\ a(1) atLeastAtMost_iff) qed from frequently_at_topE'[OF this filterlim_real_sequentially] obtain s where s: "\i. flow0 x (s i) \ A \ flow0 x (s i) \ B" "strict_mono s" "\n. real n \ s n" by blast then have "\n. (flow0 x \ s) n \ K" by (smt atLeast_iff comp_apply flowimg image_subset_iff of_nat_0_le_iff trapped_forward_def x(2)) from seq_compactE[OF props(2) this] obtain r l where r: "l \ K" "strict_mono r" "(flow0 x \ s \ r) \ l" by blast have "filterlim s at_top sequentially" using s filterlim_at_top_mono filterlim_real_sequentially not_eventuallyD by blast from filterlim_at_top_strict_mono[OF s(2) r(2) this] have "filterlim (s \ r) at_top sequentially" . then have "\_limit_point x l" unfolding \_limit_point_def using props(1) r by (metis comp_assoc) moreover have "l \ A" using s(1) r(3) \open A\ unfolding tendsto_def by auto moreover have "l \ B" using s(1) r(3) \open B\ unfolding tendsto_def by auto ultimately show False using \\_limit_set x \ A \ B\ unfolding \_limit_set_def by auto qed lemma \_limit_set_in_compact_\_limit_set_contained: shows "\y \ \_limit_set x. \_limit_set y \ \_limit_set x" proof safe fix y z assume "y \ \_limit_set x" "z \ \_limit_set y" then have "\_limit_point y z" unfolding \_limit_set_def by auto from \_limit_pointE[OF this] obtain s where s: "(flow0 y \ s) \ z" . have "\n. (flow0 y \ s) n \ \_limit_set x" using \y \ \_limit_set x\ invariant_def \_limit_set_in_compact_existence \_limit_set_invariant trapped_iff_on_existence_ivl0 by force thus "z \ \_limit_set x" using closed_sequential_limits s \_limit_set_closed by blast qed lemma \_limit_set_in_compact_\_limit_set_contained: assumes zpx: "z \ \_limit_set x" shows "\_limit_set z \ \_limit_set x" proof fix w assume "w \ \_limit_set z" then obtain s where s: "(flow0 z \ s) \ w" unfolding \_limit_set_def \_limit_point_def by auto from \_limit_set_invariant have "invariant (\_limit_set x)" . then have "\n. (flow0 z \ s) n \ \_limit_set x" using \_limit_set_in_compact_existence[OF zpx] zpx using invariant_def trapped_iff_on_existence_ivl0 by fastforce - from closed_sequentially[OF \_limit_set_closed this s] - show "w \ \_limit_set x" . + from closed_sequentially[OF \_limit_set_closed] this s + show "w \ \_limit_set x" + by blast qed end text \ Fundamental properties of the negative limit set \ end context auto_ll_on_open begin context fixes x K assumes x: "x \ X" "trapped_backward x K" assumes K: "compact K" "K \ X" begin private lemma xrev: "x \ X" "rev.trapped_forward x K" using trapped_backward_iff_rev_trapped_forward x(2) by (auto simp: rev_existence_ivl_eq0 rev_eq_flow x(1)) lemma \_limit_set_in_compact_subset: "\_limit_set x \ K" and \_limit_set_in_compact_compact: "compact (\_limit_set x)" and \_limit_set_in_compact_nonempty: "\_limit_set x \ {}" and \_limit_set_in_compact_connected: "connected (\_limit_set x)" and \_limit_set_in_compact_\_limit_set_contained: "\y \ \_limit_set x. \_limit_set y \ \_limit_set x" and \_limit_set_in_compact_tendsto: "((\t. infdist (flow0 x t) (\_limit_set x)) \ 0) at_bot" using rev.\_limit_set_in_compact_subset[OF K xrev] using rev.\_limit_set_in_compact_compact[OF K xrev] using rev.\_limit_set_in_compact_nonempty[OF K xrev] using rev.\_limit_set_in_compact_connected[OF K xrev] using rev.\_limit_set_in_compact_\_limit_set_contained[OF K xrev] using rev.\_limit_set_in_compact_tendsto[OF K xrev] unfolding invariant_eq_rev \_limit_set_eq_rev existence_ivl_eq_rev flow_eq_rev0 filterlim_at_bot_mirror minus_minus . lemma \_limit_set_in_compact_existence: shows "\y. y \ \_limit_set x \ existence_ivl0 y = UNIV" using rev.\_limit_set_in_compact_existence[OF K xrev] unfolding \_limit_set_eq_rev existence_ivl_eq_rev0 by auto end end end diff --git a/thys/Poincare_Bendixson/ODE_Misc.thy b/thys/Poincare_Bendixson/ODE_Misc.thy --- a/thys/Poincare_Bendixson/ODE_Misc.thy +++ b/thys/Poincare_Bendixson/ODE_Misc.thy @@ -1,1111 +1,1112 @@ section \Additions to the ODE Library\ theory ODE_Misc imports Ordinary_Differential_Equations.ODE_Analysis Analysis_Misc begin lemma local_lipschitz_compact_bicomposeE: assumes ll: "local_lipschitz T X f" assumes cf: "\x. x \ X \ continuous_on I (\t. f t x)" assumes cI: "compact I" assumes "I \ T" assumes cv: "continuous_on I v" assumes cw: "continuous_on I w" assumes v: "v ` I \ X" assumes w: "w ` I \ X" obtains L where "L > 0" "\x. x \ I \ dist (f x (v x)) (f x (w x)) \ L * dist (v x) (w x)" proof - from v w have "v ` I \ w ` I \ X" by auto with ll \I \ T\ have llI:"local_lipschitz I (v ` I \ w ` I) f" by (rule local_lipschitz_subset) have cvwI: "compact (v ` I \ w ` I)" by (auto intro!: compact_continuous_image cv cw cI) from local_lipschitz_compact_implies_lipschitz[OF llI cvwI \compact I\ cf] obtain L where L: "\t. t \ I \ L-lipschitz_on (v ` I \ w ` I) (f t)" using v w by blast define L' where "L' = max L 1" with L have "L' > 0" "\x. x \ I \ dist (f x (v x)) (f x (w x)) \ L' * dist (v x) (w x)" apply (auto simp: lipschitz_on_def L'_def) by (smt Un_iff image_eqI mult_right_mono zero_le_dist) then show ?thesis .. qed subsection \Comparison Principle\ lemma comparison_principle_le: fixes f::"real \ real \ real" and \ \::"real \ real" assumes ll: "local_lipschitz X Y f" assumes cf: "\x. x \ Y \ continuous_on {a..b} (\t. f t x)" assumes abX: "{a .. b} \ X" assumes \': "\x. x \ {a .. b} \ (\ has_real_derivative \' x) (at x)" assumes \': "\x. x \ {a .. b} \ (\ has_real_derivative \' x) (at x)" assumes \_in: "\ ` {a..b} \ Y" assumes \_in: "\ ` {a..b} \ Y" assumes init: "\ a \ \ a" assumes defect: "\x. x \ {a .. b} \ \' x - f x (\ x) \ \' x - f x (\ x)" shows "\x \ {a .. b}. \ x \ \ x" (is "?th1") (* "(\x \ {a .. b}. \ x < \ x) \ (\c\{a..b}. (\x\{a..c}. \ x \ \ x) \ (\x\{c<..b}. \ x < \ x))" (is "?th2") *) unfolding atomize_conj apply (cases "a \ b") defer subgoal by simp proof - assume "a \ b" note \_cont = has_real_derivative_imp_continuous_on[OF \'] note \_cont = has_real_derivative_imp_continuous_on[OF \'] from local_lipschitz_compact_bicomposeE[OF ll cf compact_Icc abX \_cont \_cont \_in \_in] obtain L where L: "L > 0" "\x. x \ {a..b} \ dist (f x (\ x)) (f x (\ x)) \ L * dist (\ x) (\ x)" by blast define w where "w x = \ x - \ x" for x have w'[derivative_intros]: "\x. x \ {a .. b} \ (w has_real_derivative \' x - \' x) (at x)" using \' \' by (auto simp: has_vderiv_on_def w_def[abs_def] intro!: derivative_eq_intros) note w_cont[continuous_intros] = has_real_derivative_imp_continuous_on[OF w', THEN continuous_on_compose2] have "w d \ 0" if "d \ {a .. b}" for d proof (rule ccontr, unfold not_le) assume "w d < 0" let ?N = "(w -` {..0} \ {a .. d})" from \w d < 0\ that have "d \ ?N" by auto then have "?N \ {}" by auto have "closed ?N" unfolding compact_eq_bounded_closed using that by (intro conjI closed_vimage_Int) (auto intro!: continuous_intros) let ?N' = "{a0 \ {a .. d}. w ` {a0 .. d} \ {..0}}" from \w d < 0\ that have "d \ ?N'" by simp then have "?N' \ {}" by auto have "compact ?N'" unfolding compact_eq_bounded_closed proof have "?N' \ {a .. d}" using that by auto then show "bounded ?N'" by (rule bounded_subset[rotated]) simp have "w u \ 0" if "(\n. x n \ ?N')" "x \ l" "l \ u" "u \ d" for x l u proof cases assume "l = u" have "\n. x n \ ?N" using that(1) by force - from closed_sequentially[OF \closed ?N\ this \x \ l\] - show ?thesis by (auto simp: \l = u\) + from closed_sequentially[OF \closed ?N\] this \x \ l\ + show ?thesis + using \l = u\ by blast next assume "l \ u" with that have "l < u" by auto from order_tendstoD(2)[OF \x \ l\ \l < u\] obtain n where "x n < u" by (auto dest: eventually_happens) with that show ?thesis using \l < u\ by (auto dest!: spec[where x=n] simp: image_subset_iff) qed then show "closed ?N'" unfolding closed_sequential_limits by (auto simp: Lim_bounded Lim_bounded2) qed from compact_attains_inf[OF \compact ?N'\ \?N' \ {}\] obtain a0 where a0: "a \ a0" "a0 \ d" "w ` {a0..d} \ {..0}" and a0_least: "\x. a \ x \ x \ d \ w ` {x..d} \ {..0} \ a0 \ x" by auto have a0d: "{a0 .. d} \ {a .. b}" using that a0 by auto have L_w_bound: "L * w x \ \' x - \' x" if "x \ {a0 .. d}" for x proof - from set_mp[OF a0d that] have "x \ {a .. b}" . from defect[OF this] have "\' x - \' x \ dist (f x (\ x)) (f x (\ x))" by (simp add: dist_real_def) also have "\ \ L * dist (\ x) (\ x)" using \x \ {a .. b}\ by (rule L) also have "\ \ -L * w x" using \0 < L\ a0 that by (force simp add: dist_real_def abs_real_def w_def algebra_split_simps ) finally show ?thesis by simp qed have mono: "mono_on {a0..d} (\x. w x * exp(-L*x))" apply (rule mono_onI) apply (rule DERIV_nonneg_imp_nondecreasing, assumption) using a0d by (auto intro!: exI[where x="(\' x - \' x) * exp (- (L * x)) - exp (- (L * x)) * L * w x" for x] derivative_eq_intros L_w_bound simp: ) then have "w a0 * exp (-L * a0) \ w d * exp (-L * d)" by (rule mono_onD) (use that a0 in auto) also have "\ < 0" using \w d < 0\ by (simp add: algebra_split_simps) finally have "w a0 * exp (- L * a0) < 0" . then have "w a0 < 0" by (simp add: algebra_split_simps) have "a0 \ a" proof (rule ccontr, unfold not_le) assume "a < a0" have "continuous_on {a.. a0} w" by (rule continuous_intros, assumption) (use a0 a0d in auto) from continuous_on_Icc_at_leftD[OF this \a < a0\] have "(w \ w a0) (at_left a0)" . from order_tendstoD(2)[OF this \w a0 < 0\] have "\\<^sub>F x in at_left a0. w x < 0" . moreover have "\\<^sub>F x in at_left a0. a < x" by (rule order_tendstoD) (auto intro!: \a < a0\) ultimately have "\\<^sub>F x in at_left a0. a < x \ w x < 0" by eventually_elim auto then obtain a1' where "a1'y. y > a1' \ y < a0 \ a < y \ w y < 0" unfolding eventually_at_left_field by auto define a1 where "a1 = (a1' + a0)/2" have "a1 < a0" using \a1' < a0\ by (auto simp: a1_def) have "a \ a1" using \a < a0\ a1_neg by (force simp: a1_def) moreover have "a1 \ d" using \a1' < a0\ a0(2) by (auto simp: a1_def) moreover have "w ` {a1..a0} \ {..0}" using \w a0 < 0\ a1_neg a0(3) by (auto simp: a1_def) smt moreover have "w ` {a0..d} \ {..0}" using a0 by auto ultimately have "a0 \ a1" apply (intro a0_least) apply assumption apply assumption by (smt atLeastAtMost_iff image_subset_iff) with \a1 show False by simp qed then have "a0 = a" using \a \ a0\ by simp with \w a0 < 0\ have "w a < 0" by simp with init show False by (auto simp: w_def) qed then show ?thesis by (auto simp: w_def) qed lemma local_lipschitz_mult: shows "local_lipschitz (UNIV::real set) (UNIV::real set) (*)" apply (auto intro!: c1_implies_local_lipschitz[where f'="\p. blinfun_mult_left (fst p)"]) apply (simp add: has_derivative_mult_right mult_commute_abs) by (auto intro!: continuous_intros) lemma comparison_principle_le_linear: fixes \ :: "real \ real" assumes "continuous_on {a..b} g" assumes "(\t. t \ {a..b} \ (\ has_real_derivative \' t) (at t))" assumes "\ a \ 0" assumes "(\t. t \ {a..b} \ \' t \ g t *\<^sub>R \ t)" shows "\t\{a..b}. \ t \ 0" proof - have *: "\x. continuous_on {a..b} (\t. g t * x)" using assms(1) continuous_on_mult_right by blast then have "local_lipschitz (g`{a..b}) UNIV (*)" using local_lipschitz_subset[OF local_lipschitz_mult] by blast from local_lipschitz_compose1[OF this assms(1)] have "local_lipschitz {a..b} UNIV (\t. (*) (g t))" . from comparison_principle_le[OF this _ _ assms(2) _ _ _ assms(3), of b "\t.0"] * assms(4) show ?thesis by auto qed subsection \Locally Lipschitz ODEs\ context ll_on_open_it begin lemma flow_lipschitzE: assumes "{a .. b} \ existence_ivl t0 x" obtains L where "L-lipschitz_on {a .. b} (flow t0 x)" proof - have f': "(flow t0 x has_derivative (\i. i *\<^sub>R f t (flow t0 x t))) (at t within {a .. b})" if "t \ {a .. b}" for t using flow_has_derivative[of t x] assms that by (auto simp: has_derivative_at_withinI) have "compact ((\t. f t (flow t0 x t)) ` {a .. b})" using assms apply (auto intro!: compact_continuous_image continuous_intros) using local.existence_ivl_empty2 apply fastforce apply (meson atLeastAtMost_iff general.existence_ivl_subset in_mono) by (simp add: general.flow_in_domain subset_iff) then obtain C where "t \ {a .. b} \ norm (f t (flow t0 x t)) \ C" for t by (fastforce dest!: compact_imp_bounded simp: bounded_iff intro: that) then have "t \ {a..b} \ onorm (\i. i *\<^sub>R f t (flow t0 x t)) \ max 0 C" for t apply (subst onorm_scaleR_left) apply (auto simp: onorm_id max_def) by (metis diff_0_right diff_mono diff_self norm_ge_zero) from bounded_derivative_imp_lipschitz[OF f' _ this] have "(max 0 C)-lipschitz_on {a..b} (flow t0 x)" by auto then show ?thesis .. qed lemma flow_undefined0: "t \ existence_ivl t0 x \ flow t0 x t = 0" unfolding flow_def by auto lemma csols_undefined: "x \ X \ csols t0 x = {}" apply (auto simp: csols_def) using general.existence_ivl_empty2 general.existence_ivl_maximal_segment apply blast done lemmas existence_ivl_undefined = existence_ivl_empty2 end subsection \Reverse flow as Sublocale\ lemma range_preflect_0[simp]: "range (preflect 0) = UNIV" by (auto simp: preflect_def) lemma range_uminus[simp]: "range uminus = (UNIV::'a::ab_group_add set)" by auto context auto_ll_on_open begin sublocale rev: auto_ll_on_open "-f" rewrites "-(-f) = f" apply unfold_locales using auto_local_lipschitz auto_open_domain unfolding fun_Compl_def local_lipschitz_minus by auto lemma existence_ivl_eq_rev0: "existence_ivl0 y = uminus ` rev.existence_ivl0 y" for y by (auto simp: existence_ivl_eq_rev rev.existence_ivl0_def preflect_def) lemma rev_existence_ivl_eq0: "rev.existence_ivl0 y = uminus ` existence_ivl0 y" for y using uminus_uminus_image[of "rev.existence_ivl0 y"] by (simp add: existence_ivl_eq_rev0) lemma flow_eq_rev0: "flow0 y t = rev.flow0 y (-t)" for y t apply (cases "t \ existence_ivl0 y") subgoal apply (subst flow_eq_rev(2), assumption) apply (subst rev.flow0_def) by (simp add: preflect_def) subgoal apply (frule flow_undefined0) by (auto simp: existence_ivl_eq_rev0 rev.flow_undefined0) done lemma rev_eq_flow: "rev.flow0 y t = flow0 y (-t)" for y t apply (subst flow_eq_rev0) using uminus_uminus_image[of "rev.existence_ivl0 y"] apply - apply (subst (asm) existence_ivl_eq_rev0[symmetric]) by auto lemma rev_flow_image_eq: "rev.flow0 x ` S = flow0 x ` (uminus ` S)" unfolding rev_eq_flow[abs_def] by force lemma flow_image_eq_rev: "flow0 x ` S = rev.flow0 x ` (uminus ` S)" unfolding rev_eq_flow[abs_def] by force end context c1_on_open begin sublocale rev: c1_on_open "-f" "-f'" rewrites "-(-f) = f" and "-(-f') = f'" by (rule c1_on_open_rev) auto end context c1_on_open_euclidean begin sublocale rev: c1_on_open_euclidean "-f" "-f'" rewrites "-(-f) = f" and "-(-f') = f'" by unfold_locales auto end subsection \Autonomous LL ODE : Existence Interval and trapping on the interval\ lemma bdd_above_is_intervalI: "bdd_above I" if "is_interval I" "a \ b" "a \ I" "b \ I" for I::"real set" by (meson bdd_above_def is_interval_1 le_cases that) lemma bdd_below_is_intervalI: "bdd_below I" if "is_interval I" "a \ b" "a \ I" "b \ I" for I::"real set" by (meson bdd_below_def is_interval_1 le_cases that) context auto_ll_on_open begin lemma open_existence_ivl0: assumes x : "x \ X" shows "\a b. a < 0 \ 0 < b \ {a..b} \ existence_ivl0 x" proof - have a1:"0 \ existence_ivl0 x" by (simp add: x) have a2: "open (existence_ivl0 x)" by (simp add: x) from a1 a2 obtain d where "d > 0" "ball 0 d \ existence_ivl0 x" using openE by blast have "{-d/2..d/2} \ ball 0 d" using \0 < d\ dist_norm mem_ball by auto thus ?thesis by (smt \0 < d\ \ball 0 d \ existence_ivl0 x\ divide_minus_left half_gt_zero order_trans) qed lemma open_existence_ivl': assumes x : "x \ X" obtains a where "a > 0" "{-a..a} \ existence_ivl0 x" proof - from open_existence_ivl0[OF assms(1)] obtain a b where ab: "a < 0" "0 < b" "{a..b} \ existence_ivl0 x" by auto then have "min (-a) b > 0" by linarith have "{-min (-a) b .. min(-a) b} \ {a..b}" by auto thus ?thesis using ab(3) that[OF \min (-a) b > 0\] by blast qed lemma open_existence_ivl_on_compact: assumes C: "C \ X" and "compact C" "C \ {}" obtains a where "a > 0" "\x. x \ C \ {-a..a} \ existence_ivl0 x" proof - from existence_ivl_cballs have "\x\C. \e>0. \t>0. \y\cball x e. cball 0 t\existence_ivl0 y" by (metis (full_types) C Int_absorb1 Int_iff UNIV_I) then obtain d' t' where *: "\x\C. 0 < d' x \ t' x > 0 \ (\y\cball x (d' x). cball 0 (t' x) \ existence_ivl0 y)" by metis with compactE_image[OF \compact C\, of C "\x. ball x (d' x)"] obtain C' where "C' \ C" and [simp]: "finite C'" and C_subset: "C \ (\c\C'. ball c (d' c))" by force from C_subset \C \ {}\ have [simp]: "C' \ {}" by auto define d where "d = Min (d' ` C')" define t where "t = Min (t' ` C')" have "t > 0" using * \C' \ C\ by (auto simp: t_def) moreover have "{-t .. t} \ existence_ivl0 x" if "x \ C" for x proof - from C_subset that \C' \ C\ obtain c where c: "c \ C'" "x \ ball c (d' c)" "c \ C" by force then have "{-t .. t} \ cball 0 (t' c)" by (auto simp: abs_real_def t_def minus_le_iff) also from c have "cball 0 (t' c) \ existence_ivl0 x" using *[rule_format, OF \c \ C\] by auto finally show ?thesis . qed ultimately show ?thesis .. qed definition "trapped_forward x K \ (flow0 x ` (existence_ivl0 x \ {0..}) \ K)" \ \TODO: use this for backwards trapped, invariant, and all assumptions\ definition "trapped_backward x K \ (flow0 x ` (existence_ivl0 x \ {..0}) \ K)" definition "trapped x K \ trapped_forward x K \ trapped_backward x K" lemma trapped_iff_on_existence_ivl0: "trapped x K \ (flow0 x ` (existence_ivl0 x) \ K)" unfolding trapped_def trapped_forward_def trapped_backward_def apply (auto) by (metis IntI atLeast_iff atMost_iff image_subset_iff less_eq_real_def linorder_not_less) end context auto_ll_on_open begin lemma infinite_rev_existence_ivl0_rewrites: "{0..} \ rev.existence_ivl0 x \ {..0} \ existence_ivl0 x" "{..0} \ rev.existence_ivl0 x \ {0..} \ existence_ivl0 x" apply (auto simp add: rev.rev_existence_ivl_eq0 subset_iff) using neg_le_0_iff_le apply fastforce using neg_0_le_iff_le by fastforce lemma trapped_backward_iff_rev_trapped_forward: "trapped_backward x K \ rev.trapped_forward x K" unfolding trapped_backward_def rev.trapped_forward_def by (auto simp add: rev_flow_image_eq existence_ivl_eq_rev0 image_subset_iff) text \If solution is trapped in a compact set at some time on its existence interval then it is trapped forever\ lemma trapped_sol_right: \ \TODO: when building on afp-devel (??? outdated): \<^url>\https://bitbucket.org/isa-afp/afp-devel/commits/0c3edf9248d5389197f248c723b625c419e4d3eb\\ assumes "compact K" "K \ X" assumes "x \ X" "trapped_forward x K" shows "{0..} \ existence_ivl0 x" proof (rule ccontr) assume "\ {0..} \ existence_ivl0 x" from this obtain t where "0 \ t" "t \ existence_ivl0 x" by blast then have bdd: "bdd_above (existence_ivl0 x)" by (auto intro!: bdd_above_is_intervalI \x \ X\) from flow_leaves_compact_ivl_right [OF UNIV_I \x \ X\ bdd UNIV_I assms(1-2)] show False by (metis assms(4) trapped_forward_def IntI atLeast_iff image_subset_iff) qed lemma trapped_sol_right_gen: assumes "compact K" "K \ X" assumes "t \ existence_ivl0 x" "trapped_forward (flow0 x t) K" shows "{t..} \ existence_ivl0 x" proof - have "x \ X" using assms(3) local.existence_ivl_empty_iff by fastforce have xtk: "flow0 x t \ X" by (simp add: assms(3) local.flow_in_domain) from trapped_sol_right[OF assms(1-2) xtk assms(4)] have "{0..} \ existence_ivl0 (flow0 x t)" . thus "{t..} \ existence_ivl0 x" using existence_ivl_trans[OF assms(3)] by (metis add.commute atLeast_iff diff_add_cancel le_add_same_cancel1 subset_iff) qed lemma trapped_sol_left: \ \TODO: when building on afp-devel: \<^url>\https://bitbucket.org/isa-afp/afp-devel/commits/0c3edf9248d5389197f248c723b625c419e4d3eb\\ assumes "compact K" "K \ X" assumes "x \ X" "trapped_backward x K" shows "{..0} \ existence_ivl0 x" proof (rule ccontr) assume "\ {..0} \ existence_ivl0 x" from this obtain t where "t \ 0" "t \ existence_ivl0 x" by blast then have bdd: "bdd_below (existence_ivl0 x)" by (auto intro!: bdd_below_is_intervalI \x \ X\) from flow_leaves_compact_ivl_left [OF UNIV_I \x \ X\ bdd UNIV_I assms(1-2)] show False by (metis IntI assms(4) atMost_iff auto_ll_on_open.trapped_backward_def auto_ll_on_open_axioms image_subset_iff) qed lemma trapped_sol_left_gen: assumes "compact K" "K \ X" assumes "t \ existence_ivl0 x" "trapped_backward (flow0 x t) K" shows "{..t} \ existence_ivl0 x" proof - have "x \ X" using assms(3) local.existence_ivl_empty_iff by fastforce have xtk: "flow0 x t \ X" by (simp add: assms(3) local.flow_in_domain) from trapped_sol_left[OF assms(1-2) xtk assms(4)] have "{..0} \ existence_ivl0 (flow0 x t)" . thus "{..t} \ existence_ivl0 x" using existence_ivl_trans[OF assms(3)] by (metis add.commute add_le_same_cancel1 atMost_iff diff_add_cancel subset_eq) qed lemma trapped_sol: assumes "compact K" "K \ X" assumes "x \ X" "trapped x K" shows "existence_ivl0 x = UNIV" by (metis (mono_tags, lifting) assms existence_ivl_zero image_subset_iff interval local.existence_ivl_initial_time_iff local.existence_ivl_subset local.subset_mem_compact_implies_subset_existence_interval order_refl subset_antisym trapped_iff_on_existence_ivl0) (* Follows from rectification *) lemma regular_locally_noteq:\ \TODO: should be true in \ll_on_open_it\\ assumes "x \ X" "f x \ 0" shows "eventually (\t. flow0 x t \ x) (at 0)" proof - have nf:"norm (f x) > 0" by (simp add: assms(2)) (* By continuity of solutions and f probably *) obtain a where a: "a>0" "{-a--a} \ existence_ivl0 x" "0 \ {-a--a}" "\t. t \ {-a--a} \ norm(f (flow0 x t) - f (flow0 x 0)) \ norm(f x)/2" proof - from open_existence_ivl'[OF assms(1)] obtain a1 where a1: "a1 > 0" "{-a1..a1} \ existence_ivl0 x" . have "continuous (at 0) (\t. norm(f (flow0 x t) - f (flow0 x 0) ))" apply (auto intro!: continuous_intros) by (simp add: assms(1) local.f_flow_continuous) then obtain a2 where "a2>0" "\t. norm t < a2 \ norm (f (flow0 x t) - f (flow0 x 0)) < norm(f x)/2" unfolding continuous_at_real_range by (metis abs_norm_cancel cancel_comm_monoid_add_class.diff_cancel diff_zero half_gt_zero nf norm_zero) then have t: "\t. t \ {-a2<-- norm(f (flow0 x t) - f (flow0 x 0)) \ norm(f x)/2" by (smt open_segment_bound(2) open_segment_bound1 real_norm_def) define a where "a = min a1 (a2/2)" have t1:"a > 0" unfolding a_def using \a1 > 0\ \a2 > 0\ by auto then have t3:"0 \{-a--a}" using closed_segment_eq_real_ivl by auto have "{-a--a} \ {-a1..a1}" unfolding a_def using \a1 > 0\ \a2 > 0\ using ODE_Auxiliarities.closed_segment_eq_real_ivl by auto then have t2:"{-a--a} \ existence_ivl0 x" using a1 by auto have "{-a--a} \ {-a2<--a1 > 0\ \a2 > 0\ by (smt Diff_iff closed_segment_eq_real_ivl atLeastAtMost_iff empty_iff half_gt_zero insert_iff pos_half_less segment(1) subset_eq) then have t4:"\t. t \ {-a--a} \ norm(f (flow0 x t) - f (flow0 x 0)) \ norm(f x)/2" using t by auto show ?thesis using t1 t2 t3 t4 that by auto qed have "\t. t \ {-a--a} \ (flow0 x has_vector_derivative f (flow0 x t)) (at t within {-a--a})" apply (rule has_vector_derivative_at_within) using a(2) by (auto intro!:flow_has_vector_derivative) from vector_differentiable_bound_linearization[OF this _ a(4)] have nb:"\c d. {c--d} \ {-a--a} \ norm (flow0 x d - flow0 x c - (d - c) *\<^sub>R f (flow0 x 0)) \ norm (d - c) * (norm (f x) / 2)" using a(3) by blast have "\t. dist t 0 < a \ t \ 0 \ flow0 x t \ x" proof (rule ccontr) fix t assume "dist t 0 < a" "t \ 0" "\ flow0 x t \ x" then have tx:"flow0 x t = x" by auto have "t \ {-a--a}" using closed_segment_eq_real_ivl \dist t 0 < a\ by auto have "t > 0 \ t < 0" using \t \ 0\ by linarith moreover { assume "t > 0" then have "{0--t} \ {-a--a}" by (simp add: \t \ {-a--a}\ a(3) subset_closed_segment) from nb[OF this] have "norm (flow0 x t - x - t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" by (simp add: assms(1)) then have "norm (t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" using tx by auto then have False using nf using \0 < t\ by auto } moreover { assume "t < 0" then have "{t--0} \ {-a--a}" by (simp add: \t \ {-a--a}\ a(3) subset_closed_segment) from nb[OF this] have "norm (x - flow0 x t + t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" by (simp add: assms(1)) then have "norm (t *\<^sub>R f x) \ norm t * (norm (f x) / 2)" using tx by auto then have False using nf using \t < 0\ by auto } ultimately show False by blast qed thus ?thesis unfolding eventually_at using a(1) by blast qed lemma compact_max_time_flow_in_closed: assumes "closed M" and t_ex: "t \ existence_ivl0 x" shows "compact {s \ {0..t}. flow0 x ` {0..s} \ M}" (is "compact ?C") unfolding compact_eq_bounded_closed proof have "bounded {0 .. t}" by auto then show "bounded ?C" by (rule bounded_subset) auto show "closed ?C" unfolding closed_def proof (rule topological_space_class.openI, clarsimp) \ \TODO: there must be a more abstract argument for this, e.g., with @{thm continuous_on_closed_vimageI} and then reasoning about the connected component around 0?\ fix s assume notM: "s \ t \ 0 \ s \ \ flow0 x ` {0..s} \ M" consider "0 \ s" "s \ t" "flow0 x s \ M" | "0 \ s" "s \ t" "flow0 x s \ M" | "s < 0" | "s > t" by arith then show "\T. open T \ s \ T \ T \ - {s. 0 \ s \ s \ t \ flow0 x ` {0..s} \ M}" proof cases assume s: "0 \ s" "s \ t" and sM: "flow0 x s \ M" have "isCont (flow0 x) s" using s ivl_subset_existence_ivl[OF t_ex] by (auto intro!: flow_continuous) from this[unfolded continuous_at_open, rule_format, of "-M"] sM \closed M\ obtain S where "open S" "s \ S" "(\x'\S. flow0 x x' \ - M)" by auto then show ?thesis by (force intro!: exI[where x=S]) next assume s: "0 \ s" "s \ t" and sM: "flow0 x s \ M" from this notM obtain s0 where s0: "0 \ s0" "s0 < s" "flow0 x s0 \ M" by force from order_tendstoD(1)[OF tendsto_ident_at \s0 < s\, of UNIV, unfolded eventually_at_topological] obtain S where "open S" "s \ S" "\x. x \ S \ x \ s \ s0 < x" by auto then show ?thesis using s0 by (auto simp: intro!: exI[where x=S]) (smt atLeastAtMost_iff image_subset_iff) qed (force intro: exI[where x="{t<..}"] exI[where x="{..<0}"])+ qed qed lemma flow_in_closed_max_timeE: assumes "closed M" "t \ existence_ivl0 x" "0 \ t" "x \ M" obtains T where "0 \ T" "T \ t" "flow0 x ` {0..T} \ M" "\s'. 0 \ s' \ s' \ t \ flow0 x ` {0..s'} \ M \ s' \ T" proof - let ?C = "{s \ {0..t}. flow0 x ` {0..s} \ M}" have "?C \ {}" using assms using local.mem_existence_ivl_iv_defined by (auto intro!: exI[where x=0]) from compact_max_time_flow_in_closed[OF assms(1,2)] have "compact ?C" . from compact_attains_sup[OF this \?C \ {}\] obtain s where s: "0 \ s" "s \ t" "flow0 x ` {0..s} \ M" and s_max: "\s'. 0 \ s' \ s' \ t \ flow0 x ` {0..s'} \ M \ s' \ s" by auto then show ?thesis .. qed lemma flow_leaves_closed_at_frontierE: assumes "closed M" and t_ex: "t \ existence_ivl0 x" and "0 \ t" "x \ M" "flow0 x t \ M" obtains s where "0 \ s" "s < t" "flow0 x ` {0..s} \ M" "flow0 x s \ frontier M" "\\<^sub>F s' in at_right s. flow0 x s' \ M" proof - from flow_in_closed_max_timeE[OF assms(1-4)] assms(5) obtain s where s: "0 \ s" "s < t" "flow0 x ` {0..s} \ M" and s_max: "\s'. 0 \ s' \ s' \ t \ flow0 x ` {0..s'} \ M \ s' \ s" by (smt atLeastAtMost_iff image_subset_iff) note s moreover have "flow0 x s \ interior M" proof assume interior: "flow0 x s \ interior M" have "s \ existence_ivl0 x" using ivl_subset_existence_ivl[OF \t \ _\] s by auto from flow_continuous[OF this, THEN isContD, THEN topological_tendstoD, OF open_interior interior] have "\\<^sub>F s' in at s. flow0 x s' \ interior M" by auto then have "\\<^sub>F s' in at_right s. flow0 x s' \ interior M" by (auto simp: eventually_at_split) moreover have "\\<^sub>F s' in at_right s. s' < t" using tendsto_ident_at \s < t\ by (rule order_tendstoD) ultimately have "\\<^sub>F s' in at_right s. flow0 x s' \ M \ s' < t" by eventually_elim (use interior_subset[of M] in auto) then obtain s' where s': "s < s'" "s' < t" "\y. y > s \ y \ s' \ flow0 x y \ M" by (auto simp: eventually_at_right_field_le) have s'_ivl: "flow0 x ` {0..s'} \ M" proof safe fix s'' assume "s'' \ {0 .. s'}" then show "flow0 x s'' \ M" using s interior_subset[of M] s' by (cases "s'' \ s") auto qed with s_max[of s'] \s' < t\ \0 \ s\ \s < s'\ show False by auto qed then have "flow0 x s \ frontier M" using s closure_subset[of M] by (force simp: frontier_def) moreover have "compact (flow0 x -` M \ {s..t})" (is "compact ?C") unfolding compact_eq_bounded_closed proof have "bounded {s .. t}" by simp then show "bounded ?C" by (rule bounded_subset) auto show "closed ?C" using \closed M\ assms mem_existence_ivl_iv_defined(2)[OF t_ex] ivl_subset_existence_ivl[OF t_ex] \0 \ s\ by (intro closed_vimage_Int) (auto intro!: continuous_intros) qed have "\\<^sub>F s' in at_right s. flow0 x s' \ M" apply (rule ccontr) unfolding not_frequently proof - assume "\\<^sub>F s' in at_right s. \ flow0 x s' \ M" moreover have "\\<^sub>F s' in at_right s. s' < t" using tendsto_ident_at \s < t\ by (rule order_tendstoD) ultimately have "\\<^sub>F s' in at_right s. flow0 x s' \ M \ s' < t" by eventually_elim auto then obtain s' where s': "s < s'" "\y. y > s \ y < s' \ flow0 x y \ M" "\y. y > s \ y < s' \ y < t" by (auto simp: eventually_at_right_field) define s'' where "s'' = (s + s') / 2" have "0 \ s''" "s'' \ t" "s < s''" "s'' < s'" using s s' by (auto simp del: divide_le_eq_numeral1 le_divide_eq_numeral1 simp: s''_def) fastforce then have "flow0 x ` {0..s''} \ M" using s s' apply auto subgoal for u by (cases "u\s") auto done from s_max[OF \0 \ s''\ \s''\ t\ this] \s'' > s\ show False by simp qed ultimately show ?thesis .. qed subsection \Connectedness\ lemma fcontX: shows "continuous_on X f" using auto_local_lipschitz local_lipschitz_continuous_on by blast lemma fcontx: assumes "x \ X" shows "continuous (at x) f" proof - have "open X" by simp from continuous_on_eq_continuous_at[OF this] show ?thesis using fcontX assms(1) by blast qed lemma continuous_at_imp_cball: assumes "continuous (at x) g" assumes "g x > (0::real)" obtains r where "r > 0" "\y \ cball x r. g y > 0" proof - from assms(1) obtain d where "d>0" "g ` (ball x d) \ ball (g x) ((g x)/2)" by (meson assms(2) continuous_at_ball half_gt_zero) then have "\y \ cball x (d/2). g y > 0" by (smt assms(2) dist_norm image_subset_iff mem_ball mem_cball pos_half_less real_norm_def) thus ?thesis using \0 < d\ that half_gt_zero by blast qed text \ \flow0\ is \path_connected\ \ lemma flow0_path_connected_time: assumes "ts \ existence_ivl0 x" "path_connected ts" shows "path_connected (flow0 x ` ts)" proof - have "continuous_on ts (flow0 x)" by (meson assms continuous_on_sequentially flow_continuous_on subsetD) from path_connected_continuous_image[OF this assms(2)] show ?thesis . qed lemma flow0_path_connected: assumes "path_connected D" "path_connected ts" "\x. x \ D \ ts \ existence_ivl0 x" shows "path_connected ( (\(x, y). flow0 x y) ` (D \ ts))" proof - have "D \ ts \ Sigma X existence_ivl0" using assms(3) subset_iff by fastforce then have a1:"continuous_on (D \ ts) (\(x, y). flow0 x y)" using flow_continuous_on_state_space continuous_on_subset by blast have a2 : "path_connected (D \ ts)" using path_connected_Times assms by auto from path_connected_continuous_image[OF a1 a2] show ?thesis . qed end subsection \Return Time and Implicit Function Theorem\ context c1_on_open_euclidean begin lemma flow_implicit_function: \ \TODO: generalization of @{thm returns_to_implicit_function}!\ fixes s::"'a::euclidean_space \ real" and S::"'a set" assumes t: "t \ existence_ivl0 x" and x: "x \ X" and st: "s (flow0 x t) = 0" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "isCont Ds (flow0 x t)" assumes nz: "Ds (flow0 x t) (f (flow0 x t)) \ 0" obtains u e where "s (flow0 x (u x)) = 0" "u x = t" "(\y. y \ cball x e \ s (flow0 y (u y)) = 0)" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ Sigma X existence_ivl0" "0 < e" "(u has_derivative (- blinfun_scaleR_left (inverse (blinfun_apply (Ds (flow0 x t)) (f (flow0 x t)))) o\<^sub>L (Ds (flow0 x t) o\<^sub>L flowderiv x t) o\<^sub>L embed1_blinfun)) (at x)" proof - note [derivative_intros] = has_derivative_compose[OF _ Ds] have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds]) note cls[simp, intro] = closed_levelset[OF cont_s] then have xt1: "(x, t) \ Sigma X existence_ivl0" by (auto simp: t x) have D: "(\x. x \ Sigma X existence_ivl0 \ ((\(x, t). s (flow0 x t)) has_derivative blinfun_apply (Ds (flow0 (fst x) (snd x)) o\<^sub>L (flowderiv (fst x) (snd x)))) (at x))" by (auto intro!: derivative_eq_intros) have C: "isCont (\x. Ds (flow0 (fst x) (snd x)) o\<^sub>L flowderiv (fst x) (snd x)) (x, t)" using flowderiv_continuous_on[unfolded continuous_on_eq_continuous_within, rule_format, OF xt1] using at_within_open[OF xt1 open_state_space] by (auto intro!: continuous_intros tendsto_eq_intros x t isCont_tendsto_compose[OF DsC, unfolded poincare_map_def] simp: split_beta' isCont_def) have Z: "(case (x, t) of (x, t) \ s (flow0 x t)) = 0" by (auto simp: st) have I1: "blinfun_scaleR_left (inverse (Ds (flow0 x t)(f (flow0 x t)))) o\<^sub>L ((Ds (flow0 (fst (x, t)) (snd (x, t))) o\<^sub>L flowderiv (fst (x, t)) (snd (x, t))) o\<^sub>L embed2_blinfun) = 1\<^sub>L" using nz by (auto intro!: blinfun_eqI simp: flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) have I2: "((Ds (flow0 (fst (x, t)) (snd (x, t))) o\<^sub>L flowderiv (fst (x, t)) (snd (x, t))) o\<^sub>L embed2_blinfun) o\<^sub>L blinfun_scaleR_left (inverse (Ds (flow0 x t)(f (flow0 x t)))) = 1\<^sub>L" using nz by (auto intro!: blinfun_eqI simp: flowderiv_def blinfun.bilinear_simps inverse_eq_divide poincare_map_def) show ?thesis apply (rule implicit_function_theorem[where f="\(x, t). s (flow0 x t)" and S="Sigma X existence_ivl0", OF D xt1 open_state_space order_refl C Z I1 I2]) apply blast unfolding split_beta' fst_conv snd_conv poincare_map_def[symmetric] .. qed lemma flow_implicit_function_at: fixes s::"'a::euclidean_space \ real" and S::"'a set" assumes x: "x \ X" and st: "s x = 0" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" assumes DsC: "isCont Ds x" assumes nz: "Ds x (f x) \ 0" assumes pos: "e > 0" obtains u d where "0 < d" "u x = 0" "\y. y \ cball x d \ s (flow0 y (u y)) = 0" "\y. y \ cball x d \ \u y\ < e" "\y. y \ cball x d \ u y \ existence_ivl0 y" "continuous_on (cball x d) u" "(u has_derivative -Ds x /\<^sub>R (Ds x) (f x)) (at x)" proof - have x0: "flow0 x 0 = x" by (simp add: x) from flow_implicit_function[OF existence_ivl_zero[OF x] x, unfolded x0, of s, OF st Ds DsC nz] obtain u d0 where s0: "s (flow0 x (u x)) = 0" and u0: "u x = 0" and u: "\y. y \ cball x d0 \ s (flow0 y (u y)) = 0" and uc: "continuous_on (cball x d0) u" and uex: "(\t. (t, u t)) ` cball x d0 \ Sigma X existence_ivl0" and d0: "0 < d0" and u': "(u has_derivative blinfun_apply (- blinfun_scaleR_left (inverse (blinfun_apply (Ds x) (f x))) o\<^sub>L (Ds x o\<^sub>L flowderiv x 0) o\<^sub>L embed1_blinfun)) (at x)" by blast have "at x within cball x d0 = at x" by (rule at_within_interior) (auto simp: \0 < d0\) then have "(u \ 0) (at x)" using uc d0 by (auto simp: continuous_on_def u0 dest!: bspec[where x=x]) from tendstoD[OF this \0 < e\] pos u0 obtain d1 where d1: "0 < d1" "\xa. dist xa x \ d1 \ \u xa\ < e" unfolding eventually_at_le by force define d where "d = min d0 d1" have "0 < d" by (auto simp: d_def d0 d1) moreover note u0 moreover have "\y. y \ cball x d \ s (flow0 y (u y)) = 0" by (auto intro!: u simp: d_def) moreover have "\y. y \ cball x d \ \u y\ < e" using d1 by (auto simp: d_def dist_commute) moreover have "\y. y \ cball x d \ u y \ existence_ivl0 y" using uex by (force simp: d_def) moreover have "continuous_on (cball x d) u" using uc by (rule continuous_on_subset) (auto simp: d_def) moreover have "(u has_derivative -Ds x /\<^sub>R (Ds x) (f x)) (at x)" using u' by (rule has_derivative_subst) (auto intro!: ext simp: x x0 flowderiv_def blinfun.bilinear_simps) ultimately show ?thesis .. qed lemma returns_to_implicit_function_gen: \ \TODO: generalizes proof of @{thm returns_to_implicit_function}!\ fixes s::"'a::euclidean_space \ real" assumes rt: "returns_to {x \ S. s x = 0} x" (is "returns_to ?P x") assumes cS: "closed S" assumes Ds: "\x. (s has_derivative blinfun_apply (Ds x)) (at x)" "isCont Ds (poincare_map ?P x)" "Ds (poincare_map ?P x) (f (poincare_map ?P x)) \ 0" obtains u e where "s (flow0 x (u x)) = 0" "u x = return_time ?P x" "(\y. y \ cball x e \ s (flow0 y (u y)) = 0)" "continuous_on (cball x e) u" "(\t. (t, u t)) ` cball x e \ Sigma X existence_ivl0" "0 < e" "(u has_derivative (- blinfun_scaleR_left (inverse (blinfun_apply (Ds (poincare_map ?P x)) (f (poincare_map ?P x)))) o\<^sub>L (Ds (poincare_map ?P x) o\<^sub>L flowderiv x (return_time ?P x)) o\<^sub>L embed1_blinfun)) (at x)" proof - note [derivative_intros] = has_derivative_compose[OF _ Ds(1)] have cont_s: "continuous_on UNIV s" by (rule has_derivative_continuous_on[OF Ds(1)]) note cls[simp, intro] = closed_levelset[OF cont_s] let ?t1 = "return_time ?P x" have cls[simp, intro]: "closed {x \ S. s x = 0}" by (rule closed_levelset_within) (auto intro!: cS continuous_on_subset[OF cont_s]) have *: "poincare_map ?P x = flow0 x (return_time {x \ S. s x = 0} x)" by (simp add: poincare_map_def) have "return_time {x \ S. s x = 0} x \ existence_ivl0 x" "x \ X" "s (poincare_map ?P x) = 0" using poincare_map_returns rt by (auto intro!: return_time_exivl rt) note E = flow_implicit_function[of "return_time ?P x" x s Ds, OF this[unfolded *] Ds[unfolded *], folded *] show ?thesis by (rule E) rule qed text \c.f. Perko Section 3.7 Lemma 2 part 1.\ lemma flow_transversal_surface_finite_intersections: fixes s::"'a \ 'b::real_normed_vector" and Ds::"'a \ 'a \\<^sub>L 'b" assumes "closed S" assumes "\x. (s has_derivative (Ds x)) (at x)" assumes "\x. x \ S \ s x = 0 \ Ds x (f x) \ 0" assumes "a \ b" "{a .. b} \ existence_ivl0 x" shows "finite {t\{a..b}. flow0 x t \ {x \ S. s x = 0}}" \ \TODO: define notion of (compact/closed)-(continuous/differentiable/C1)-surface?\ proof cases note Ds = \\x. (s has_derivative (Ds x)) (at x)\ note transversal = \\x. x \ S \ s x = 0 \ Ds x (f x) \ 0\ assume "a < b" show ?thesis proof (rule ccontr) let ?S = "{x \ S. s x = 0}" let ?T = "{t\{a..b}. flow0 x t \ {x \ S. s x = 0}}" define \ where "\ = flow0 x" have [THEN continuous_on_compose2, continuous_intros]: "continuous_on S s" by (auto simp: intro!: has_derivative_continuous_on Ds intro: has_derivative_at_withinI) assume "infinite ?T" from compact_sequentialE[OF compact_Icc[of a b] this] obtain t tl where t: "t n \ ?T" "flow0 x (t n) \ ?S" "t n \ {a .. b}" "t n \ tl" and tl: "t \ tl" "tl \ {a..b}" for n by force have tl_ex: "tl \ existence_ivl0 x" using \{a .. b} \ existence_ivl0 x\ \tl \ {a .. b}\ by auto have "closed ?S" by (auto intro!: closed_levelset_within \closed S\ continuous_intros) moreover have "\n. flow0 x (t n) \ ?S" using t by auto moreover have flow_t: "(\n. flow0 x (t n)) \ flow0 x tl" by (auto intro!: tendsto_eq_intros tl_ex tl) ultimately have "flow0 x tl \ ?S" - by (rule closed_sequentially) + by (metis (no_types, lifting) closed_sequentially) let ?qt = "\t. (flow0 x t - flow0 x tl) /\<^sub>R (t - tl)" from flow_has_vector_derivative[OF tl_ex, THEN has_vector_derivative_withinD] have qt_tendsto: "?qt \tl\ f (flow0 x tl)" . let ?q = "\n. ?qt (t n)" have "filterlim t (at tl) sequentially" using tl(1) by (rule filterlim_atI) (simp add: t) with qt_tendsto have "?q \ f (flow0 x tl)" by (rule filterlim_compose) then have "((\n. Ds (flow0 x tl) (?q n))) \ Ds (flow0 x tl) (f (flow0 x tl))" by (auto intro!: tendsto_intros) moreover from flow_lipschitzE[OF \{a .. b} \ existence_ivl0 x\] obtain L' where L': "L'-lipschitz_on {a..b} (flow0 x)" . define L where "L = L' + 1" from lipschitz_on_le[OF L', of L] lipschitz_on_nonneg[OF L'] have L: "L-lipschitz_on {a .. b} (flow0 x)" and "L > 0" by (auto simp: L_def) from flow_lipschitzE[OF \{a .. b} \ existence_ivl0 x\] obtain L' where "L'-lipschitz_on {a..b} (flow0 x)" . \ \TODO: is this reasoning (below) with this Lipschitz constant really necessary?\ have s[simp]: "s (flow0 x (t n)) = 0""s (flow0 x tl) = 0" for n using t \flow0 x tl \ ?S\ by auto from Ds(1)[of "flow0 x tl", unfolded has_derivative_within] have "(\y. (1 / norm (y - flow0 x tl)) *\<^sub>R (s y - (s (flow0 x tl) + blinfun_apply (Ds (flow0 x tl)) (y - flow0 x tl)))) \flow0 x tl\ 0" by auto then have "((\y. (1 / norm (y - flow0 x tl)) *\<^sub>R (s y - (s (flow0 x tl) + blinfun_apply (Ds (flow0 x tl)) (y - flow0 x tl)))) \ 0) (nhds (flow0 x tl))" by (rule tendsto_nhds_continuousI) simp from filterlim_compose[OF this flow_t] have "(\xa. (blinfun_apply (Ds (flow0 x tl)) (flow0 x (t xa) - flow0 x tl)) /\<^sub>R norm (flow0 x (t xa) - flow0 x tl)) \ 0" using t by (auto simp: inverse_eq_divide tendsto_minus_cancel_right) from tendsto_mult[OF tendsto_const[of "L"] tendsto_norm[OF this, simplified, simplified divide_inverse_commute[symmetric]]]\ \TODO: uuugly\ have Ds0: "(\xa. norm (blinfun_apply (Ds (flow0 x tl)) (flow0 x (t xa) - flow0 x tl)) / (norm (flow0 x (t xa) - flow0 x tl)/(L))) \ 0" by (auto simp: ac_simps) from _ Ds0 have "((\n. Ds (flow0 x tl) (?q n)) \ 0)" apply (rule Lim_null_comparison) apply (rule eventuallyI) unfolding norm_scaleR blinfun.scaleR_right abs_inverse divide_inverse_commute[symmetric] subgoal for n apply (cases "flow0 x (t n) = flow0 x tl") subgoal by (simp add: blinfun.bilinear_simps) subgoal apply (rule divide_left_mono) using lipschitz_onD[OF L, of "t n" tl] \0 < L\ t(3) tl(2) by (auto simp: algebra_split_simps zero_less_divide_iff dist_norm pos_divide_le_eq intro!: add_pos_nonneg) done done ultimately have "Ds (flow0 x tl) (f (flow0 x tl)) = 0" by (rule LIMSEQ_unique) moreover have "Ds (flow0 x tl) (f (flow0 x tl)) \ 0" by (rule transversal) (use \flow0 x tl \ ?S\ in auto) ultimately show False by auto qed qed (use assms in auto) lemma uniform_limit_flow0_state:\ \TODO: is that something more general?\ assumes "compact C" assumes "C \ X" shows "uniform_limit C (\s x. flow0 x s) (\x. flow0 x 0) (at 0)" proof (cases "C = {}") case True then show ?thesis by auto next case False show ?thesis proof (rule uniform_limitI) fix e::real assume "0 < e" { fix x assume "x \ C" with assms have "x \ X" by auto from existence_ivl_cballs[OF UNIV_I \x \ X\] obtain t L u where "\y. y \ cball x u \ cball 0 t \ existence_ivl0 y" "\s y. y \ cball x u \ s \ cball 0 t \ flow0 y s \ cball y u" "L-lipschitz_on (cball 0 t\cball x u) (\(t, x). flow0 x t)" "\y. y \ cball x u \ cball y u \ X" "0 < t" "0 < u" by metis then have "\L. \u>0. \t>0. L-lipschitz_on (cball 0 t\cball x u) (\(t, x). flow0 x t)" by blast } then have "\x\C. \L. \u>0. \t>0. L-lipschitz_on (cball 0 t\cball x u) (\(t, x). flow0 x t)" .. then obtain L d' u' where L: "\x. x \ C \ (L x)-lipschitz_on (cball 0 (d' x)\cball x (u' x)) (\(t, x). flow0 x t)" and d': "\x. x \ C \ d' x > 0" and u': "\x. x \ C \ u' x > 0" by metis have "C \ (\c\C. ball c (u' c))" using u' by auto from compactE_image[OF \compact C\ _ this] obtain C' where "C' \ C" and [simp]: "finite C'" and C'_cover: "C \ (\c\C'. ball c (u' c))" by auto from C'_cover obtain c' where c': "x \ C \ x \ ball (c' x) (u' (c' x))" "x \ C \ c' x \ C'" for x by (auto simp: subset_iff) metis have "\\<^sub>F s in at 0. \x\ball c (u' c). dist (flow0 x s) (flow0 x 0) < e" if "c \ C'" for c proof - have cC: "c \ C" using c' \c \ C'\ d' \C' \ C\ by auto have *: "dist (flow0 x s) (flow0 x 0) \ L c * \s\" if "x\ball c (u' c)" "s \ cball 0 (d' c)" for x s proof - from L[OF cC, THEN lipschitz_onD, of "(0, x)" "(s, x)"] d'[OF cC] that show ?thesis by (auto simp: dist_prod_def dist_commute) qed have "\\<^sub>F s in at 0. abs s < d' c" by (rule order_tendstoD tendsto_intros)+ (use d' cC in auto) moreover have "\\<^sub>F s in at 0. L c * \s\ < e" by (rule order_tendstoD tendsto_intros)+ (use \0 < e\ in auto) ultimately show ?thesis apply eventually_elim apply (use * in auto) by smt qed then have "\\<^sub>F s in at 0. \c\C'. \x\ball c (u' c). dist (flow0 x s) (flow0 x 0) < e" by (simp add: eventually_ball_finite_distrib) then show "\\<^sub>F s in at 0. \x\C. dist (flow0 x s) (flow0 x 0) < e" apply eventually_elim apply auto subgoal for s x apply (drule bspec[where x="c' x"]) apply (simp add: c'(2)) apply (drule bspec) prefer 2 apply assumption apply auto using c'(1) by auto done qed qed end subsection \Fixpoints\ context auto_ll_on_open begin lemma fixpoint_sol: assumes "x \ X" "f x = 0" shows "existence_ivl0 x = UNIV" "flow0 x t = x" proof - have sol: "((\t::real. x) solves_ode (\_. f)) UNIV X" apply (rule solves_odeI) by(auto simp add: assms intro!: derivative_intros) from maximal_existence_flow[OF sol] have "UNIV \ existence_ivl0 x" "flow0 x t = x" by auto thus "existence_ivl0 x = UNIV" "flow0 x t = x" by auto qed end end \ No newline at end of file