Theory Poincare_Map

theory Poincare_Map
imports
  Flow
begin

abbreviation "plane n c  {x. x  n = c}"

lemma
  eventually_tendsto_compose_within:
  assumes "eventually P (at l within S)"
  assumes "P l"
  assumes "(f  l) (at x within T)"
  assumes "eventually (λx. f x  S) (at x within T)"
  shows "eventually (λx. P (f x)) (at x within T)"
proof -
  from assms(1) assms(2) obtain U where U:
    "open U" "l  U" "x. x  U  x  S  P x"
    by (force simp: eventually_at_topological)
  from topological_tendstoD[OF assms(3) open U l  U]
  have "F x in at x within T. f x  U" by auto
  then show ?thesis using assms(4)
    by eventually_elim (auto intro!: U)
qed

lemma
  eventually_eventually_withinI:― ‹aha...›
  assumes "F x in at x within A. P x" "P x"
  shows "F a in at x within S. F x in at a within A. P x"
  using assms
  unfolding eventually_at_topological
  by force

lemma eventually_not_in_closed:
  assumes "closed P"
  assumes "f t  P" "t  T"
  assumes "continuous_on T f"
  shows "F t in at t within T. f t  P"
  using assms
  unfolding Compl_iff[symmetric] closed_def continuous_on_topological eventually_at_topological
  by metis

context ll_on_open_it begin

lemma
  existence_ivl_trans':
  assumes "t + s  existence_ivl t0 x0"
    "t  existence_ivl t0 x0"
  shows "t + s  existence_ivl t (flow t0 x0 t)"
  by (meson assms(1) assms(2) general.existence_ivl_reverse general.flow_solves_ode
      general.is_interval_existence_ivl general.maximal_existence_flow(1)
      general.mem_existence_ivl_iv_defined(2) general.mem_existence_ivl_subset
      local.existence_ivl_subset subsetD)

end

context auto_ll_on_open― ‹TODO: generalize to continuous systems›
begin

definition returns_to ::"'a set  'a  bool"
  where "returns_to P x  (F t in at_right 0. flow0 x t  P)  (t>0. t  existence_ivl0 x  flow0 x t  P)"

definition return_time :: "'a set  'a  real"
  where "return_time P x =
    (if returns_to P x then (SOME t.
      t > 0 
      t  existence_ivl0 x 
      flow0 x t  P 
      (s  {0<..<t}. flow0 x s  P)) else 0)"

lemma returns_toI:
  assumes t: "t > 0" "t  existence_ivl0 x" "flow0 x t  P"
  assumes ev: "F t in at_right 0. flow0 x t  P"
  assumes "closed P"
  shows "returns_to P x"
  using assms
  by (auto simp: returns_to_def)

lemma returns_to_outsideI:
  assumes t: "t  0" "t  existence_ivl0 x" "flow0 x t  P"
  assumes ev: "x  P"
  assumes "closed P"
  shows "returns_to P x"
proof cases
  assume "t > 0"
  moreover
  have "F s in at 0 within {0 .. t}. flow0 x s  P"
    using assms mem_existence_ivl_iv_defined ivl_subset_existence_ivl[OF t  _] 0 < t
    by (auto intro!: eventually_not_in_closed flow_continuous_on continuous_intros
        simp: eventually_conj_iff)
  with order_tendstoD(2)[OF tendsto_ident_at 0 < t, of "{0<..}"]
  have "F t in at_right 0. flow0 x t  P"
    unfolding eventually_at_filter
    by eventually_elim (use t > 0 in auto)
  then show ?thesis
    by (auto intro!: returns_toI assms 0 < t)
qed (use assms in simp)

lemma returns_toE:
  assumes "returns_to P x"
  obtains t0 t1 where
    "0 < t0"
    "t0  t1"
    "t1  existence_ivl0 x"
    "flow0 x t1  P"
    "t. 0 < t  t < t0  flow0 x t  P"
proof -
  obtain t0 t1 where t0: "t0 > 0" "t. 0 < t  t < t0  flow0 x t  P"
    and t1: "t1 > 0" "t1  existence_ivl0 x" "flow0 x t1  P"
    using assms
    by (auto simp: returns_to_def eventually_at_right[OF zero_less_one])
  moreover
  have "t0  t1"
    using t0(2)[of t1] t1 t0(1)
    by force
  ultimately show ?thesis by (blast intro: that)
qed

lemma return_time_some:
  assumes "returns_to P x"
  shows "return_time P x =
    (SOME t. t > 0  t  existence_ivl0 x  flow0 x t  P  (s  {0<..<t}. flow0 x s  P))"
  using assms by (auto simp: return_time_def)

lemma return_time_ex1:
  assumes "returns_to P x"
  assumes "closed P"
  shows "∃!t. t > 0  t  existence_ivl0 x  flow0 x t  P  (s  {0<..<t}. flow0 x s  P)"
