src/HOL/Multivariate_Analysis/Path_Connected.thy
changeset 60303 00c06f1315d0
parent 59557 ebd8ecacfba6
child 60420 884f54e01427
equal deleted inserted replaced
60302:6dcb8aa0966a 60303:00c06f1315d0
     1 (*  Title:      HOL/Multivariate_Analysis/Path_Connected.thy
     1 (*  Title:      HOL/Multivariate_Analysis/Path_Connected.thy
     2     Author:     Robert Himmelmann, TU Muenchen
     2     Author:     Robert Himmelmann, TU Muenchen, and LCP with material from HOL Light
     3 *)
     3 *)
     4 
     4 
     5 section {* Continuous paths and path-connected sets *}
     5 section {* Continuous paths and path-connected sets *}
     6 
     6 
     7 theory Path_Connected
     7 theory Path_Connected
     8 imports Convex_Euclidean_Space
     8 imports Convex_Euclidean_Space
     9 begin
     9 begin
    10 
    10 
    11 subsection {* Paths. *}
    11 (*FIXME move up?*)
       
    12 lemma image_add_atLeastAtMost [simp]:
       
    13   fixes d::"'a::linordered_idom" shows "(op + d ` {a..b}) = {a+d..b+d}"
       
    14   apply auto
       
    15   apply (rule_tac x="x-d" in rev_image_eqI, auto)
       
    16   done
       
    17 
       
    18 lemma image_diff_atLeastAtMost [simp]:
       
    19   fixes d::"'a::linordered_idom" shows "(op - d ` {a..b}) = {d-b..d-a}"
       
    20   apply auto
       
    21   apply (rule_tac x="d-x" in rev_image_eqI, auto)
       
    22   done
       
    23 
       
    24 lemma image_mult_atLeastAtMost [simp]:
       
    25   fixes d::"'a::linordered_field"
       
    26   assumes "d>0" shows "(op * d ` {a..b}) = {d*a..d*b}"
       
    27   using assms
       
    28   apply auto
       
    29   apply (rule_tac x="x/d" in rev_image_eqI)
       
    30   apply (auto simp: field_simps)
       
    31   done
       
    32 
       
    33 lemma image_affinity_interval:
       
    34   fixes c :: "'a::ordered_real_vector"
       
    35   shows "((\<lambda>x. m *\<^sub>R x + c) ` {a..b}) = (if {a..b}={} then {}
       
    36             else if 0 <= m then {m *\<^sub>R a + c .. m  *\<^sub>R b + c}
       
    37             else {m *\<^sub>R b + c .. m *\<^sub>R a + c})"
       
    38   apply (case_tac "m=0", force)
       
    39   apply (auto simp: scaleR_left_mono)
       
    40   apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI, auto simp: pos_le_divideR_eq le_diff_eq scaleR_left_mono_neg)
       
    41   apply (metis diff_le_eq inverse_inverse_eq order.not_eq_order_implies_strict pos_le_divideR_eq positive_imp_inverse_positive)
       
    42   apply (rule_tac x="inverse m *\<^sub>R (x-c)" in rev_image_eqI, auto simp: not_le neg_le_divideR_eq diff_le_eq)
       
    43   using le_diff_eq scaleR_le_cancel_left_neg 
       
    44   apply fastforce
       
    45   done
       
    46 
       
    47 lemma image_affinity_atLeastAtMost:
       
    48   fixes c :: "'a::linordered_field"
       
    49   shows "((\<lambda>x. m*x + c) ` {a..b}) = (if {a..b}={} then {}
       
    50             else if 0 \<le> m then {m*a + c .. m *b + c}
       
    51             else {m*b + c .. m*a + c})"
       
    52   apply (case_tac "m=0", auto)
       
    53   apply (rule_tac x="inverse m*(x-c)" in rev_image_eqI, auto simp: field_simps)
       
    54   apply (rule_tac x="inverse m*(x-c)" in rev_image_eqI, auto simp: field_simps)
       
    55   done
       
    56 
       
    57 lemma image_affinity_atLeastAtMost_diff:
       
    58   fixes c :: "'a::linordered_field"
       
    59   shows "((\<lambda>x. m*x - c) ` {a..b}) = (if {a..b}={} then {}
       
    60             else if 0 \<le> m then {m*a - c .. m*b - c}
       
    61             else {m*b - c .. m*a - c})"
       
    62   using image_affinity_atLeastAtMost [of m "-c" a b]
       
    63   by simp
       
    64 
       
    65 lemma image_affinity_atLeastAtMost_div_diff:
       
    66   fixes c :: "'a::linordered_field"
       
    67   shows "((\<lambda>x. x/m - c) ` {a..b}) = (if {a..b}={} then {}
       
    68             else if 0 \<le> m then {a/m - c .. b/m - c}
       
    69             else {b/m - c .. a/m - c})"
       
    70   using image_affinity_atLeastAtMost_diff [of "inverse m" c a b]
       
    71   by (simp add: field_class.field_divide_inverse algebra_simps)
       
    72 
       
    73 lemma closed_segment_real_eq:
       
    74   fixes u::real shows "closed_segment u v = (\<lambda>x. (v - u) * x + u) ` {0..1}"
       
    75   by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl)
       
    76 
       
    77 subsection {* Paths and Arcs *}
    12 
    78 
    13 definition path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
    79 definition path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
    14   where "path g \<longleftrightarrow> continuous_on {0..1} g"
    80   where "path g \<longleftrightarrow> continuous_on {0..1} g"
    15 
    81 
    16 definition pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
    82 definition pathstart :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> 'a"
    29     (infixr "+++" 75)
    95     (infixr "+++" 75)
    30   where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
    96   where "g1 +++ g2 = (\<lambda>x. if x \<le> 1/2 then g1 (2 * x) else g2 (2 * x - 1))"
    31 
    97 
    32 definition simple_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
    98 definition simple_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
    33   where "simple_path g \<longleftrightarrow>
    99   where "simple_path g \<longleftrightarrow>
    34     (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
   100      path g \<and> (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
    35 
   101 
    36 definition injective_path :: "(real \<Rightarrow> 'a::topological_space) \<Rightarrow> bool"
   102 definition arc :: "(real \<Rightarrow> 'a :: topological_space) \<Rightarrow> bool"
    37   where "injective_path g \<longleftrightarrow> (\<forall>x\<in>{0..1}. \<forall>y\<in>{0..1}. g x = g y \<longrightarrow> x = y)"
   103   where "arc g \<longleftrightarrow> path g \<and> inj_on g {0..1}"
    38 
   104 
    39 
   105 
    40 subsection {* Some lemmas about these concepts. *}
   106 subsection{*Invariance theorems*}
    41 
   107 
    42 lemma injective_imp_simple_path: "injective_path g \<Longrightarrow> simple_path g"
   108 lemma path_eq: "path p \<Longrightarrow> (\<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t) \<Longrightarrow> path q"
    43   unfolding injective_path_def simple_path_def
   109   using continuous_on_eq path_def by blast
    44   by auto
   110 
       
   111 lemma path_continuous_image: "path g \<Longrightarrow> continuous_on (path_image g) f \<Longrightarrow> path(f o g)"
       
   112   unfolding path_def path_image_def
       
   113   using continuous_on_compose by blast
       
   114 
       
   115 lemma path_translation_eq:
       
   116   fixes g :: "real \<Rightarrow> 'a :: real_normed_vector"
       
   117   shows "path((\<lambda>x. a + x) o g) = path g"
       
   118 proof -
       
   119   have g: "g = (\<lambda>x. -a + x) o ((\<lambda>x. a + x) o g)"
       
   120     by (rule ext) simp
       
   121   show ?thesis
       
   122     unfolding path_def
       
   123     apply safe
       
   124     apply (subst g)
       
   125     apply (rule continuous_on_compose)
       
   126     apply (auto intro: continuous_intros)
       
   127     done
       
   128 qed
       
   129 
       
   130 lemma path_linear_image_eq:
       
   131   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   132    assumes "linear f" "inj f"
       
   133      shows "path(f o g) = path g"
       
   134 proof -
       
   135   from linear_injective_left_inverse [OF assms]
       
   136   obtain h where h: "linear h" "h \<circ> f = id"
       
   137     by blast
       
   138   then have g: "g = h o (f o g)"
       
   139     by (metis comp_assoc id_comp)
       
   140   show ?thesis
       
   141     unfolding path_def
       
   142     using h assms
       
   143     by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear)
       
   144 qed
       
   145 
       
   146 lemma pathstart_translation: "pathstart((\<lambda>x. a + x) o g) = a + pathstart g"
       
   147   by (simp add: pathstart_def)
       
   148 
       
   149 lemma pathstart_linear_image_eq: "linear f \<Longrightarrow> pathstart(f o g) = f(pathstart g)"
       
   150   by (simp add: pathstart_def)
       
   151 
       
   152 lemma pathfinish_translation: "pathfinish((\<lambda>x. a + x) o g) = a + pathfinish g"
       
   153   by (simp add: pathfinish_def)
       
   154 
       
   155 lemma pathfinish_linear_image: "linear f \<Longrightarrow> pathfinish(f o g) = f(pathfinish g)"
       
   156   by (simp add: pathfinish_def)
       
   157 
       
   158 lemma path_image_translation: "path_image((\<lambda>x. a + x) o g) = (\<lambda>x. a + x) ` (path_image g)"
       
   159   by (simp add: image_comp path_image_def)
       
   160 
       
   161 lemma path_image_linear_image: "linear f \<Longrightarrow> path_image(f o g) = f ` (path_image g)"
       
   162   by (simp add: image_comp path_image_def)
       
   163 
       
   164 lemma reversepath_translation: "reversepath((\<lambda>x. a + x) o g) = (\<lambda>x. a + x) o reversepath g"
       
   165   by (rule ext) (simp add: reversepath_def)
       
   166 
       
   167 lemma reversepath_linear_image: "linear f \<Longrightarrow> reversepath(f o g) = f o reversepath g"
       
   168   by (rule ext) (simp add: reversepath_def)
       
   169 
       
   170 lemma joinpaths_translation:
       
   171     "((\<lambda>x. a + x) o g1) +++ ((\<lambda>x. a + x) o g2) = (\<lambda>x. a + x) o (g1 +++ g2)"
       
   172   by (rule ext) (simp add: joinpaths_def)
       
   173 
       
   174 lemma joinpaths_linear_image: "linear f \<Longrightarrow> (f o g1) +++ (f o g2) = f o (g1 +++ g2)"
       
   175   by (rule ext) (simp add: joinpaths_def)
       
   176 
       
   177 lemma simple_path_translation_eq: 
       
   178   fixes g :: "real \<Rightarrow> 'a::euclidean_space"
       
   179   shows "simple_path((\<lambda>x. a + x) o g) = simple_path g"
       
   180   by (simp add: simple_path_def path_translation_eq)
       
   181 
       
   182 lemma simple_path_linear_image_eq:
       
   183   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   184   assumes "linear f" "inj f"
       
   185     shows "simple_path(f o g) = simple_path g"
       
   186   using assms inj_on_eq_iff [of f]
       
   187   by (auto simp: path_linear_image_eq simple_path_def path_translation_eq)
       
   188 
       
   189 lemma arc_translation_eq:
       
   190   fixes g :: "real \<Rightarrow> 'a::euclidean_space"
       
   191   shows "arc((\<lambda>x. a + x) o g) = arc g"
       
   192   by (auto simp: arc_def inj_on_def path_translation_eq)
       
   193 
       
   194 lemma arc_linear_image_eq:
       
   195   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
       
   196    assumes "linear f" "inj f"
       
   197      shows  "arc(f o g) = arc g"
       
   198   using assms inj_on_eq_iff [of f]
       
   199   by (auto simp: arc_def inj_on_def path_linear_image_eq)
       
   200 
       
   201 subsection{*Basic lemmas about paths*}
       
   202 
       
   203 lemma arc_imp_simple_path: "arc g \<Longrightarrow> simple_path g"
       
   204   by (simp add: arc_def inj_on_def simple_path_def)
       
   205 
       
   206 lemma arc_imp_path: "arc g \<Longrightarrow> path g"
       
   207   using arc_def by blast
       
   208 
       
   209 lemma simple_path_imp_path: "simple_path g \<Longrightarrow> path g"
       
   210   using simple_path_def by blast
       
   211 
       
   212 lemma simple_path_cases: "simple_path g \<Longrightarrow> arc g \<or> pathfinish g = pathstart g"
       
   213   unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def
       
   214   by (force)
       
   215 
       
   216 lemma simple_path_imp_arc: "simple_path g \<Longrightarrow> pathfinish g \<noteq> pathstart g \<Longrightarrow> arc g"
       
   217   using simple_path_cases by auto
       
   218 
       
   219 lemma arc_distinct_ends: "arc g \<Longrightarrow> pathfinish g \<noteq> pathstart g"
       
   220   unfolding arc_def inj_on_def pathfinish_def pathstart_def
       
   221   by fastforce
       
   222 
       
   223 lemma arc_simple_path: "arc g \<longleftrightarrow> simple_path g \<and> pathfinish g \<noteq> pathstart g"
       
   224   using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast
       
   225 
       
   226 lemma simple_path_eq_arc: "pathfinish g \<noteq> pathstart g \<Longrightarrow> (simple_path g = arc g)"
       
   227   by (simp add: arc_simple_path)
    45 
   228 
    46 lemma path_image_nonempty: "path_image g \<noteq> {}"
   229 lemma path_image_nonempty: "path_image g \<noteq> {}"
    47   unfolding path_image_def image_is_empty box_eq_empty
   230   unfolding path_image_def image_is_empty box_eq_empty
    48   by auto
   231   by auto
    49 
   232 
    55   unfolding pathfinish_def path_image_def
   238   unfolding pathfinish_def path_image_def
    56   by auto
   239   by auto
    57 
   240 
    58 lemma connected_path_image[intro]: "path g \<Longrightarrow> connected (path_image g)"
   241 lemma connected_path_image[intro]: "path g \<Longrightarrow> connected (path_image g)"
    59   unfolding path_def path_image_def
   242   unfolding path_def path_image_def
    60   apply (erule connected_continuous_image)
   243   using connected_continuous_image connected_Icc by blast
    61   apply (rule convex_connected, rule convex_real_interval)
       
    62   done
       
    63 
   244 
    64 lemma compact_path_image[intro]: "path g \<Longrightarrow> compact (path_image g)"
   245 lemma compact_path_image[intro]: "path g \<Longrightarrow> compact (path_image g)"
    65   unfolding path_def path_image_def
   246   unfolding path_def path_image_def
    66   apply (erule compact_continuous_image)
   247   using compact_continuous_image connected_Icc by blast
    67   apply (rule compact_Icc)
       
    68   done
       
    69 
   248 
    70 lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g"
   249 lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g"
    71   unfolding reversepath_def
   250   unfolding reversepath_def
    72   by auto
   251   by auto
    73 
   252 
    89 
   268 
    90 lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g"
   269 lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g"
    91 proof -
   270 proof -
    92   have *: "\<And>g. path_image (reversepath g) \<subseteq> path_image g"
   271   have *: "\<And>g. path_image (reversepath g) \<subseteq> path_image g"
    93     unfolding path_image_def subset_eq reversepath_def Ball_def image_iff
   272     unfolding path_image_def subset_eq reversepath_def Ball_def image_iff
    94     apply rule
   273     by force
    95     apply rule
       
    96     apply (erule bexE)
       
    97     apply (rule_tac x="1 - xa" in bexI)
       
    98     apply auto
       
    99     done
       
   100   show ?thesis
   274   show ?thesis
   101     using *[of g] *[of "reversepath g"]
   275     using *[of g] *[of "reversepath g"]
   102     unfolding reversepath_reversepath
   276     unfolding reversepath_reversepath
   103     by auto
   277     by auto
   104 qed
   278 qed
   116   show ?thesis
   290   show ?thesis
   117     using *[of "reversepath g"] *[of g]
   291     using *[of "reversepath g"] *[of g]
   118     unfolding reversepath_reversepath
   292     unfolding reversepath_reversepath
   119     by (rule iffI)
   293     by (rule iffI)
   120 qed
   294 qed
       
   295 
       
   296 lemma arc_reversepath:
       
   297   assumes "arc g" shows "arc(reversepath g)"
       
   298 proof -
       
   299   have injg: "inj_on g {0..1}"
       
   300     using assms
       
   301     by (simp add: arc_def)
       
   302   have **: "\<And>x y::real. 1-x = 1-y \<Longrightarrow> x = y"
       
   303     by simp
       
   304   show ?thesis
       
   305     apply (auto simp: arc_def inj_on_def path_reversepath)
       
   306     apply (simp add: arc_imp_path assms)
       
   307     apply (rule **)
       
   308     apply (rule inj_onD [OF injg])
       
   309     apply (auto simp: reversepath_def)
       
   310     done
       
   311 qed
       
   312 
       
   313 lemma simple_path_reversepath: "simple_path g \<Longrightarrow> simple_path (reversepath g)"
       
   314   apply (simp add: simple_path_def)
       
   315   apply (force simp: reversepath_def)
       
   316   done
   121 
   317 
   122 lemmas reversepath_simps =
   318 lemmas reversepath_simps =
   123   path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
   319   path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath
   124 
   320 
   125 lemma path_join[simp]:
   321 lemma path_join[simp]:
   160     apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros)
   356     apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros)
   161     apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2)
   357     apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2)
   162     done
   358     done
   163 qed
   359 qed
   164 
   360 
       
   361 section {*Path Images*}
       
   362 
       
   363 lemma bounded_path_image: "path g \<Longrightarrow> bounded(path_image g)"
       
   364   by (simp add: compact_imp_bounded compact_path_image)
       
   365 
       
   366 lemma closed_path_image: 
       
   367   fixes g :: "real \<Rightarrow> 'a::t2_space"
       
   368   shows "path g \<Longrightarrow> closed(path_image g)"
       
   369   by (metis compact_path_image compact_imp_closed)
       
   370 
       
   371 lemma connected_simple_path_image: "simple_path g \<Longrightarrow> connected(path_image g)"
       
   372   by (metis connected_path_image simple_path_imp_path)
       
   373 
       
   374 lemma compact_simple_path_image: "simple_path g \<Longrightarrow> compact(path_image g)"
       
   375   by (metis compact_path_image simple_path_imp_path)
       
   376 
       
   377 lemma bounded_simple_path_image: "simple_path g \<Longrightarrow> bounded(path_image g)"
       
   378   by (metis bounded_path_image simple_path_imp_path)
       
   379 
       
   380 lemma closed_simple_path_image:
       
   381   fixes g :: "real \<Rightarrow> 'a::t2_space"
       
   382   shows "simple_path g \<Longrightarrow> closed(path_image g)"
       
   383   by (metis closed_path_image simple_path_imp_path)
       
   384 
       
   385 lemma connected_arc_image: "arc g \<Longrightarrow> connected(path_image g)"
       
   386   by (metis connected_path_image arc_imp_path)
       
   387 
       
   388 lemma compact_arc_image: "arc g \<Longrightarrow> compact(path_image g)"
       
   389   by (metis compact_path_image arc_imp_path)
       
   390 
       
   391 lemma bounded_arc_image: "arc g \<Longrightarrow> bounded(path_image g)"
       
   392   by (metis bounded_path_image arc_imp_path)
       
   393 
       
   394 lemma closed_arc_image:
       
   395   fixes g :: "real \<Rightarrow> 'a::t2_space"
       
   396   shows "arc g \<Longrightarrow> closed(path_image g)"
       
   397   by (metis closed_path_image arc_imp_path)
       
   398 
   165 lemma path_image_join_subset: "path_image (g1 +++ g2) \<subseteq> path_image g1 \<union> path_image g2"
   399 lemma path_image_join_subset: "path_image (g1 +++ g2) \<subseteq> path_image g1 \<union> path_image g2"
   166   unfolding path_image_def joinpaths_def
   400   unfolding path_image_def joinpaths_def
   167   by auto
   401   by auto
   168 
   402 
   169 lemma subset_path_image_join:
   403 lemma subset_path_image_join:
   172   shows "path_image (g1 +++ g2) \<subseteq> s"
   406   shows "path_image (g1 +++ g2) \<subseteq> s"
   173   using path_image_join_subset[of g1 g2] and assms
   407   using path_image_join_subset[of g1 g2] and assms
   174   by auto
   408   by auto
   175 
   409 
   176 lemma path_image_join:
   410 lemma path_image_join:
   177   assumes "pathfinish g1 = pathstart g2"
   411     "pathfinish g1 = pathstart g2 \<Longrightarrow> path_image(g1 +++ g2) = path_image g1 \<union> path_image g2"
   178   shows "path_image (g1 +++ g2) = path_image g1 \<union> path_image g2"
   412   apply (rule subset_antisym [OF path_image_join_subset])
   179   apply rule
   413   apply (auto simp: pathfinish_def pathstart_def path_image_def joinpaths_def image_def)
   180   apply (rule path_image_join_subset)
   414   apply (drule sym)
   181   apply rule
   415   apply (rule_tac x="xa/2" in bexI, auto)
   182   unfolding Un_iff
   416   apply (rule ccontr)
   183 proof (erule disjE)
   417   apply (drule_tac x="(xa+1)/2" in bspec)
   184   fix x
   418   apply (auto simp: field_simps)
   185   assume "x \<in> path_image g1"
   419   apply (drule_tac x="1/2" in bspec, auto)
   186   then obtain y where y: "y \<in> {0..1}" "x = g1 y"
   420   done
   187     unfolding path_image_def image_iff by auto
       
   188   then show "x \<in> path_image (g1 +++ g2)"
       
   189     unfolding joinpaths_def path_image_def image_iff
       
   190     apply (rule_tac x="(1/2) *\<^sub>R y" in bexI)
       
   191     apply auto
       
   192     done
       
   193 next
       
   194   fix x
       
   195   assume "x \<in> path_image g2"
       
   196   then obtain y where y: "y \<in> {0..1}" "x = g2 y"
       
   197     unfolding path_image_def image_iff by auto
       
   198   then show "x \<in> path_image (g1 +++ g2)"
       
   199     unfolding joinpaths_def path_image_def image_iff
       
   200     apply (rule_tac x="(1/2) *\<^sub>R (y + 1)" in bexI)
       
   201     using assms(1)[unfolded pathfinish_def pathstart_def]
       
   202     apply (auto simp add: add_divide_distrib)
       
   203     done
       
   204 qed
       
   205 
   421 
   206 lemma not_in_path_image_join:
   422 lemma not_in_path_image_join:
   207   assumes "x \<notin> path_image g1"
   423   assumes "x \<notin> path_image g1"
   208     and "x \<notin> path_image g2"
   424     and "x \<notin> path_image g2"
   209   shows "x \<notin> path_image (g1 +++ g2)"
   425   shows "x \<notin> path_image (g1 +++ g2)"
   210   using assms and path_image_join_subset[of g1 g2]
   426   using assms and path_image_join_subset[of g1 g2]
   211   by auto
   427   by auto
   212 
   428 
   213 lemma simple_path_reversepath:
   429 lemma pathstart_compose: "pathstart(f o p) = f(pathstart p)"
   214   assumes "simple_path g"
   430   by (simp add: pathstart_def)
   215   shows "simple_path (reversepath g)"
   431 
       
   432 lemma pathfinish_compose: "pathfinish(f o p) = f(pathfinish p)"
       
   433   by (simp add: pathfinish_def)
       
   434 
       
   435 lemma path_image_compose: "path_image (f o p) = f ` (path_image p)"
       
   436   by (simp add: image_comp path_image_def)
       
   437 
       
   438 lemma path_compose_join: "f o (p +++ q) = (f o p) +++ (f o q)"
       
   439   by (rule ext) (simp add: joinpaths_def)
       
   440 
       
   441 lemma path_compose_reversepath: "f o reversepath p = reversepath(f o p)"
       
   442   by (rule ext) (simp add: reversepath_def)
       
   443 
       
   444 lemma join_paths_eq:
       
   445   "(\<And>t. t \<in> {0..1} \<Longrightarrow> p t = p' t) \<Longrightarrow>
       
   446    (\<And>t. t \<in> {0..1} \<Longrightarrow> q t = q' t)
       
   447    \<Longrightarrow>  t \<in> {0..1} \<Longrightarrow> (p +++ q) t = (p' +++ q') t"
       
   448   by (auto simp: joinpaths_def)
       
   449 
       
   450 lemma simple_path_inj_on: "simple_path g \<Longrightarrow> inj_on g {0<..<1}"
       
   451   by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def)
       
   452 
       
   453 
       
   454 subsection{*Simple paths with the endpoints removed*}
       
   455 
       
   456 lemma simple_path_endless:
       
   457     "simple_path c \<Longrightarrow> path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}"
       
   458   apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def Bex_def image_def)
       
   459   apply (metis eq_iff le_less_linear)
       
   460   apply (metis leD linear)
       
   461   using less_eq_real_def zero_le_one apply blast
       
   462   using less_eq_real_def zero_le_one apply blast
       
   463   done
       
   464 
       
   465 lemma connected_simple_path_endless:
       
   466     "simple_path c \<Longrightarrow> connected(path_image c - {pathstart c,pathfinish c})"
       
   467 apply (simp add: simple_path_endless)
       
   468 apply (rule connected_continuous_image)
       
   469 apply (meson continuous_on_subset greaterThanLessThan_subseteq_atLeastAtMost_iff le_numeral_extra(3) le_numeral_extra(4) path_def simple_path_imp_path)
       
   470 by auto
       
   471 
       
   472 lemma nonempty_simple_path_endless:
       
   473     "simple_path c \<Longrightarrow> path_image c - {pathstart c,pathfinish c} \<noteq> {}"
       
   474   by (simp add: simple_path_endless)
       
   475 
       
   476 
       
   477 subsection{* The operations on paths*}
       
   478 
       
   479 lemma path_image_subset_reversepath: "path_image(reversepath g) \<le> path_image g"
       
   480   by (auto simp: path_image_def reversepath_def)
       
   481 
       
   482 lemma continuous_on_op_minus: "continuous_on (s::real set) (op - x)"
       
   483   by (rule continuous_intros | simp)+
       
   484 
       
   485 lemma path_imp_reversepath: "path g \<Longrightarrow> path(reversepath g)"
       
   486   apply (auto simp: path_def reversepath_def)
       
   487   using continuous_on_compose [of "{0..1}" "\<lambda>x. 1 - x" g]
       
   488   apply (auto simp: continuous_on_op_minus)
       
   489   done
       
   490 
       
   491 lemma forall_01_trivial: "(\<forall>x\<in>{0..1}. x \<le> 0 \<longrightarrow> P x) \<longleftrightarrow> P (0::real)"
       
   492   by auto
       
   493 
       
   494 lemma forall_half1_trivial: "(\<forall>x\<in>{1/2..1}. x * 2 \<le> 1 \<longrightarrow> P x) \<longleftrightarrow> P (1/2::real)"
       
   495   by auto (metis add_divide_distrib mult_2_right real_sum_of_halves)
       
   496 
       
   497 lemma continuous_on_joinpaths:
       
   498   assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2"
       
   499     shows "continuous_on {0..1} (g1 +++ g2)"
       
   500 proof -
       
   501   have *: "{0..1::real} = {0..1/2} \<union> {1/2..1}"
       
   502     by auto
       
   503   have gg: "g2 0 = g1 1"
       
   504     by (metis assms(3) pathfinish_def pathstart_def)
       
   505   have 1: "continuous_on {0..1 / 2} (g1 +++ g2)"
       
   506     apply (rule continuous_on_eq [of _ "g1 o (\<lambda>x. 2*x)"])
       
   507     apply (simp add: joinpaths_def)
       
   508     apply (rule continuous_intros | simp add: assms)+
       
   509     done
       
   510   have 2: "continuous_on {1 / 2..1} (g1 +++ g2)"
       
   511     apply (rule continuous_on_eq [of _ "g2 o (\<lambda>x. 2*x-1)"])
       
   512     apply (simp add: joinpaths_def)
       
   513     apply (rule continuous_intros | simp add: forall_half1_trivial gg)+
       
   514     apply (rule continuous_on_subset)
       
   515     apply (rule assms, auto)
       
   516     done
       
   517   show ?thesis
       
   518     apply (subst *)
       
   519     apply (rule continuous_on_union)
       
   520     using 1 2
       
   521     apply auto
       
   522     done
       
   523 qed
       
   524 
       
   525 lemma path_join_imp: "\<lbrakk>path g1; path g2; pathfinish g1 = pathstart g2\<rbrakk> \<Longrightarrow> path(g1 +++ g2)"
       
   526   by (simp add: path_join)
       
   527 
       
   528 lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join
       
   529 
       
   530 lemma simple_path_join_loop:
       
   531   assumes "arc g1" "arc g2" 
       
   532           "pathfinish g1 = pathstart g2"  "pathfinish g2 = pathstart g1" 
       
   533           "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
       
   534   shows "simple_path(g1 +++ g2)"
       
   535 proof -
       
   536   have injg1: "inj_on g1 {0..1}"
       
   537     using assms
       
   538     by (simp add: arc_def)
       
   539   have injg2: "inj_on g2 {0..1}"
       
   540     using assms
       
   541     by (simp add: arc_def)
       
   542   have g12: "g1 1 = g2 0" 
       
   543    and g21: "g2 1 = g1 0" 
       
   544    and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g1 0, g2 0}"
       
   545     using assms
       
   546     by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def)
       
   547   { fix x and y::real
       
   548     assume xyI: "x = 1 \<longrightarrow> y \<noteq> 0" 
       
   549        and xy: "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1" "g2 (2 * x - 1) = g1 (2 * y)"
       
   550     have g1im: "g1 (2 * y) \<in> g1 ` {0..1} \<inter> g2 ` {0..1}"
       
   551       using xy
       
   552       apply simp
       
   553       apply (rule_tac x="2 * x - 1" in image_eqI, auto)
       
   554       done
       
   555     have False
       
   556       using subsetD [OF sb g1im] xy 
       
   557       apply auto
       
   558       apply (drule inj_onD [OF injg1])
       
   559       using g21 [symmetric] xyI
       
   560       apply (auto dest: inj_onD [OF injg2])
       
   561       done
       
   562    } note * = this
       
   563   { fix x and y::real
       
   564     assume xy: "y \<le> 1" "0 \<le> x" "\<not> y * 2 \<le> 1" "x * 2 \<le> 1" "g1 (2 * x) = g2 (2 * y - 1)"
       
   565     have g1im: "g1 (2 * x) \<in> g1 ` {0..1} \<inter> g2 ` {0..1}"
       
   566       using xy
       
   567       apply simp
       
   568       apply (rule_tac x="2 * x" in image_eqI, auto)
       
   569       done
       
   570     have "x = 0 \<and> y = 1"
       
   571       using subsetD [OF sb g1im] xy 
       
   572       apply auto
       
   573       apply (force dest: inj_onD [OF injg1])
       
   574       using  g21 [symmetric]
       
   575       apply (auto dest: inj_onD [OF injg2])
       
   576       done
       
   577    } note ** = this
       
   578   show ?thesis
       
   579     using assms
       
   580     apply (simp add: arc_def simple_path_def path_join, clarify)
       
   581     apply (simp add: joinpaths_def split: split_if_asm)
       
   582     apply (force dest: inj_onD [OF injg1])
       
   583     apply (metis *)
       
   584     apply (metis **)
       
   585     apply (force dest: inj_onD [OF injg2])
       
   586     done
       
   587 qed
       
   588 
       
   589 lemma arc_join:
       
   590   assumes "arc g1" "arc g2" 
       
   591           "pathfinish g1 = pathstart g2"
       
   592           "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
       
   593     shows "arc(g1 +++ g2)"
       
   594 proof -
       
   595   have injg1: "inj_on g1 {0..1}"
       
   596     using assms
       
   597     by (simp add: arc_def)
       
   598   have injg2: "inj_on g2 {0..1}"
       
   599     using assms
       
   600     by (simp add: arc_def)
       
   601   have g11: "g1 1 = g2 0"
       
   602    and sb:  "g1 ` {0..1} \<inter> g2 ` {0..1} \<subseteq> {g2 0}"
       
   603     using assms
       
   604     by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def)
       
   605   { fix x and y::real
       
   606     assume xy: "x \<le> 1" "0 \<le> y" " y * 2 \<le> 1" "\<not> x * 2 \<le> 1" "g2 (2 * x - 1) = g1 (2 * y)"       
       
   607     have g1im: "g1 (2 * y) \<in> g1 ` {0..1} \<inter> g2 ` {0..1}"
       
   608       using xy
       
   609       apply simp
       
   610       apply (rule_tac x="2 * x - 1" in image_eqI, auto)
       
   611       done
       
   612     have False
       
   613       using subsetD [OF sb g1im] xy 
       
   614       by (auto dest: inj_onD [OF injg2])
       
   615    } note * = this
       
   616   show ?thesis
       
   617     apply (simp add: arc_def inj_on_def)
       
   618     apply (clarsimp simp add: arc_imp_path assms path_join)
       
   619     apply (simp add: joinpaths_def split: split_if_asm)
       
   620     apply (force dest: inj_onD [OF injg1])
       
   621     apply (metis *)
       
   622     apply (metis *)
       
   623     apply (force dest: inj_onD [OF injg2])
       
   624     done
       
   625 qed
       
   626 
       
   627 lemma reversepath_joinpaths:
       
   628     "pathfinish g1 = pathstart g2 \<Longrightarrow> reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1"
       
   629   unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def
       
   630   by (rule ext) (auto simp: mult.commute)
       
   631 
       
   632 
       
   633 subsection{* Choosing a subpath of an existing path*}
       
   634     
       
   635 definition subpath :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> real \<Rightarrow> 'a::real_normed_vector"
       
   636   where "subpath a b g \<equiv> \<lambda>x. g((b - a) * x + a)"
       
   637 
       
   638 lemma path_image_subpath_gen [simp]: 
       
   639   fixes g :: "real \<Rightarrow> 'a::real_normed_vector"
       
   640   shows "path_image(subpath u v g) = g ` (closed_segment u v)"
       
   641   apply (simp add: closed_segment_real_eq path_image_def subpath_def)
       
   642   apply (subst o_def [of g, symmetric])
       
   643   apply (simp add: image_comp [symmetric])
       
   644   done
       
   645 
       
   646 lemma path_image_subpath [simp]:
       
   647   fixes g :: "real \<Rightarrow> 'a::real_normed_vector"
       
   648   shows "path_image(subpath u v g) = (if u \<le> v then g ` {u..v} else g ` {v..u})"
       
   649   by (simp add: closed_segment_eq_real_ivl)
       
   650 
       
   651 lemma path_subpath [simp]:
       
   652   fixes g :: "real \<Rightarrow> 'a::real_normed_vector"
       
   653   assumes "path g" "u \<in> {0..1}" "v \<in> {0..1}"
       
   654     shows "path(subpath u v g)"
       
   655 proof -
       
   656   have "continuous_on {0..1} (g o (\<lambda>x. ((v-u) * x+ u)))"
       
   657     apply (rule continuous_intros | simp)+
       
   658     apply (simp add: image_affinity_atLeastAtMost [where c=u])
       
   659     using assms
       
   660     apply (auto simp: path_def continuous_on_subset)
       
   661     done
       
   662   then show ?thesis
       
   663     by (simp add: path_def subpath_def)
       
   664 qed
       
   665 
       
   666 lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)"
       
   667   by (simp add: pathstart_def subpath_def)
       
   668 
       
   669 lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)"
       
   670   by (simp add: pathfinish_def subpath_def)
       
   671 
       
   672 lemma subpath_trivial [simp]: "subpath 0 1 g = g"
       
   673   by (simp add: subpath_def)
       
   674 
       
   675 lemma subpath_reversepath: "subpath 1 0 g = reversepath g"
       
   676   by (simp add: reversepath_def subpath_def)
       
   677 
       
   678 lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g"
       
   679   by (simp add: reversepath_def subpath_def algebra_simps)
       
   680 
       
   681 lemma subpath_translation: "subpath u v ((\<lambda>x. a + x) o g) = (\<lambda>x. a + x) o subpath u v g"
       
   682   by (rule ext) (simp add: subpath_def)
       
   683 
       
   684 lemma subpath_linear_image: "linear f \<Longrightarrow> subpath u v (f o g) = f o subpath u v g"
       
   685   by (rule ext) (simp add: subpath_def)
       
   686 
       
   687 lemma affine_ineq: 
       
   688   fixes x :: "'a::linordered_idom" 
       
   689   assumes "x \<le> 1" "v < u"
       
   690     shows "v + x * u \<le> u + x * v"
       
   691 proof -
       
   692   have "(1-x)*(u-v) \<ge> 0"
       
   693     using assms by auto
       
   694   then show ?thesis
       
   695     by (simp add: algebra_simps)
       
   696 qed
       
   697 
       
   698 lemma simple_path_subpath_eq: 
       
   699   "simple_path(subpath u v g) \<longleftrightarrow>
       
   700      path(subpath u v g) \<and> u\<noteq>v \<and>
       
   701      (\<forall>x y. x \<in> closed_segment u v \<and> y \<in> closed_segment u v \<and> g x = g y
       
   702                 \<longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u)"
       
   703     (is "?lhs = ?rhs")
       
   704 proof (rule iffI)
       
   705   assume ?lhs
       
   706   then have p: "path (\<lambda>x. g ((v - u) * x + u))"
       
   707         and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk> 
       
   708                   \<Longrightarrow> x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0)"
       
   709     by (auto simp: simple_path_def subpath_def)
       
   710   { fix x y
       
   711     assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y"
       
   712     then have "x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u"
       
   713     using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p
       
   714     by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost divide_simps 
       
   715        split: split_if_asm)
       
   716   } moreover
       
   717   have "path(subpath u v g) \<and> u\<noteq>v"
       
   718     using sim [of "1/3" "2/3"] p
       
   719     by (auto simp: subpath_def)
       
   720   ultimately show ?rhs
       
   721     by metis
       
   722 next
       
   723   assume ?rhs
       
   724   then 
       
   725   have d1: "\<And>x y. \<lbrakk>g x = g y; u \<le> x; x \<le> v; u \<le> y; y \<le> v\<rbrakk> \<Longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u"
       
   726    and d2: "\<And>x y. \<lbrakk>g x = g y; v \<le> x; x \<le> u; v \<le> y; y \<le> u\<rbrakk> \<Longrightarrow> x = y \<or> x = u \<and> y = v \<or> x = v \<and> y = u"
       
   727    and ne: "u < v \<or> v < u"
       
   728    and psp: "path (subpath u v g)"
       
   729     by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost)
       
   730   have [simp]: "\<And>x. u + x * v = v + x * u \<longleftrightarrow> u=v \<or> x=1"
       
   731     by algebra
       
   732   show ?lhs using psp ne
       
   733     unfolding simple_path_def subpath_def
       
   734     by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2)
       
   735 qed
       
   736 
       
   737 lemma arc_subpath_eq: 
       
   738   "arc(subpath u v g) \<longleftrightarrow> path(subpath u v g) \<and> u\<noteq>v \<and> inj_on g (closed_segment u v)"
       
   739     (is "?lhs = ?rhs")
       
   740 proof (rule iffI)
       
   741   assume ?lhs
       
   742   then have p: "path (\<lambda>x. g ((v - u) * x + u))"
       
   743         and sim: "(\<And>x y. \<lbrakk>x\<in>{0..1}; y\<in>{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\<rbrakk> 
       
   744                   \<Longrightarrow> x = y)"
       
   745     by (auto simp: arc_def inj_on_def subpath_def)
       
   746   { fix x y
       
   747     assume "x \<in> closed_segment u v" "y \<in> closed_segment u v" "g x = g y"
       
   748     then have "x = y"
       
   749     using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p
       
   750     by (force simp add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost divide_simps 
       
   751        split: split_if_asm)
       
   752   } moreover
       
   753   have "path(subpath u v g) \<and> u\<noteq>v"
       
   754     using sim [of "1/3" "2/3"] p
       
   755     by (auto simp: subpath_def)
       
   756   ultimately show ?rhs
       
   757     unfolding inj_on_def   
       
   758     by metis
       
   759 next
       
   760   assume ?rhs
       
   761   then 
       
   762   have d1: "\<And>x y. \<lbrakk>g x = g y; u \<le> x; x \<le> v; u \<le> y; y \<le> v\<rbrakk> \<Longrightarrow> x = y"
       
   763    and d2: "\<And>x y. \<lbrakk>g x = g y; v \<le> x; x \<le> u; v \<le> y; y \<le> u\<rbrakk> \<Longrightarrow> x = y"
       
   764    and ne: "u < v \<or> v < u"
       
   765    and psp: "path (subpath u v g)"
       
   766     by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost)
       
   767   show ?lhs using psp ne
       
   768     unfolding arc_def subpath_def inj_on_def
       
   769     by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2)
       
   770 qed
       
   771 
       
   772 
       
   773 lemma simple_path_subpath: 
       
   774   assumes "simple_path g" "u \<in> {0..1}" "v \<in> {0..1}" "u \<noteq> v"
       
   775   shows "simple_path(subpath u v g)"
   216   using assms
   776   using assms
   217   unfolding simple_path_def reversepath_def
   777   apply (simp add: simple_path_subpath_eq simple_path_imp_path)
   218   apply -
   778   apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce)
   219   apply (rule ballI)+
   779   done
   220   apply (erule_tac x="1-x" in ballE)
   780 
   221   apply (erule_tac x="1-y" in ballE)
   781 lemma arc_simple_path_subpath:
   222   apply auto
   782     "\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; g u \<noteq> g v\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
   223   done
   783   by (force intro: simple_path_subpath simple_path_imp_arc)
   224 
   784 
   225 lemma simple_path_join_loop:
   785 lemma arc_subpath_arc:
   226   assumes "injective_path g1"
   786     "\<lbrakk>arc g; u \<in> {0..1}; v \<in> {0..1}; u \<noteq> v\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
   227     and "injective_path g2"
   787   by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD)
   228     and "pathfinish g2 = pathstart g1"
   788 
   229     and "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g1, pathstart g2}"
   789 lemma arc_simple_path_subpath_interior: 
   230   shows "simple_path (g1 +++ g2)"
   790     "\<lbrakk>simple_path g; u \<in> {0..1}; v \<in> {0..1}; u \<noteq> v; \<bar>u-v\<bar> < 1\<rbrakk> \<Longrightarrow> arc(subpath u v g)"
   231   unfolding simple_path_def
   791     apply (rule arc_simple_path_subpath)
   232 proof (intro ballI impI)
   792     apply (force simp: simple_path_def)+
   233   let ?g = "g1 +++ g2"
   793     done
   234   note inj = assms(1,2)[unfolded injective_path_def, rule_format]
   794 
   235   fix x y :: real
   795 lemma path_image_subpath_subset: 
   236   assume xy: "x \<in> {0..1}" "y \<in> {0..1}" "?g x = ?g y"
   796     "\<lbrakk>path g; u \<in> {0..1}; v \<in> {0..1}\<rbrakk> \<Longrightarrow> path_image(subpath u v g) \<subseteq> path_image g"
   237   show "x = y \<or> x = 0 \<and> y = 1 \<or> x = 1 \<and> y = 0"
   797   apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost)
   238   proof (cases "x \<le> 1/2", case_tac[!] "y \<le> 1/2", unfold not_le)
   798   apply (auto simp: path_image_def)
   239     assume as: "x \<le> 1 / 2" "y \<le> 1 / 2"
   799   done
   240     then have "g1 (2 *\<^sub>R x) = g1 (2 *\<^sub>R y)"
   800 
   241       using xy(3)
   801 lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p"
   242       unfolding joinpaths_def
   802   by (rule ext) (simp add: joinpaths_def subpath_def divide_simps)
   243       by auto
       
   244     moreover have "2 *\<^sub>R x \<in> {0..1}" "2 *\<^sub>R y \<in> {0..1}"
       
   245       using xy(1,2) as
       
   246       by auto
       
   247     ultimately show ?thesis
       
   248       using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"]
       
   249       by auto
       
   250   next
       
   251     assume as: "x > 1 / 2" "y > 1 / 2"
       
   252     then have "g2 (2 *\<^sub>R x - 1) = g2 (2 *\<^sub>R y - 1)"
       
   253       using xy(3)
       
   254       unfolding joinpaths_def
       
   255       by auto
       
   256     moreover have "2 *\<^sub>R x - 1 \<in> {0..1}" "2 *\<^sub>R y - 1 \<in> {0..1}"
       
   257       using xy(1,2) as
       
   258       by auto
       
   259     ultimately show ?thesis
       
   260       using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] by auto
       
   261   next
       
   262     assume as: "x \<le> 1 / 2" "y > 1 / 2"
       
   263     then have "?g x \<in> path_image g1" "?g y \<in> path_image g2"
       
   264       unfolding path_image_def joinpaths_def
       
   265       using xy(1,2) by auto
       
   266     moreover have "?g y \<noteq> pathstart g2"
       
   267       using as(2)
       
   268       unfolding pathstart_def joinpaths_def
       
   269       using inj(2)[of "2 *\<^sub>R y - 1" 0] and xy(2)
       
   270       by (auto simp add: field_simps)
       
   271     ultimately have *: "?g x = pathstart g1"
       
   272       using assms(4)
       
   273       unfolding xy(3)
       
   274       by auto
       
   275     then have "x = 0"
       
   276       unfolding pathstart_def joinpaths_def
       
   277       using as(1) and xy(1)
       
   278       using inj(1)[of "2 *\<^sub>R x" 0]
       
   279       by auto
       
   280     moreover have "y = 1"
       
   281       using *
       
   282       unfolding xy(3) assms(3)[symmetric]
       
   283       unfolding joinpaths_def pathfinish_def
       
   284       using as(2) and xy(2)
       
   285       using inj(2)[of "2 *\<^sub>R y - 1" 1]
       
   286       by auto
       
   287     ultimately show ?thesis
       
   288       by auto
       
   289   next
       
   290     assume as: "x > 1 / 2" "y \<le> 1 / 2"
       
   291     then have "?g x \<in> path_image g2" and "?g y \<in> path_image g1"
       
   292       unfolding path_image_def joinpaths_def
       
   293       using xy(1,2) by auto
       
   294     moreover have "?g x \<noteq> pathstart g2"
       
   295       using as(1)
       
   296       unfolding pathstart_def joinpaths_def
       
   297       using inj(2)[of "2 *\<^sub>R x - 1" 0] and xy(1)
       
   298       by (auto simp add: field_simps)
       
   299     ultimately have *: "?g y = pathstart g1"
       
   300       using assms(4)
       
   301       unfolding xy(3)
       
   302       by auto
       
   303     then have "y = 0"
       
   304       unfolding pathstart_def joinpaths_def
       
   305       using as(2) and xy(2)
       
   306       using inj(1)[of "2 *\<^sub>R y" 0]
       
   307       by auto
       
   308     moreover have "x = 1"
       
   309       using *
       
   310       unfolding xy(3)[symmetric] assms(3)[symmetric]
       
   311       unfolding joinpaths_def pathfinish_def using as(1) and xy(1)
       
   312       using inj(2)[of "2 *\<^sub>R x - 1" 1]
       
   313       by auto
       
   314     ultimately show ?thesis
       
   315       by auto
       
   316   qed
       
   317 qed
       
   318 
       
   319 lemma injective_path_join:
       
   320   assumes "injective_path g1"
       
   321     and "injective_path g2"
       
   322     and "pathfinish g1 = pathstart g2"
       
   323     and "path_image g1 \<inter> path_image g2 \<subseteq> {pathstart g2}"
       
   324   shows "injective_path (g1 +++ g2)"
       
   325   unfolding injective_path_def
       
   326 proof (rule, rule, rule)
       
   327   let ?g = "g1 +++ g2"
       
   328   note inj = assms(1,2)[unfolded injective_path_def, rule_format]
       
   329   fix x y
       
   330   assume xy: "x \<in> {0..1}" "y \<in> {0..1}" "(g1 +++ g2) x = (g1 +++ g2) y"
       
   331   show "x = y"
       
   332   proof (cases "x \<le> 1/2", case_tac[!] "y \<le> 1/2", unfold not_le)
       
   333     assume "x \<le> 1 / 2" and "y \<le> 1 / 2"
       
   334     then show ?thesis
       
   335       using inj(1)[of "2*\<^sub>R x" "2*\<^sub>R y"] and xy
       
   336       unfolding joinpaths_def by auto
       
   337   next
       
   338     assume "x > 1 / 2" and "y > 1 / 2"
       
   339     then show ?thesis
       
   340       using inj(2)[of "2*\<^sub>R x - 1" "2*\<^sub>R y - 1"] and xy
       
   341       unfolding joinpaths_def by auto
       
   342   next
       
   343     assume as: "x \<le> 1 / 2" "y > 1 / 2"
       
   344     then have "?g x \<in> path_image g1" and "?g y \<in> path_image g2"
       
   345       unfolding path_image_def joinpaths_def
       
   346       using xy(1,2)
       
   347       by auto
       
   348     then have "?g x = pathfinish g1" and "?g y = pathstart g2"
       
   349       using assms(4)
       
   350       unfolding assms(3) xy(3)
       
   351       by auto
       
   352     then show ?thesis
       
   353       using as and inj(1)[of "2 *\<^sub>R x" 1] inj(2)[of "2 *\<^sub>R y - 1" 0] and xy(1,2)
       
   354       unfolding pathstart_def pathfinish_def joinpaths_def
       
   355       by auto
       
   356   next
       
   357     assume as:"x > 1 / 2" "y \<le> 1 / 2"
       
   358     then have "?g x \<in> path_image g2" and "?g y \<in> path_image g1"
       
   359       unfolding path_image_def joinpaths_def
       
   360       using xy(1,2)
       
   361       by auto
       
   362     then have "?g x = pathstart g2" and "?g y = pathfinish g1"
       
   363       using assms(4)
       
   364       unfolding assms(3) xy(3)
       
   365       by auto
       
   366     then show ?thesis using as and inj(2)[of "2 *\<^sub>R x - 1" 0] inj(1)[of "2 *\<^sub>R y" 1] and xy(1,2)
       
   367       unfolding pathstart_def pathfinish_def joinpaths_def
       
   368       by auto
       
   369   qed
       
   370 qed
       
   371 
       
   372 lemmas join_paths_simps = path_join path_image_join pathstart_join pathfinish_join
       
   373 
   803 
   374 
   804 
   375 subsection {* Reparametrizing a closed curve to start at some chosen point *}
   805 subsection {* Reparametrizing a closed curve to start at some chosen point *}
   376 
   806 
   377 definition shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
   807 definition shiftpath :: "real \<Rightarrow> (real \<Rightarrow> 'a::topological_space) \<Rightarrow> real \<Rightarrow> 'a"
   498   unfolding path_def
   928   unfolding path_def
   499   by (rule continuous_on_linepath)
   929   by (rule continuous_on_linepath)
   500 
   930 
   501 lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b"
   931 lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b"
   502   unfolding path_image_def segment linepath_def
   932   unfolding path_image_def segment linepath_def
   503   apply (rule set_eqI)
   933   by auto
   504   apply rule
       
   505   defer
       
   506   unfolding mem_Collect_eq image_iff
       
   507   apply (erule exE)
       
   508   apply (rule_tac x="u *\<^sub>R 1" in bexI)
       
   509   apply auto
       
   510   done
       
   511 
   934 
   512 lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a"
   935 lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a"
   513   unfolding reversepath_def linepath_def
   936   unfolding reversepath_def linepath_def
   514   by auto
   937   by auto
   515 
   938 
   516 lemma injective_path_linepath:
   939 lemma arc_linepath:
   517   assumes "a \<noteq> b"
   940   assumes "a \<noteq> b"
   518   shows "injective_path (linepath a b)"
   941   shows "arc (linepath a b)"
   519 proof -
   942 proof -
   520   {
   943   {
   521     fix x y :: "real"
   944     fix x y :: "real"
   522     assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b"
   945     assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b"
   523     then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b"
   946     then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b"
   524       by (simp add: algebra_simps)
   947       by (simp add: algebra_simps)
   525     with assms have "x = y"
   948     with assms have "x = y"
   526       by simp
   949       by simp
   527   }
   950   }
   528   then show ?thesis
   951   then show ?thesis
   529     unfolding injective_path_def linepath_def
   952     unfolding arc_def inj_on_def 
   530     by (auto simp add: algebra_simps)
   953     by (simp add:  path_linepath) (force simp: algebra_simps linepath_def)
   531 qed
   954 qed
   532 
   955 
   533 lemma simple_path_linepath[intro]: "a \<noteq> b \<Longrightarrow> simple_path (linepath a b)"
   956 lemma simple_path_linepath[intro]: "a \<noteq> b \<Longrightarrow> simple_path (linepath a b)"
   534   by (auto intro!: injective_imp_simple_path injective_path_linepath)
   957   by (simp add: arc_imp_simple_path arc_linepath)
   535 
   958 
   536 
   959 
   537 subsection {* Bounding a point away from a path *}
   960 subsection {* Bounding a point away from a path *}
   538 
   961 
   539 lemma not_on_path_ball:
   962 lemma not_on_path_ball:
   621 text {* Can also consider it as a set, as the name suggests. *}
  1044 text {* Can also consider it as a set, as the name suggests. *}
   622 
  1045 
   623 lemma path_component_set:
  1046 lemma path_component_set:
   624   "{y. path_component s x y} =
  1047   "{y. path_component s x y} =
   625     {y. (\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)}"
  1048     {y. (\<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)}"
   626   apply (rule set_eqI)
  1049   unfolding mem_Collect_eq path_component_def
   627   unfolding mem_Collect_eq
       
   628   unfolding path_component_def
       
   629   apply auto
  1050   apply auto
   630   done
  1051   done
   631 
  1052 
   632 lemma path_component_subset: "{y. path_component s x y} \<subseteq> s"
  1053 lemma path_component_subset: "{y. path_component s x y} \<subseteq> s"
   633   apply rule
  1054   by (auto simp add: path_component_mem(2))
   634   apply (rule path_component_mem(2))
       
   635   apply auto
       
   636   done
       
   637 
  1055 
   638 lemma path_component_eq_empty: "{y. path_component s x y} = {} \<longleftrightarrow> x \<notin> s"
  1056 lemma path_component_eq_empty: "{y. path_component s x y} = {} \<longleftrightarrow> x \<notin> s"
   639   apply rule
  1057   using path_component_mem path_component_refl_eq
   640   apply (drule equals0D[of _ x])
  1058     by fastforce
   641   defer
       
   642   apply (rule equals0I)
       
   643   unfolding mem_Collect_eq
       
   644   apply (drule path_component_mem(1))
       
   645   using path_component_refl
       
   646   apply auto
       
   647   done
       
   648 
       
   649 
  1059 
   650 subsection {* Path connectedness of a space *}
  1060 subsection {* Path connectedness of a space *}
   651 
  1061 
   652 definition "path_connected s \<longleftrightarrow>
  1062 definition "path_connected s \<longleftrightarrow>
   653   (\<forall>x\<in>s. \<forall>y\<in>s. \<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
  1063   (\<forall>x\<in>s. \<forall>y\<in>s. \<exists>g. path g \<and> path_image g \<subseteq> s \<and> pathstart g = x \<and> pathfinish g = y)"
   654 
  1064 
   655 lemma path_connected_component: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. path_component s x y)"
  1065 lemma path_connected_component: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. \<forall>y\<in>s. path_component s x y)"
   656   unfolding path_connected_def path_component_def by auto
  1066   unfolding path_connected_def path_component_def by auto
   657 
  1067 
   658 lemma path_connected_component_set: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. {y. path_component s x y} = s)"
  1068 lemma path_connected_component_set: "path_connected s \<longleftrightarrow> (\<forall>x\<in>s. {y. path_component s x y} = s)"
   659   unfolding path_connected_component
  1069   unfolding path_connected_component path_component_subset 
   660   apply rule
       
   661   apply rule
       
   662   apply rule
       
   663   apply (rule path_component_subset)
       
   664   unfolding subset_eq mem_Collect_eq Ball_def
       
   665   apply auto
  1070   apply auto
   666   done
  1071   using path_component_mem(2) by blast
   667 
       
   668 
  1072 
   669 subsection {* Some useful lemmas about path-connectedness *}
  1073 subsection {* Some useful lemmas about path-connectedness *}
   670 
  1074 
   671 lemma convex_imp_path_connected:
  1075 lemma convex_imp_path_connected:
   672   fixes s :: "'a::real_normed_vector set"
  1076   fixes s :: "'a::real_normed_vector set"