Theory Flow
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⇩L")
lemma blinfun_inverse_left:
fixes f::"'a::euclidean_space ⇒⇩L 'a" and f'
shows "f o⇩L f' = 1⇩L ⟷ f' o⇩L f = 1⇩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⇩L 1⇩L = x"
and blinfun_compose_1_right[simp]: "1⇩L o⇩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) *⇩R c"
by (auto simp: closed_segment_eq_real_ivl content_closed_segment_real)
lemmas [integrable_on_simps] =
integrable_on_empty
integrable_on_refl integrable_on_refl_ivl integrable_on_refl_closed_segment
integrable_const integrable_const_ivl integrable_const_ivl_closed_segment
ident_integrable_on integrable_ident_ivl integrable_ident_cbox
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_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 *⇩R f t (flow t0 x0 t))) (at t)"
proof -
have "(flow t0 x0 has_derivative (λi. i *⇩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--<b} ⟹ 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--<b})"
and limsol: "((λ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" "∀⇩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 "∀⇩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 "∀⇩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"
"∀⇩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