proof -
  from returns_toE[OF returns_to P x]
  obtain t0 t1 where
    t1: "t1  t0" "t1  existence_ivl0 x" "flow0 x t1  P"
    and t0: "t0 > 0" "t. 0 < t  t < t0  flow0 x t  P"
    by metis
  from flow_continuous_on have cont: "continuous_on {0 .. t1} (flow0 x)"
    by (rule continuous_on_subset) (intro ivl_subset_existence_ivl t1)
  from cont have cont': "continuous_on {t0 .. t1} (flow0 x)"
    by (rule continuous_on_subset) (use 0 < t0 in auto)
  have "compact (flow0 x -` P  {t0 .. t1})"
    using closed P cont'
    by (auto simp: compact_eq_bounded_closed bounded_Int bounded_closed_interval
        intro!: closed_vimage_Int)

  have "flow0 x -` P  {t0..t1}  {}"
    using t1 t0 by auto
  from compact_attains_inf[OF compact _ this] t0 t1
  obtain rt where rt: "t0  rt" "rt  t1" "flow0 x rt  P"
    and least: "t'. flow0 x t'  P  t0  t'  t'  t1  rt  t'"
    by auto
  have "0 < rt" "flow0 x rt  P" "rt  existence_ivl0 x"
    and "0 < t'  t' < rt  flow0 x t'  P" for t'
    using ivl_subset_existence_ivl[OF t1  existence_ivl0 x] t0 t1 rt least[of t']
    by force+
  then show ?thesis
    by (intro ex_ex1I) force+
qed

lemma
  return_time_pos_returns_to:
  "return_time P x > 0  returns_to P x"
  by (auto simp: return_time_def split: if_splits)

lemma
  assumes ret: "returns_to P x"
  assumes "closed P"
  shows return_time_pos: "return_time P x > 0"
  using someI_ex[OF return_time_ex1[OF assms, THEN ex1_implies_ex]]
  unfolding return_time_some[OF ret, symmetric]
  by auto

lemma returns_to_return_time_pos:
  assumes "closed P"
  shows "returns_to P x  return_time P x > 0"
  by (auto intro!: return_time_pos assms) (auto simp: return_time_def split: if_splits)

lemma return_time:
  assumes ret: "returns_to P x"
  assumes "closed P"
  shows "return_time P x > 0"
    and return_time_exivl: "return_time P x  existence_ivl0 x"
    and return_time_returns: "flow0 x (return_time P x)  P"
    and return_time_least: "s. 0 < s  s < return_time P x  flow0 x s  P"
  using someI_ex[OF return_time_ex1[OF assms, THEN ex1_implies_ex]]
  unfolding return_time_some[OF ret, symmetric]
  by auto

lemma returns_to_earlierI:
  assumes ret: "returns_to P (flow0 x t)" "closed P"
  assumes "t  0" "t  existence_ivl0 x"
  assumes ev: "F t in at_right 0. flow0 x t  P"
  shows "returns_to P x"
proof -
  from return_time[OF ret]
  have rt: "0 < return_time P (flow0 x t)" "flow0 (flow0 x t) (return_time P (flow0 x t))  P"
    and "0 < s  s < return_time P (flow0 x t)  flow0 (flow0 x t) s  P" for s
    by auto
  let ?t = "t + return_time P  (flow0 x t)"
  show ?thesis
  proof (rule returns_toI[of ?t])
    show "0 < ?t" by (auto intro!: add_nonneg_pos rt t  0)
    show "?t  existence_ivl0 x"
      by (intro existence_ivl_trans return_time_exivl assms)
    have "flow0 x (t + return_time P (flow0 x t)) = flow0 (flow0 x t) (return_time P (flow0 x t))"
      by (intro flow_trans assms return_time_exivl)
    also have "  P"
      by (rule return_time_returns[OF ret])
    finally show "flow0 x (t + return_time P (flow0 x t))  P" .
    show "closed P" by fact
    show "F t in at_right 0. flow0 x t  P" by fact
  qed
qed

lemma return_time_gt:
  assumes ret: "returns_to P x" "closed P"
  assumes flow_not: "s. 0 < s  s  t  flow0 x s  P"
  shows "t < return_time P x"
  using flow_not[of "return_time P x"] return_time_pos[OF ret] return_time_returns[OF ret] by force

lemma return_time_le:
  assumes ret: "returns_to P x" "closed P"
  assumes flow_not: "flow0 x t  P" "t > 0"
  shows "return_time P x  t"
  using return_time_least[OF assms(1,2), of t] flow_not
  by force

lemma returns_to_laterI:
  assumes ret: "returns_to P x" "closed P"
  assumes t: "t > 0" "t  existence_ivl0 x"
  assumes flow_not: "s. 0 < s  s  t  flow0 x s  P"
  shows "returns_to P (flow0 x t)"
  apply (rule returns_toI[of "return_time P x - t"])
  subgoal using flow_not by (auto intro!: return_time_gt ret)
  subgoal by (auto intro!: existence_ivl_trans' return_time_exivl ret t)
  subgoal by (subst flow_trans[symmetric])
      (auto intro!: existence_ivl_trans' return_time_exivl ret t return_time_returns)
  subgoal
  proof -
    have "F y in nhds 0. y  existence_ivl0 (flow0 x t)"
      apply (rule eventually_nhds_in_open[OF open_existence_ivl[of "flow0 x t"] existence_ivl_zero])
      apply (rule flow_in_domain)
      apply fact
      done
    then have "F s in at_right 0. s  existence_ivl0 (flow0 x t)"
      unfolding eventually_at_filter
      by eventually_elim auto
    moreover
    have "F s in at_right 0. t + s < return_time P x"
      using return_time_gt[OF ret flow_not, of t]
      by (auto simp: eventually_at_right[OF zero_less_one] intro!: exI[of _ "return_time P x - t"])
    moreover
    have "F s in at_right 0. 0 < t + s"
      by (metis (mono_tags) eventually_at_rightI greaterThanLessThan_iff pos_add_strict t(1))
    ultimately show ?thesis
      apply eventually_elim
      apply (subst flow_trans[symmetric])
      using return_time_least[OF ret]
      by (auto intro!: existence_ivl_trans' t)
    qed
  subgoal by fact
  done

lemma never_returns:
  assumes "¬returns_to P x"
  assumes "closed P" "t  0" "t  existence_ivl0 x"
  assumes ev: "F t in at_right 0. flow0 x t  P"
  shows "¬returns_to P (flow0 x t)"
  using returns_to_earlierI[OF _ assms(2-5)] assms(1)
  by blast

lemma return_time_eqI:
  assumes "closed P"
    and t_pos: "t > 0"
    and ex: "t  existence_ivl0 x"
    and ret: "flow0 x t  P"
    and least: "s. 0 < s  s < t  flow0 x s  P"
  shows "return_time P x = t"
proof -
  from least t_pos have "F t in at_right 0. flow0 x t  P"
    by (auto simp: eventually_at_right[OF zero_less_one])
  then have "returns_to P x"
    by (auto intro!: returns_toI[of t] assms)
  then show ?thesis
    using least
    by (auto simp: return_time_def t_pos ex ret
        intro!: some1_equality[OF return_time_ex1[OF returns_to _ _ closed _]])
qed

lemma return_time_step:
  assumes "returns_to P (flow0 x t)"
  assumes "closed P"
  assumes flow_not: "s. 0 < s  s  t  flow0 x s  P"
  assumes t: "t > 0" "t  existence_ivl0 x"
  shows "return_time P (flow0 x t) = return_time P x - t"
proof -
  from flow_not t have "F t in at_right 0. flow0 x t  P"
    by (auto simp: eventually_at_right[OF zero_less_one])
  from returns_to_earlierI[OF assms(1,2) less_imp_le, OF t this]
  have ret: "returns_to P x" .
  from return_time_gt[OF ret closed P flow_not]
  have "t < return_time P x" by simp
  moreover
  have "0 < s  s < return_time P x - t  flow0 (flow0 x t) s = flow0 x (t + s)" for s
    using ivl_subset_existence_ivl[OF return_time_exivl[OF ret closed _]] t
    by (subst flow_trans) (auto intro!: existence_ivl_trans')
  ultimately show ?thesis
    using flow_not assms(1) ret return_time_least t(1)
    by (auto intro!: return_time_eqI return_time_returns ret
        simp: flow_trans[symmetric] closed P t(2) existence_ivl_trans' return_time_exivl)
qed

definition "poincare_map P x = flow0 x (return_time P x)"

lemma poincare_map_step_flow:
  assumes ret: "returns_to P x" "closed P"
  assumes flow_not: "s. 0 < s  s  t  flow0 x s  P"
  assumes t: "t > 0" "t  existence_ivl0 x"
  shows "poincare_map P (flow0 x t) = poincare_map P x"
  unfolding poincare_map_def
  apply (subst flow_trans[symmetric])
  subgoal by fact
  subgoal using flow_not by (auto intro!: return_time_exivl returns_to_laterI t ret)
  subgoal
    using flow_not
    by (subst return_time_step) (auto intro!: return_time_exivl returns_to_laterI t ret)
  done

lemma poincare_map_returns:
  assumes "returns_to P x" "closed P"
  shows "poincare_map P x  P"
  by (auto intro!: return_time_returns assms simp: poincare_map_def)

lemma poincare_map_onto:
  assumes "closed P"
  assumes "0 < t" "t  existence_ivl0 x" "F t in at_right 0. flow0 x t  P"
  assumes "flow0 x t  P"
  shows "poincare_map P x  flow0 x ` {0 <.. t}  P"
proof (rule IntI)
  have "returns_to P x"
    by (rule returns_toI) (rule assms)+
  then have "return_time P x  {0<..t}"
    by (auto intro!: return_time_pos assms return_time_le)
  then show "poincare_map P x  flow0 x ` {0<..t}"
    by (auto simp: poincare_map_def)
  show "poincare_map P x  P"
    by (auto intro!: poincare_map_returns returns_to _ _ closed _)
qed

end


lemma isCont_blinfunD:
  fixes f'::"'a::metric_space  'b::real_normed_vector L 'c::real_normed_vector"
  assumes "isCont f' a" "0 < e"
  shows "d>0. x. dist a x < d  onorm (λv. blinfun_apply (f' x) v - blinfun_apply (f' a) v) < e"
proof -
  have "F x in at a. dist (f' x) (f' a) < e"
    using assms isCont_def tendsto_iff by blast
  then show ?thesis
    using e > 0 norm_eq_zero
    by (force simp: eventually_at dist_commute dist_norm norm_blinfun.rep_eq
        simp flip: blinfun.bilinear_simps)
qed

proposition has_derivative_locally_injective_blinfun:
  fixes f :: "'n::euclidean_space  'm::euclidean_space"
    and f'::"'n  'n L 'm"
    and g'::"'m L 'n"
  assumes "a  s"
      and "open s"
      and g': "g' oL (f' a) = 1L"
      and f': "x. x  s  (f has_derivative f' x) (at x)"
      and c: "isCont f' a"
    obtains r where "r > 0" "ball a r  s" "inj_on f (ball a r)"
proof -
  have bl: "bounded_linear (blinfun_apply g')"
    by (auto simp: blinfun.bounded_linear_right)
  from g' have g': "blinfun_apply g'  blinfun_apply (f' a) = id"
    by transfer (simp add: id_def)
  from has_derivative_locally_injective[OF a  s open s bl g' f' isCont_blinfunD[OF c]]
  obtain r where "0 < r" "ball a r  s" "inj_on f (ball a r)"
    by auto
  then show ?thesis ..
qed

lift_definition embed1_blinfun::"'a::real_normed_vector L ('a*'b::real_normed_vector)" is "λx. (x, 0)"
  by standard (auto intro!: exI[where x=1])
lemma blinfun_apply_embed1_blinfun[simp]: "blinfun_apply embed1_blinfun x = (x, 0)"
  by transfer simp

lift_definition embed2_blinfun::"'a::real_normed_vector L ('b::real_normed_vector*'a)" is "λx. (0, x)"
  by standard (auto intro!: exI[where x=1])
lemma blinfun_apply_embed2_blinfun[simp]: "blinfun_apply embed2_blinfun x = (0, x)"
  by transfer simp

lemma blinfun_inverseD: "f oL f' = 1L  f (f' x) = x"
  apply transfer
  unfolding o_def
  by meson

lemmas continuous_on_open_vimageI = continuous_on_open_vimage[THEN iffD1, rule_format]
lemmas continuous_on_closed_vimageI = continuous_on_closed_vimage[THEN iffD1, rule_format]

lemma ball_times_subset: "ball a (c/2) × ball b (c/2)  ball (a, b) c"
proof -
  {
    fix a' b'
    have "sqrt ((dist a a')2 + (dist b b')2)  dist a a' + dist b b'"
      by (rule real_le_lsqrt) (auto simp: power2_eq_square algebra_simps)
    also assume "a'  ball a (c / 2)"
    then have "dist a a' < c / 2" by (simp add:)
    also assume "b'  ball b (c / 2)"
    then have "dist b b' < c / 2" by (simp add:)
    finally have "sqrt ((dist a a')2 + (dist b b')2) < c"
      by simp
  } thus ?thesis by (auto simp: dist_prod_def mem_cball)
qed

lemma linear_inverse_blinop_lemma:
  fixes w::"'a::{banach, perfect_space} blinop"
  assumes "norm w < 1"
  shows
    "summable (λn. (-1)^n *R w^n)" (is ?C)
    "(n. (-1)^n *R w^n) * (1 + w) = 1" (is ?I1)
    "(1 + w) * (n. (-1)^n *R w^n) = 1" (is ?I2)
    "norm ((n. (-1)^n *R w^n) - 1 + w)  (norm w)2/(1 - norm (w))" (is ?L)
proof -
  have "summable (λn. norm w ^ n)"
    apply (rule summable_geometric)
    using assms by auto
  then have "summable (λn. norm (w ^ n))"
    by (rule summable_comparison_test'[where N=0]) (auto intro!: norm_power_ineq)
  then show ?C
    by (rule summable_comparison_test'[where N=0]) (auto simp: norm_power )
  {
    fix N
    have 1: "(1 + w) * sum (λn. (-1)^n *R w^n) {..<N} = sum (λn. (-1)^n *R w^n) {..<N} * (1 + w)"
      by (auto simp: algebra_simps sum_distrib_left sum_distrib_right sum.distrib power_commutes)
    also have " = sum (λn. (-1)^n *R w^n - (-1)^Suc n *R w^ Suc n) {..<N}"
      by (auto simp: algebra_simps sum_distrib_left sum_distrib_right sum.distrib power_commutes)
    also have " = 1 - (-1)^N *R w^N"
      by (subst sum_lessThan_telescope') (auto simp: algebra_simps)
    finally have "(1 + w) * (n<N. (- 1) ^ n *R w ^ n) = 1 - (- 1) ^ N *R w ^ N" .
    note 1 this
  } note nu = this
  show "?I1"
    apply (subst suminf_mult2, fact)
    apply (subst suminf_eq_lim)
    apply (subst sum_distrib_right[symmetric])
    apply (rule limI)
    apply (subst nu(1)[symmetric])
    apply (subst nu(2))
    apply (rule tendsto_eq_intros)
      apply (rule tendsto_intros)
     apply (rule tendsto_norm_zero_cancel)
     apply auto
    apply (rule Lim_transform_bound[where g="λi. norm w ^ i"])
     apply (rule eventuallyI)
    apply simp apply (rule norm_power_ineq)
    apply (auto intro!: LIMSEQ_power_zero assms)
    done
  show "?I2"
    apply (subst suminf_mult[symmetric], fact)
    apply (subst suminf_eq_lim)
    apply (subst sum_distrib_left[symmetric])
    apply (rule limI)
    apply (subst nu(2))
    apply (rule tendsto_eq_intros)
      apply (rule tendsto_intros)
     apply (rule tendsto_norm_zero_cancel)
     apply auto
    apply (rule Lim_transform_bound[where g="λi. norm w ^ i"])
     apply (rule eventuallyI)
    apply simp apply (rule norm_power_ineq)
    apply (auto intro!: LIMSEQ_power_zero assms)
    done
  have *: "(n. (- 1) ^ n *R w ^ n) - 1 + w = (w2 * (n. (- 1) ^ n *R w ^ n))"
    apply (subst suminf_split_initial_segment[where k=2], fact)
    apply (subst suminf_mult[symmetric], fact)
    by (auto simp: power2_eq_square algebra_simps eval_nat_numeral)
  also have "norm   (norm w)2 / (1 - norm w)"
  proof -
    have §: "norm (n. (- 1) ^ n *R w ^ n)  1 / (1 - norm w)"
      apply (rule order_trans[OF summable_norm])
       apply auto
       apply fact
      apply (rule order_trans[OF suminf_le])
         apply (rule norm_power_ineq)
        apply fact
       apply fact
      by (auto simp: suminf_geometric assms)
    show ?thesis
      apply (rule order_trans[OF norm_mult_ineq])
      apply (subst divide_inverse)
      apply (rule mult_mono)
         apply (auto simp: norm_power_ineq inverse_eq_divide assms §)
      done
  qed
  finally show ?L .
qed

lemma linear_inverse_blinfun_lemma:
  fixes w::"'a L 'a::{banach, perfect_space}"
  assumes "norm w < 1"
  obtains I where
    "I oL (1L + w) = 1L" "(1L + w) oL I = 1L"
    "norm (I - 1L + w)  (norm w)2/(1 - norm (w))"
proof -
  define v::"'a blinop" where "v = Blinop w"
  have "norm v = norm w"
    unfolding v_def
    apply transfer
    by (simp add: bounded_linear_Blinfun_apply norm_blinfun.rep_eq)
  with assms have "norm v < 1" by simp
  from linear_inverse_blinop_lemma[OF this]
  have v: "(n. (- 1) ^ n *R v ^ n) * (1 + v) = 1"
    "(1 + v) * (n. (- 1) ^ n *R v ^ n) = 1"
    "norm ((n. (- 1) ^ n *R v ^ n) - 1 + v)  (norm v)2 / (1 - norm v)"
    by auto
  define J::"'a blinop" where "J = (n. (- 1) ^ n *R v ^ n)"
  define I::"'a L 'a" where "I = Blinfun J"
  have "Blinfun (blinop_apply J) - 1L + w = Rep_blinop (J - 1 + Blinop (blinfun_apply w))"
    by transfer' (auto simp: blinfun_apply_inverse)
  then have ne: "norm (Blinfun (blinop_apply J) - 1L + w) =
    norm (J - 1 + Blinop (blinfun_apply w))"
    by (auto simp: norm_blinfun_def norm_blinop_def)
  from v have
    "I oL (1L + w) = 1L" "(1L + w) oL I = 1L"
    "norm (I - 1L + w)  (norm w)2/(1 - norm (w))"
      apply (auto simp: I_def J_def[symmetric])
    unfolding v_def
      apply (auto simp: blinop.bounded_linear_right bounded_linear_Blinfun_apply
        intro!: blinfun_eqI)
    subgoal by transfer
       (auto simp: blinfun_ext blinfun.bilinear_simps bounded_linear_Blinfun_apply)
    subgoal
      by transfer (auto simp: Transfer.Rel_def
          blinfun_ext blinfun.bilinear_simps bounded_linear_Blinfun_apply)
    subgoal
      apply (auto simp: ne)
      apply transfer
      by (auto simp: norm_blinfun_def bounded_linear_Blinfun_apply)
    done
  then show ?thesis ..
qed

definition "invertibles_blinfun = {w. wi. w oL wi = 1L  wi oL w = 1L}"

lemma blinfun_inverse_open:― ‹8.3.2 in Dieudonne, TODO: add continuity and derivative›
  shows "open (invertibles_blinfun::
      ('a::{banach, perfect_space} L 'b::banach) set)"
proof (rule openI)
  fix u0::"'a L 'b"
  assume "u0  invertibles_blinfun"
  then obtain u0i where u0i: "u0 oL u0i = 1L" "u0i oL u0 = 1L"
    by (auto simp: invertibles_blinfun_def)
  then have [simp]: "u0i  0"
    apply (auto)
    by (metis one_blinop.abs_eq zero_blinop.abs_eq zero_neq_one)
  let ?e = "inverse (norm u0i)"
  show "e>0. ball u0 e  invertibles_blinfun"
    apply (clarsimp intro!: exI[where x = ?e] simp: invertibles_blinfun_def)
    subgoal premises prems for u0s
    proof -
      define s where "s = u0s - u0"
      have u0s: "u0s = u0 + s"
        by (auto simp: s_def)
      have "norm (u0i oL s) < 1"
        using prems by (auto simp: dist_norm u0s
        divide_simps ac_simps intro!: le_less_trans[OF norm_blinfun_compose])
      from linear_inverse_blinfun_lemma[OF this]
      obtain I where I:
        "I oL 1L + (u0i oL s) = 1L"
        "1L + (u0i oL s) oL I = 1L"
        "norm (I - 1L + (u0i oL s))  (norm (u0i oL s))2 / (1 - norm (u0i oL s))"
        by auto
      have u0s_eq: "u0s = u0 oL (1L + (u0i oL s))"
        using u0i
        by (auto simp: s_def blinfun.bilinear_simps blinfun_ext)
      show ?thesis
        apply (rule exI[where x="I oL u0i"])
        using I u0i
        apply (auto simp: u0s_eq)
        by (auto simp:  algebra_simps blinfun_ext blinfun.bilinear_simps)
    qed
    done
  qed

lemma blinfun_compose_assoc[ac_simps]: "a oL b oL c = a oL (b oL c)"
  by (auto intro!: blinfun_eqI)

text ‹TODO: move @{thm norm_minus_cancel} to class!›
lemma (in real_normed_vector) norm_minus_cancel [simp]: "norm (- x) = norm x"
proof -
  have scaleR_minus_left: "- a *R x = - (a *R x)" for a x
  proof -
    have "x1 x2. (x2::real) + x1 = x1 + x2"
      by auto
    then have f1: "r ra a. (ra + r) *R (a::'a) = r *R a + ra *R a"
      using local.scaleR_add_left by presburger
    have f2: "a + a = 2 * a"
      by force
    have f3: "2 * a + - 1 * a = a"
      by auto
    have "- a = - 1 * a"
      by auto
    then show ?thesis
      using f3 f2 f1 by (metis local.add_minus_cancel local.add_right_imp_eq)
  qed
  have "norm (- x) = norm (scaleR (- 1) x)"
    by (simp only: scaleR_minus_left scaleR_one)
  also have " = ¦- 1¦ * norm x"
    by (rule norm_scaleR)
  finally show ?thesis by simp
qed

text ‹TODO: move @{thm norm_minus_commute} to class!›
lemma (in real_normed_vector) norm_minus_commute: "norm (a - b) = norm (b - a)"
proof -
  have "norm (- (b - a)) = norm (b - a)"
    by (rule norm_minus_cancel)
  then show ?thesis by simp
qed

instance euclidean_space  banach
  by standard

lemma blinfun_apply_Pair_split:
  "blinfun_apply g (a, b) = blinfun_apply g (a, 0) + blinfun_apply g (0, b)"
  unfolding blinfun.bilinear_simps[symmetric] by simp

lemma blinfun_apply_Pair_add2: "blinfun_apply f (0, a + b) = blinfun_apply f (0, a) + blinfun_apply f (0, b)"
  unfolding blinfun.bilinear_simps[symmetric] by simp

lemma blinfun_apply_Pair_add1: "blinfun_apply f (a + b, 0) = blinfun_apply f (a, 0) + blinfun_apply f (b, 0)"
  unfolding blinfun.bilinear_simps[symmetric] by simp

lemma blinfun_apply_Pair_minus2: "blinfun_apply f (0, a - b) = blinfun_apply f (0, a) - blinfun_apply f (0, b)"
  unfolding blinfun.bilinear_simps[symmetric] by simp

lemma blinfun_apply_Pair_minus1: "blinfun_apply f (a - b, 0) = blinfun_apply f (a, 0) - blinfun_apply f (b, 0)"
  unfolding blinfun.bilinear_simps[symmetric] by simp

lemma implicit_function_theorem:
  fixes f::"'a::euclidean_space * 'b::euclidean_space  'c::euclidean_space"― ‹TODO: generalize?!›
  assumes [derivative_intros]: "x. x  S  (f has_derivative blinfun_apply (f' x)) (at x)"
  assumes S: "(x, y)  S" "open S"
  assumes "DIM('c)  DIM('b)"
  assumes f'C: "isCont f' (x, y)"
  assumes "f (x, y) = 0"
  assumes T2: "T oL (f' (x, y) oL embed2_blinfun) = 1L"
  assumes T1: "(f' (x, y) oL embed2_blinfun) oL T = 1L"― ‹TODO: reduce?!›
  obtains u e r
  where "f (x, u x) = 0" "u x = y"
    "s. s  cball x e  f (s, u s) = 0"
    "continuous_on (cball x e) u"
    "(λt. (t, u t)) ` cball x e  S"
    "e > 0"
    "(u has_derivative - T oL f' (x, y) oL embed1_blinfun) (at x)"

    "r > 0"
    "U v s. v x = y  (s. s  U  f (s, v s) = 0)  U  cball x e 
      continuous_on U v  s  U  (s, v s)  ball (x, y) r  u s = v s"
proof -
  define H where "H  λ(x, y). (x, f (x, y))"
  define H' where "H'  λx. (embed1_blinfun oL fst_blinfun) + (embed2_blinfun oL (f' x))"
  have f'_inv: "f' (x, y) oL embed2_blinfun  invertibles_blinfun"
    using T1 T2 by (auto simp: invertibles_blinfun_def ac_simps intro!: exI[where x=T])
  from openE[OF blinfun_inverse_open this]
  obtain d0 where e0: "0 < d0"
    "ball (f' (x, y) oL embed2_blinfun) d0  invertibles_blinfun"
    by auto
  have "isCont (λs. f' s oL embed2_blinfun) (x, y)"
    by (auto intro!: continuous_intros f'C)
  from this[unfolded isCont_def, THEN tendstoD, OF 0 < d0]
  have "F s in at (x, y). f' s oL embed2_blinfun  invertibles_blinfun"
    apply eventually_elim
    using e0 by (auto simp: subset_iff dist_commute)
  then obtain e0 where "e0 > 0"
      "xa  (x, y)  dist xa (x, y) < e0 
        f' xa oL embed2_blinfun  invertibles_blinfun" for xa
    unfolding eventually_at
    by auto
  then have e0: "e0 > 0"
    "dist xa (x, y) < e0  f' xa oL embed2_blinfun  invertibles_blinfun" for xa
    apply -
    subgoal by simp
    using f'_inv
    apply (cases "xa = (x, y)")
    by auto

  have H': "x  S  (H has_derivative H' x) (at x)" for x
    unfolding H_def  H'_def
    by (auto intro!: derivative_eq_intros ext simp: blinfun.bilinear_simps)
  have cH': "isCont H' (x, y)"
    unfolding H'_def
    by (auto intro!: continuous_intros assms)
  have linear_H': "s. s  S  linear (H' s)"
    using H' assms(2) has_derivative_linear by blast
  have *: "blinfun_apply T (blinfun_apply (f' (x, y)) (0, b)) = b" for b
    using blinfun_inverseD[OF T2, of b]
    by simp
  have "inj (f' (x, y) oL embed2_blinfun)"
    by (metis (no_types, lifting) "*" blinfun_apply_blinfun_compose embed2_blinfun.rep_eq injI)
  then have [simp]: "blinfun_apply (f' (x, y)) (0, b) = 0  b = 0" for b
    apply (subst (asm) linear_injective_0)
    subgoal
      apply (rule bounded_linear.linear)
      apply (rule blinfun.bounded_linear_right)
      done
    subgoal by simp
    done
  have "inj (H' (x, y))"
    apply (subst linear_injective_0)
     apply (rule linear_H')
     apply fact
    apply (auto simp: H'_def blinfun.bilinear_simps zero_prod_def)
    done
  define Hi where "Hi = (embed1_blinfun oL fst_blinfun) + ((embed2_blinfun oL T oL (snd_blinfun - (f' (x, y) oL embed1_blinfun oL fst_blinfun))))"
  have Hi': "(λu. snd (blinfun_apply Hi (u, 0))) = - T oL f' (x, y) oL embed1_blinfun"
    by (auto simp: Hi_def blinfun.bilinear_simps)
  have Hi: "Hi oL H' (x, y) = 1L"
    apply (auto simp: H'_def fun_eq_iff blinfun.bilinear_simps Hi_def
        intro!: ext blinfun_eqI)
    apply (subst blinfun_apply_Pair_split)
    by (auto simp: * blinfun.bilinear_simps)
  from has_derivative_locally_injective_blinfun[OF S this H' cH']
  obtain r0 where r0: "0 < r0" "ball (x, y) r0  S" and inj: "inj_on H (ball (x, y) r0)"
    by auto
  define r where "r = min r0 e0"
  have r: "0 < r" "ball (x, y) r  S" and inj: "inj_on H (ball (x, y) r)"
    and r_inv: "s. s  ball (x, y) r  f' s oL embed2_blinfun  invertibles_blinfun"
    subgoal using e0 r0 by (auto simp: r_def)
    subgoal using e0 r0 by (auto simp: r_def)
    subgoal using inj apply (rule inj_on_subset)
      using e0 r0 by (auto simp: r_def)
    subgoal for s
      using e0 r0 by (auto simp: r_def dist_commute)
    done
  obtain i::'a where "i  Basis"
    using nonempty_Basis by blast
  define undef where "undef  (x, y) + r *R (i, 0)"― ‹really??›
  have ud: "¬ dist (x, y) undef < r"
    using r > 0 i  Basis by (auto simp: undef_def dist_norm)
  define G where "G  the_inv_into (ball (x, y) r) H"
  {
    fix u v
    assume [simp]: "(u, v)  H ` ball (x, y) r"
    note [simp] = inj
    have "(u, v) = H (G (u, v))"
      unfolding G_def
      by (subst f_the_inv_into_f[where f=H]) auto
    moreover have " = H (G (u, v))"
      by (auto simp: G_def)
    moreover have " = (fst (G (u, v)), f (G (u, v)))"
      by (auto simp: H_def split_beta')
    ultimately have "u = fst (G (u, v))" "v = f (G (u, v))" by simp_all
    then have "f (u, snd (G(u, v))) = v" "u = fst (G (u, v))"
      by (metis prod.collapse)+
  } note uvs = this
  note uv = uvs(1)
  moreover
  have "f (x, snd (G (x, 0))) = 0"
    apply (rule uv)
    by (metis (mono_tags, lifting) H_def assms(6) case_prod_beta' centre_in_ball fst_conv image_iff r(1) snd_conv)
  moreover
  have cH: "continuous_on S H"
    apply (rule has_derivative_continuous_on)
    apply (subst at_within_open)
      apply (auto intro!: H' assms)
    done
  have inj2: "inj_on H (ball (x, y) (r / 2))"
    apply (rule inj_on_subset, rule inj)
    using r by auto
  have oH: "open (H ` ball (x, y) (r/2))"
    apply (rule invariance_of_domain_gen)
       apply (auto simp: assms inj)
    apply (rule continuous_on_subset)
      apply fact
    using r
     apply auto
    using inj2 apply simp
    done
  have "(x, f (x, y))  H ` ball (x, y) (r/2)"
    using r > 0 by (auto simp: H_def)
  from open_contains_cball[THEN iffD1, OF oH, rule_format, OF this]
  obtain e' where e': "e' > 0" "cball (x, f (x, y)) e'  H ` ball (x, y) (r/2)"
    by auto

  have inv_subset: "the_inv_into (ball (x, y) r) H a = the_inv_into R H a"
    if "a  H ` R" "R  (ball (x, y) r)"
    for a R
    apply (rule the_inv_into_f_eq[OF inj])
     apply (rule f_the_inv_into_f)
      apply (rule inj_on_subset[OF inj])
      apply fact
     apply fact
    apply (rule the_inv_into_into)
      apply (rule inj_on_subset[OF inj])
      apply fact
     apply fact
    apply (rule order_trans)
     apply fact
    using r apply auto
    done
  have GH: "G (H z) = z" if "dist (x, y) z < r" for z
    by (auto simp: G_def the_inv_into_f_f inj that)
  define e where "e = min (e' / 2) e0"
  define r2 where "r2 = r / 2"
  have r2: "r2 > 0" "r2 < r"
    using r > 0 by (auto simp: r2_def)
  have "e > 0" using e' e0 by (auto simp: e_def)
  from cball_times_subset[of "x" e' "f (x, y)"] e'
  have "cball x e × cball (f (x, y)) e  H ` ball (x, y) (r/2)"
    by (force simp: e_def)
  then have e_r_subset: "z  cball x e  (z, 0)  H ` ball (x, y) (r/2)" for z
    using 0 < e assms(6)
    by (auto simp: H_def subset_iff)
  have u0: "(u, 0)  H ` ball (x, y) r" if "u  cball x e" for u
    apply (rule rev_subsetD)
     apply (rule e_r_subset)
     apply fact
    unfolding r2_def using r2 by auto
  have G_r: "G (u, 0)  ball (x, y) r" if "u  cball x e" for u
    unfolding G_def
    apply (rule the_inv_into_into)
      apply fact
     apply (auto)
    apply (rule u0, fact)
    done
  note e_r_subset
  ultimately have G2:
    "f (x, snd (G (x, 0))) = 0" "snd (G (x, 0)) = y"
    "u. u  cball x e  f (u, snd (G (u, 0))) = 0"
    "continuous_on (cball x e) (λu. snd (G (u, 0)))"
    "(λt. (t, snd (G (t, 0)))) ` cball x e  S"
    "e > 0"
    "((λu. snd (G (u, 0))) has_derivative (λu. snd (Hi (u, 0)))) (at x)"
       apply (auto simp: G_def split_beta'
        intro!: continuous_intros continuous_on_compose2[OF cH])
    subgoal premises prems
    proof -
      have "the_inv_into (ball (x, y) r) H (x, 0) = (x, y)"
        apply (rule the_inv_into_f_eq)
          apply fact
         by (auto simp: H_def assms r > 0)
       then show ?thesis
         by auto
    qed
    using r2(2) r2_def apply fastforce
    apply (subst continuous_on_cong[OF refl])
     apply (rule inv_subset[where R="cball (x, y) r2"])
    subgoal
      using r2
      apply auto
      using r2_def by force
    subgoal using r2 by (force simp:)
    subgoal
      apply (rule continuous_on_compose2[OF continuous_on_inv_into])
      using r(2) r2(2)
        apply (auto simp: r2_def[symmetric]
          intro!: continuous_on_compose2[OF cH] continuous_intros)
       apply (rule inj_on_subset)
        apply (rule inj)
      using r(2) r2(2) apply force
      apply force
      done
    subgoal premises prems for u
    proof -
      from prems have u: "u  cball x e" by auto
      note G_r[OF u]
      also have "ball (x, y) r  S"
        using r by simp
      finally have "(G (u, 0))  S" .
      then show ?thesis
        unfolding G_def[symmetric]
        using uvs(2)[OF u0, OF u]
        by (metis prod.collapse)
    qed
    subgoal using e > 0 by simp
    subgoal premises prems
    proof -
      have "(x, y)  cball (x, y) r2"
        using r2
        by auto
      moreover
      have "H (x, y)  interior (H ` cball (x, y) r2)"
        apply (rule interiorI[OF oH])
        using r2 by (auto simp: r2_def)
      moreover
      have "cball (x, y) r2  S"
        using r r2 by auto
      moreover have "z. z  cball (x, y) r2  G (H z) = z"
        using r2 by (auto intro!: GH)
      ultimately have "(G has_derivative Hi) (at (H (x, y)))"
      proof (rule has_derivative_inverse[where g = G and f = H,
              OF compact_cball _ _ continuous_on_subset[OF cH] _ H' _ _])
        show "blinfun_apply Hi  blinfun_apply (H' (x, y)) = id"
          using Hi by transfer auto
      qed (use S blinfun.bounded_linear_right in auto)
      then have g': "(G has_derivative Hi) (at (x, 0))"
        by (auto simp: H_def assms)
      show ?thesis
        unfolding G_def[symmetric] H_def[symmetric]
        apply (auto intro!: derivative_eq_intros)
         apply (rule has_derivative_compose[where g=G and f="λx. (x, 0)"])
         apply (auto intro!: g' derivative_eq_intros)
        done
    qed
    done
  moreover
  note r > 0
  moreover
  define u where "u  λx. snd (G (x, 0))"
  have local_unique: "u s = v s"
    if solves: "(s. s  U  f (s, v s) = 0)"
    and i: "v x = y"
    and v: "continuous_on U v"
    and s: "s  U"
    and s': "(s, v s)  ball (x, y) r"
    and U: "U  cball x e"
    for U v s
  proof -
    have H_eq: "H (s, v s) = H (s, u s)"
      apply (auto simp: H_def solves[OF s])
      unfolding u_def
      apply (rule G2)
      apply (rule subsetD; fact)
      done
    have "(s, snd (G (s, 0))) = (G (s, 0))"
      using GH H_def s s' solves by fastforce
    also have "  ball (x, y) r"
      unfolding G_def
      apply (rule the_inv_into_into)
        apply fact
       apply (rule u0)
       apply (rule subsetD; fact)
      apply (rule order_refl)
      done
    finally have "(s, u s)  ball (x, y) r" unfolding u_def .
    from inj_onD[OF inj H_eq s' this]
    show "u s = v s"
      by auto
  qed
  ultimately show ?thesis
    unfolding u_def Hi' ..
qed

lemma implicit_function_theorem_unique:
  fixes f::"'a::euclidean_space * 'b::euclidean_space  'c::euclidean_space"― ‹TODO: generalize?!›
  assumes f'[derivative_intros]: "x. x  S  (f has_derivative blinfun_apply (f' x)) (at x)"
  assumes S: "(x, y)  S" "open S"
  assumes D: "DIM('c)  DIM('b)"
  assumes f'C: "continuous_on S f'"
  assumes z: "f (x, y) = 0"
  assumes T2: "T oL (f' (x, y) oL embed2_blinfun) = 1L"
  assumes T1: "(f' (x, y) oL embed2_blinfun) oL T = 1L"― ‹TODO: reduce?!›
  obtains u e
  where "f (x, u x) = 0" "u x = y"
    "s. s  cball x e  f (s, u s) = 0"
    "continuous_on (cball x e) u"
    "(λt. (t, u t)) ` cball x e  S"
    "e > 0"
    "(u has_derivative (- T oL f' (x, y) oL embed1_blinfun)) (at x)"
    "s. s  cball x e  f' (s, u s) oL embed2_blinfun  invertibles_blinfun"
    "U v s. (s. s  U  f (s, v s) = 0) 
      u x = v x 
      continuous_on U v  s  U  x  U  U  cball x e  connected U  open U  u s = v s"
proof -
  from T1 T2 have f'I: "f' (x, y) oL embed2_blinfun  invertibles_blinfun"
    by (auto simp: invertibles_blinfun_def)
  from assms have f'Cg: "s  S  isCont f' s" for s
    by (auto simp: continuous_on_eq_continuous_at[OF open S])
  then have f'C: "isCont f' (x, y)" by (auto simp: S)
  obtain u e1 r
    where u: "f (x, u x) = 0" "u x = y"
      "s. s  cball x e1  f (s, u s) = 0"
      "continuous_on (cball x e1) u"
      "(λt. (t, u t)) ` cball x e1  S"
      "e1 > 0"
    "(u has_derivative (- T oL f' (x, y) oL embed1_blinfun)) (at x)"
    and unique_u: "r > 0"
      "(v s U. v x = y 
        (s. s  U  f (s, v s) = 0) 
        continuous_on U v  s  U  U  cball x e1  (s, v s)  ball (x, y) r  u s = v s)"
    by (rule implicit_function_theorem[OF f' S D f'C z T2 T1]; blast)

  from openE[OF blinfun_inverse_open f'I] obtain d where d:
    "0 < d" "ball (f' (x, y) oL embed2_blinfun) d  invertibles_blinfun"
    by auto
  note [continuous_intros] = continuous_at_compose[OF _ f'Cg, unfolded o_def]
  from continuous_on _ u
  have "continuous_on (ball x e1) u" by (rule continuous_on_subset) auto
  then have "s. s  ball x e1  isCont u s"
    unfolding continuous_on_eq_continuous_at[OF open_ball] by auto
  note [continuous_intros] = continuous_at_compose[OF _ this, unfolded o_def]
  from assms have f'Ce: "isCont (λs. f' (s, u s) oL embed2_blinfun) x"
    by (auto simp: u intro!: continuous_intros)
  from f'Ce[unfolded isCont_def, THEN tendstoD, OF 0 < d] d
  obtain e0 where "e0 > 0" "s. s  x  s  ball x e0 
      (f' (s, u s) oL embed2_blinfun)  invertibles_blinfun"
    by (auto simp: eventually_at dist_commute subset_iff u)
  then have e0: "s  ball x e0  (f' (s, u s) oL embed2_blinfun)  invertibles_blinfun" for s
    by (cases "s = x") (auto simp: f'I 0 < d u)


  define e where "e = min (e0/2) (e1/2)"
  have e: "f (x, u x) = 0"
      "u x = y"
      "s. s  cball x e  f (s, u s) = 0"
      "continuous_on (cball x e) u"
      "(λt. (t, u t)) ` cball x e  S"
      "e > 0"
      "(u has_derivative (- T oL f' (x, y) oL embed1_blinfun)) (at x)"
      "s. s  cball x e  f' (s, u s) oL embed2_blinfun  invertibles_blinfun"
    using e0 u e0 > 0 by (auto simp: e_def intro: continuous_on_subset)

  from u(4) have "continuous_on (ball x e1) u"
    apply (rule continuous_on_subset)
    using e1 > 0
    by (auto simp: e_def)
  then have "s. s  cball x e  isCont u s"
    using e0 > 0 e1 > 0
    unfolding continuous_on_eq_continuous_at[OF open_ball] by (auto simp: e_def Ball_def dist_commute)
  note [continuous_intros] = continuous_at_compose[OF _ this, unfolded o_def]

  have "u s = v s"
    if solves: "(s. s  U  f (s, v s) = 0)"
    and i: "u x = v x"
    and v: "continuous_on U v"
    and s: "s  U" and U: "x  U" "U  cball x e" "connected U" "open U"
    for U v s
  proof -
    define M where "M = {s  U. u s = v s}"
    have "x  M" using i U by (auto simp: M_def)
    moreover
    have "continuous_on U (λs. u s - v s)"
      by (auto intro!: continuous_intros v continuous_on_subset[OF e(4) U(2)])
    from continuous_closedin_preimage[OF this closed_singleton[where a=0]]
    have "closedin (top_of_set U) M"
      by (auto simp: M_def vimage_def Collect_conj_eq)
    moreover
    have "s. s  U   isCont v s"
      using v
      unfolding continuous_on_eq_continuous_at[OF open U] by auto
    note [continuous_intros] = continuous_at_compose[OF _ this, unfolded o_def]
    {
      fix a assume "a  M"
      then have aU: "a  U" and u_v: "u a = v a"
        by (auto simp: M_def)
      then have a_ball: "a  cball x e" and a_dist: "dist x a  e" using U by auto
      then have a_S: "(a, u a)  S"
        using e by auto
      have fa_z: "f (a, u a) = 0"
        using a  cball x e by (auto intro!: e)
      from e(8)[OF a  cball _ _]
      obtain Ta where Ta: "Ta oL (f' (a, u a) oL embed2_blinfun) = 1L" "f' (a, u a) oL embed2_blinfun oL Ta = 1L"
        by (auto simp: invertibles_blinfun_def ac_simps)
      obtain u' e' r'
        where "r' > 0" "e' > 0"
        and u': "v s U. v a = u a 
             (s. s  U  f (s, v s) = 0) 
             continuous_on U v  s  U  U  cball a e'  (s, v s)  ball (a, u a) r'  u' s = v s"
        by (rule implicit_function_theorem[OF f' a_S open S D f'Cg[OF a_S] fa_z Ta]; blast)
      from openE[OF open U aU] obtain dU where dU: "dU > 0" "s. s  ball a dU  s  U"
        by (auto simp: dist_commute subset_iff)

      have v_tendsto: "((λs. (s, v s))  (a, u a)) (at a)"
        unfolding u_v
        by (subst continuous_at[symmetric]) (auto intro!: continuous_intros aU)
      from tendstoD[OF v_tendsto 0 < r', unfolded eventually_at]
      obtain dv where "dv > 0" "s  a  dist s a < dv  (s, v s)  ball (a, u a) r'" for s
        by (auto simp: dist_commute)
      then have dv: "dist s a < dv  (s, v s)  ball (a, u a) r'" for s
        by (cases "s = a") (auto simp: u_v 0 < r')

      have v_tendsto: "((λs. (s, u s))  (a, u a)) (at a)"
        using a_dist
        by (subst continuous_at[symmetric]) (auto intro!: continuous_intros)
      from tendstoD[OF v_tendsto 0 < r', unfolded eventually_at]
      obtain du where "du > 0" "s  a  dist s a < du  (s, u s)  ball (a, u a) r'" for s
        by (auto simp: dist_commute)
      then have du: "dist s a < du  (s, u s)  ball (a, u a) r'" for s
        by (cases "s = a") (auto simp: u_v 0 < r')
      {
        fix s assume s: "s  ball a (Min {dU, e', dv, du})"
        let ?U = "ball a (Min {dU, e', dv, du})"
        have balls: "ball a (Min {dU, e', dv, du})  cball a e'" by auto
        have dsadv: "dist s a < dv"
          using s by (auto simp: dist_commute)
        have dsadu: "dist s a < du"
          using s by (auto simp: dist_commute)
        have U_U: "s. s  ball a (Min {dU, e', dv, du})  s  U"
          using dU by auto
        have U_e: "s. s  ball a (Min {dU, e', dv, du})  s  cball x e"
          using dU U by (auto simp: dist_commute subset_iff)
        have cv: "continuous_on ?U v"
          using v
          apply (rule continuous_on_subset)
          using dU
          by auto
        have cu: "continuous_on ?U u"
          using e(4)
          apply (rule continuous_on_subset)
          using dU U(2)
          by auto
        from u'[where v=v, OF u_v[symmetric] solves[OF U_U] cv s balls dv[OF dsadv]]
          u'[where v=u, OF refl              e(3)[OF U_e]   cu s balls du[OF dsadu]]
        have "v s = u s" by auto
      } then have "dv>0. s  ball a dv. v s = u s"
        using 0 < dU 0 < e' 0 < dv 0 < du
        by (auto intro!: exI[where x="(Min {dU, e', dv, du})"])
    } note ex = this
    have "openin (top_of_set U) M"
      unfolding openin_contains_ball
      apply (rule conjI)
      subgoal using U by (auto simp: M_def)
      apply (auto simp:)
      apply (drule ex)
      apply auto
      subgoal for x d
        by (rule exI[where x=d]) (auto simp: M_def)
      done
    ultimately have "M = U"
      using connected U
      by (auto simp: connected_clopen)
    with s  U show ?thesis by (auto simp: M_def)
  qed
  from e this
  show ?thesis ..
qed

lemma uniform_limit_compose:
  assumes ul: "uniform_limit T f l F"
  assumes uc: "uniformly_continuous_on S s"
  assumes ev: "F x in F. f x ` T  S"
  assumes subs: "l ` T  S"
  shows  "uniform_limit T (λi x. s (f i x)) (λx. s (l x)) F"
proof (rule uniform_limitI)
  fix e::real assume "e > 0"
  from uniformly_continuous_onE[OF uc e > 0]
  obtain d where d: "0 < d" "t t'. t  S  t'  S  dist t' t < d  dist (s t') (s t) < e"
    by auto
  from uniform_limitD[OF ul 0 < d] have "F n in F. xT. dist (f n x) (l x) < d" .
  then show "F n in F. xT. dist (s (f n x)) (s (l x)) < e"
    using ev
    by eventually_elim (use d subs in force)
qed

lemma
  uniform_limit_in_open:
  fixes l::"'a::topological_space'b::heine_borel"
  assumes ul: "uniform_limit T f l (at x)"
  assumes cont: "continuous_on T l"
  assumes compact: "compact T" and T_ne: "T  {}"
  assumes B: "open B"
  assumes mem: "l ` T  B"
  shows "F y in at x. t  T. f y t  B"
proof -
  have l_ne: "l ` T  {}" using T_ne by auto
  have "compact (l ` T)"
    by (auto intro!: compact_continuous_image cont compact)
  from compact_in_open_separated[OF l_ne this B mem]
  obtain e where "e > 0" "{x. infdist x (l ` T)  e}  B"
    by auto
  from uniform_limitD[OF ul 0 < e]
  have "F n in at x. xT. dist (f n x) (l x) < e" .
  then show ?thesis
  proof eventually_elim
    case (elim y)
    show ?case
    proof safe
      fix t assume "t  T"
      have "infdist (f y t) (l ` T)  dist (f y t) (l t)"
        by (rule infdist_le) (use t  T in auto)
      also have " < e" using elim t  T by auto
      finally have "infdist (f y t) (l ` T)  e" by simp
      then have "(f y t)  {x. infdist x (l ` T)  e}"
        by (auto )
      also note   B
      finally show "f y t  B" .
    qed
  qed
qed

lemma
  order_uniform_limitD1:
  fixes l::"'a::topological_spacereal"― ‹TODO: generalize?!›
  assumes ul: "uniform_limit T f l (at x)"
  assumes cont: "continuous_on T l"
  assumes compact: "compact T"
  assumes less: "t. t  T  l t < b"
  shows "F y in at x. t  T. f y t < b"
proof cases
  assume ne: "T  {}"
  from compact_attains_sup[OF compact_continuous_image[OF cont compact], unfolded image_is_empty, OF ne]
  obtain tmax where tmax: "tmax  T" "s. s  T  l s  l tmax"
    by auto
  have "b - l tmax > 0"
    using ne tmax less by auto
  from uniform_limitD[OF ul this]
  have "F n in at x. xT. dist (f n x) (l x) < b - l tmax"
    by auto
  then show ?thesis
    apply eventually_elim
    using tmax
    by (force simp: dist_real_def abs_real_def split: if_splits)
qed auto

lemma
  order_uniform_limitD2:
  fixes l::"'a::topological_spacereal"― ‹TODO: generalize?!›
  assumes ul: "uniform_limit T f l (at x)"
  assumes cont: "continuous_on T l"
  assumes compact: "compact T"
  assumes less: "t. t  T  l t > b"
  shows "F y in at x. t  T. f y t > b"
proof -
  have "F y in at x. tT. (- f) y t < - b"
    by (rule order_uniform_limitD1[of "- f" T "-l" x "- b"])
      (auto simp: assms fun_Compl_def intro!: uniform_limit_eq_intros continuous_intros)
  then show ?thesis by auto
qed

lemma continuous_on_avoid_cases:
  fixes l::"'b::topological_space  'a::linear_continuum_topology"― ‹TODO: generalize!›
  assumes cont: "continuous_on T l" and conn: "connected T"
  assumes avoid: "t. t  T  l t  b"
  obtains "t. t  T  l t < b" | "t. t  T  l t > b"
  apply atomize_elim
  using connected_continuous_image[OF cont conn] using avoid
  unfolding connected_iff_interval
  apply (auto simp: image_iff)
  using leI by blast

lemma
  order_uniform_limit_ne:
  fixes l::"'a::topological_spacereal"― ‹TODO: generalize?!›
  assumes ul: "uniform_limit T f l (at x)"
  assumes cont: "continuous_on T l"
  assumes compact: "compact T" and conn: "connected T"
  assumes ne: "t. t  T  l t  b"
  shows "F y in at x. t  T. f y t  b"
proof -
  from continuous_on_avoid_cases[OF cont conn ne]
  consider "(t. t  T  l t < b)" | "(t. t  T  l t > b)"
    by blast
  then show ?thesis
  proof cases
    case 1
    from order_uniform_limitD1[OF ul cont compact 1]
    have "F y in at x. tT. f y t < b" by simp
    then show ?thesis
      by eventually_elim auto
  next
    case 2
    from order_uniform_limitD2[OF ul cont compact 2]
    have "F y in at x. tT. f y t > b" by simp
    then show ?thesis
      by eventually_elim auto
  qed
qed

lemma open_cballE:
  assumes "open S" "xS"
  obtains e where "e>0" "cball x e  S"
  using assms unfolding open_contains_cball by auto

lemma pos_half_less: fixes x::real shows "x > 0  x / 2 < x"
  by auto

lemma closed_levelset: "closed {x. s x = (c::'a::t1_space)}" if "continuous_on UNIV s"
proof -
  have "{x. s x = c} = s -` {c}" by auto
  also have "closed "
    apply (rule closed_vimage)
     apply (rule closed_singleton)
    apply (rule that)
    done
  finally show ?thesis .
qed

lemma closed_levelset_within: "closed {x  S. s x = (c::'a::t1_space)}" if "continuous_on S s" "closed S"
proof -
  have "{x  S. s x = c} = s -` {c}  S" by auto
  also have "closed "
    apply (rule continuous_on_closed_vimageI)
      apply (rule that)
     apply (rule that)
    apply simp
    done
  finally show ?thesis .
qed

context c1_on_open_euclidean
begin

lemma open_existence_ivlE:
  assumes "t  existence_ivl0 x" "t  0"
  obtains e where "e > 0" "cball x e × {0 .. t + e}  Sigma X existence_ivl0"
proof -
  from assms have "(x, t)  Sigma X existence_ivl0"
    by auto
  from open_cballE[OF open_state_space this]
  obtain e0' where e0: "0 < e0'" "cball (x, t) e0'  Sigma X existence_ivl0"
    by auto
  define e0 where "e0 = (e0' / 2)"
  from cball_times_subset[of x e0' t] pos_half_less[OF 0 < e0'] half_gt_zero[OF 0 < e0'] e0
  have "cball x e0 × cball t e0  Sigma X existence_ivl0" "0 < e0" "e0 < e0'"
    unfolding e0_def by auto
  then have "e0 > 0" "cball x e0 × {0..t + e0}  Sigma X existence_ivl0"
    apply (auto simp: subset_iff dest!: spec[where x=t])
    subgoal for a b
      apply (rule in_existence_between_zeroI)
      apply (drule spec[where x=a])
       apply (drule spec[where x="t + e0"])
       apply (auto simp: dist_real_def closed_segment_eq_real_ivl)
      done
    done
  then show ?thesis ..
qed

lemmas [derivative_intros] = flow0_comp_has_derivative

lemma flow_isCont_state_space_comp[continuous_intros]:
  "t x  existence_ivl0 (s x)  isCont s x  isCont t x  isCont (λx. flow0 (s x) (t x)) x"
  using continuous_within_compose3[where g="λ(x, t). flow0 x t"
      and f="λx. (s x, t x)" and x = x and s = UNIV]
  flow_isCont_state_space
  by auto

lemma closed_plane[simp]: "closed {x. x