3384 using \<open>a \<in> S\<close> \<open>0 < r\<close> |
3382 using \<open>a \<in> S\<close> \<open>0 < r\<close> |
3385 apply (auto simp: disjoint_iff_not_equal dist_norm) |
3383 apply (auto simp: disjoint_iff_not_equal dist_norm) |
3386 by (metis dw_le norm_minus_commute not_less order_trans rle wy) |
3384 by (metis dw_le norm_minus_commute not_less order_trans rle wy) |
3387 qed |
3385 qed |
3388 |
3386 |
3389 |
|
3390 section \<open>Homotopy of Maps\<close> (* TODO separate theory? *) |
|
3391 |
|
3392 |
|
3393 definition%important homotopic_with :: |
|
3394 "[('a::topological_space \<Rightarrow> 'b::topological_space) \<Rightarrow> bool, 'a set, 'b set, 'a \<Rightarrow> 'b, 'a \<Rightarrow> 'b] \<Rightarrow> bool" |
|
3395 where |
|
3396 "homotopic_with P X Y p q \<equiv> |
|
3397 (\<exists>h:: real \<times> 'a \<Rightarrow> 'b. |
|
3398 continuous_on ({0..1} \<times> X) h \<and> |
|
3399 h ` ({0..1} \<times> X) \<subseteq> Y \<and> |
|
3400 (\<forall>x. h(0, x) = p x) \<and> |
|
3401 (\<forall>x. h(1, x) = q x) \<and> |
|
3402 (\<forall>t \<in> {0..1}. P(\<lambda>x. h(t, x))))" |
|
3403 |
|
3404 text\<open>\<open>p\<close>, \<open>q\<close> are functions \<open>X \<rightarrow> Y\<close>, and the property \<open>P\<close> restricts all intermediate maps. |
|
3405 We often just want to require that \<open>P\<close> fixes some subset, but to include the case of a loop homotopy, |
|
3406 it is convenient to have a general property \<open>P\<close>.\<close> |
|
3407 |
|
3408 text \<open>We often want to just localize the ending function equality or whatever.\<close> |
|
3409 text%important \<open>%whitespace\<close> |
|
3410 proposition homotopic_with: |
|
3411 fixes X :: "'a::topological_space set" and Y :: "'b::topological_space set" |
|
3412 assumes "\<And>h k. (\<And>x. x \<in> X \<Longrightarrow> h x = k x) \<Longrightarrow> (P h \<longleftrightarrow> P k)" |
|
3413 shows "homotopic_with P X Y p q \<longleftrightarrow> |
|
3414 (\<exists>h :: real \<times> 'a \<Rightarrow> 'b. |
|
3415 continuous_on ({0..1} \<times> X) h \<and> |
|
3416 h ` ({0..1} \<times> X) \<subseteq> Y \<and> |
|
3417 (\<forall>x \<in> X. h(0,x) = p x) \<and> |
|
3418 (\<forall>x \<in> X. h(1,x) = q x) \<and> |
|
3419 (\<forall>t \<in> {0..1}. P(\<lambda>x. h(t, x))))" |
|
3420 unfolding homotopic_with_def |
|
3421 apply (rule iffI, blast, clarify) |
|
3422 apply (rule_tac x="\<lambda>(u,v). if v \<in> X then h(u,v) else if u = 0 then p v else q v" in exI) |
|
3423 apply auto |
|
3424 apply (force elim: continuous_on_eq) |
|
3425 apply (drule_tac x=t in bspec, force) |
|
3426 apply (subst assms; simp) |
|
3427 done |
|
3428 |
|
3429 proposition homotopic_with_eq: |
|
3430 assumes h: "homotopic_with P X Y f g" |
|
3431 and f': "\<And>x. x \<in> X \<Longrightarrow> f' x = f x" |
|
3432 and g': "\<And>x. x \<in> X \<Longrightarrow> g' x = g x" |
|
3433 and P: "(\<And>h k. (\<And>x. x \<in> X \<Longrightarrow> h x = k x) \<Longrightarrow> (P h \<longleftrightarrow> P k))" |
|
3434 shows "homotopic_with P X Y f' g'" |
|
3435 using h unfolding homotopic_with_def |
|
3436 apply safe |
|
3437 apply (rule_tac x="\<lambda>(u,v). if v \<in> X then h(u,v) else if u = 0 then f' v else g' v" in exI) |
|
3438 apply (simp add: f' g', safe) |
|
3439 apply (fastforce intro: continuous_on_eq, fastforce) |
|
3440 apply (subst P; fastforce) |
|
3441 done |
|
3442 |
|
3443 proposition homotopic_with_equal: |
|
3444 assumes contf: "continuous_on X f" and fXY: "f ` X \<subseteq> Y" |
|
3445 and gf: "\<And>x. x \<in> X \<Longrightarrow> g x = f x" |
|
3446 and P: "P f" "P g" |
|
3447 shows "homotopic_with P X Y f g" |
|
3448 unfolding homotopic_with_def |
|
3449 apply (rule_tac x="\<lambda>(u,v). if u = 1 then g v else f v" in exI) |
|
3450 using assms |
|
3451 apply (intro conjI) |
|
3452 apply (rule continuous_on_eq [where f = "f \<circ> snd"]) |
|
3453 apply (rule continuous_intros | force)+ |
|
3454 apply clarify |
|
3455 apply (case_tac "t=1"; force) |
|
3456 done |
|
3457 |
|
3458 |
|
3459 lemma image_Pair_const: "(\<lambda>x. (x, c)) ` A = A \<times> {c}" |
|
3460 by auto |
|
3461 |
|
3462 lemma homotopic_constant_maps: |
|
3463 "homotopic_with (\<lambda>x. True) s t (\<lambda>x. a) (\<lambda>x. b) \<longleftrightarrow> s = {} \<or> path_component t a b" |
|
3464 proof (cases "s = {} \<or> t = {}") |
|
3465 case True with continuous_on_const show ?thesis |
|
3466 by (auto simp: homotopic_with path_component_def) |
|
3467 next |
|
3468 case False |
|
3469 then obtain c where "c \<in> s" by blast |
|
3470 show ?thesis |
|
3471 proof |
|
3472 assume "homotopic_with (\<lambda>x. True) s t (\<lambda>x. a) (\<lambda>x. b)" |
|
3473 then obtain h :: "real \<times> 'a \<Rightarrow> 'b" |
|
3474 where conth: "continuous_on ({0..1} \<times> s) h" |
|
3475 and h: "h ` ({0..1} \<times> s) \<subseteq> t" "(\<forall>x\<in>s. h (0, x) = a)" "(\<forall>x\<in>s. h (1, x) = b)" |
|
3476 by (auto simp: homotopic_with) |
|
3477 have "continuous_on {0..1} (h \<circ> (\<lambda>t. (t, c)))" |
|
3478 apply (rule continuous_intros conth | simp add: image_Pair_const)+ |
|
3479 apply (blast intro: \<open>c \<in> s\<close> continuous_on_subset [OF conth]) |
|
3480 done |
|
3481 with \<open>c \<in> s\<close> h show "s = {} \<or> path_component t a b" |
|
3482 apply (simp_all add: homotopic_with path_component_def, auto) |
|
3483 apply (drule_tac x="h \<circ> (\<lambda>t. (t, c))" in spec) |
|
3484 apply (auto simp: pathstart_def pathfinish_def path_image_def path_def) |
|
3485 done |
|
3486 next |
|
3487 assume "s = {} \<or> path_component t a b" |
|
3488 with False show "homotopic_with (\<lambda>x. True) s t (\<lambda>x. a) (\<lambda>x. b)" |
|
3489 apply (clarsimp simp: homotopic_with path_component_def pathstart_def pathfinish_def path_image_def path_def) |
|
3490 apply (rule_tac x="g \<circ> fst" in exI) |
|
3491 apply (rule conjI continuous_intros | force)+ |
|
3492 done |
|
3493 qed |
|
3494 qed |
|
3495 |
|
3496 |
|
3497 subsection%unimportant\<open>Trivial properties\<close> |
|
3498 |
|
3499 lemma homotopic_with_imp_property: "homotopic_with P X Y f g \<Longrightarrow> P f \<and> P g" |
|
3500 unfolding homotopic_with_def Ball_def |
|
3501 apply clarify |
|
3502 apply (frule_tac x=0 in spec) |
|
3503 apply (drule_tac x=1 in spec, auto) |
|
3504 done |
|
3505 |
|
3506 lemma continuous_on_o_Pair: "\<lbrakk>continuous_on (T \<times> X) h; t \<in> T\<rbrakk> \<Longrightarrow> continuous_on X (h \<circ> Pair t)" |
|
3507 by (fast intro: continuous_intros elim!: continuous_on_subset) |
|
3508 |
|
3509 lemma homotopic_with_imp_continuous: |
|
3510 assumes "homotopic_with P X Y f g" |
|
3511 shows "continuous_on X f \<and> continuous_on X g" |
|
3512 proof - |
|
3513 obtain h :: "real \<times> 'a \<Rightarrow> 'b" |
|
3514 where conth: "continuous_on ({0..1} \<times> X) h" |
|
3515 and h: "\<forall>x. h (0, x) = f x" "\<forall>x. h (1, x) = g x" |
|
3516 using assms by (auto simp: homotopic_with_def) |
|
3517 have *: "t \<in> {0..1} \<Longrightarrow> continuous_on X (h \<circ> (\<lambda>x. (t,x)))" for t |
|
3518 by (rule continuous_intros continuous_on_subset [OF conth] | force)+ |
|
3519 show ?thesis |
|
3520 using h *[of 0] *[of 1] by auto |
|
3521 qed |
|
3522 |
|
3523 proposition homotopic_with_imp_subset1: |
|
3524 "homotopic_with P X Y f g \<Longrightarrow> f ` X \<subseteq> Y" |
|
3525 by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) |
|
3526 |
|
3527 proposition homotopic_with_imp_subset2: |
|
3528 "homotopic_with P X Y f g \<Longrightarrow> g ` X \<subseteq> Y" |
|
3529 by (simp add: homotopic_with_def image_subset_iff) (metis atLeastAtMost_iff order_refl zero_le_one) |
|
3530 |
|
3531 proposition homotopic_with_mono: |
|
3532 assumes hom: "homotopic_with P X Y f g" |
|
3533 and Q: "\<And>h. \<lbrakk>continuous_on X h; image h X \<subseteq> Y \<and> P h\<rbrakk> \<Longrightarrow> Q h" |
|
3534 shows "homotopic_with Q X Y f g" |
|
3535 using hom |
|
3536 apply (simp add: homotopic_with_def) |
|
3537 apply (erule ex_forward) |
|
3538 apply (force simp: intro!: Q dest: continuous_on_o_Pair) |
|
3539 done |
|
3540 |
|
3541 proposition homotopic_with_subset_left: |
|
3542 "\<lbrakk>homotopic_with P X Y f g; Z \<subseteq> X\<rbrakk> \<Longrightarrow> homotopic_with P Z Y f g" |
|
3543 apply (simp add: homotopic_with_def) |
|
3544 apply (fast elim!: continuous_on_subset ex_forward) |
|
3545 done |
|
3546 |
|
3547 proposition homotopic_with_subset_right: |
|
3548 "\<lbrakk>homotopic_with P X Y f g; Y \<subseteq> Z\<rbrakk> \<Longrightarrow> homotopic_with P X Z f g" |
|
3549 apply (simp add: homotopic_with_def) |
|
3550 apply (fast elim!: continuous_on_subset ex_forward) |
|
3551 done |
|
3552 |
|
3553 proposition homotopic_with_compose_continuous_right: |
|
3554 "\<lbrakk>homotopic_with (\<lambda>f. p (f \<circ> h)) X Y f g; continuous_on W h; h ` W \<subseteq> X\<rbrakk> |
|
3555 \<Longrightarrow> homotopic_with p W Y (f \<circ> h) (g \<circ> h)" |
|
3556 apply (clarsimp simp add: homotopic_with_def) |
|
3557 apply (rename_tac k) |
|
3558 apply (rule_tac x="k \<circ> (\<lambda>y. (fst y, h (snd y)))" in exI) |
|
3559 apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ |
|
3560 apply (erule continuous_on_subset) |
|
3561 apply (fastforce simp: o_def)+ |
|
3562 done |
|
3563 |
|
3564 proposition homotopic_compose_continuous_right: |
|
3565 "\<lbrakk>homotopic_with (\<lambda>f. True) X Y f g; continuous_on W h; h ` W \<subseteq> X\<rbrakk> |
|
3566 \<Longrightarrow> homotopic_with (\<lambda>f. True) W Y (f \<circ> h) (g \<circ> h)" |
|
3567 using homotopic_with_compose_continuous_right by fastforce |
|
3568 |
|
3569 proposition homotopic_with_compose_continuous_left: |
|
3570 "\<lbrakk>homotopic_with (\<lambda>f. p (h \<circ> f)) X Y f g; continuous_on Y h; h ` Y \<subseteq> Z\<rbrakk> |
|
3571 \<Longrightarrow> homotopic_with p X Z (h \<circ> f) (h \<circ> g)" |
|
3572 apply (clarsimp simp add: homotopic_with_def) |
|
3573 apply (rename_tac k) |
|
3574 apply (rule_tac x="h \<circ> k" in exI) |
|
3575 apply (rule conjI continuous_intros continuous_on_compose [where f=snd and g=h, unfolded o_def] | simp)+ |
|
3576 apply (erule continuous_on_subset) |
|
3577 apply (fastforce simp: o_def)+ |
|
3578 done |
|
3579 |
|
3580 proposition homotopic_compose_continuous_left: |
|
3581 "\<lbrakk>homotopic_with (\<lambda>_. True) X Y f g; |
|
3582 continuous_on Y h; h ` Y \<subseteq> Z\<rbrakk> |
|
3583 \<Longrightarrow> homotopic_with (\<lambda>f. True) X Z (h \<circ> f) (h \<circ> g)" |
|
3584 using homotopic_with_compose_continuous_left by fastforce |
|
3585 |
|
3586 proposition homotopic_with_Pair: |
|
3587 assumes hom: "homotopic_with p s t f g" "homotopic_with p' s' t' f' g'" |
|
3588 and q: "\<And>f g. \<lbrakk>p f; p' g\<rbrakk> \<Longrightarrow> q(\<lambda>(x,y). (f x, g y))" |
|
3589 shows "homotopic_with q (s \<times> s') (t \<times> t') |
|
3590 (\<lambda>(x,y). (f x, f' y)) (\<lambda>(x,y). (g x, g' y))" |
|
3591 using hom |
|
3592 apply (clarsimp simp add: homotopic_with_def) |
|
3593 apply (rename_tac k k') |
|
3594 apply (rule_tac x="\<lambda>z. ((k \<circ> (\<lambda>x. (fst x, fst (snd x)))) z, (k' \<circ> (\<lambda>x. (fst x, snd (snd x)))) z)" in exI) |
|
3595 apply (rule conjI continuous_intros | erule continuous_on_subset | clarsimp)+ |
|
3596 apply (auto intro!: q [unfolded case_prod_unfold]) |
|
3597 done |
|
3598 |
|
3599 lemma homotopic_on_empty [simp]: "homotopic_with (\<lambda>x. True) {} t f g" |
|
3600 by (metis continuous_on_def empty_iff homotopic_with_equal image_subset_iff) |
|
3601 |
|
3602 |
|
3603 text\<open>Homotopy with P is an equivalence relation (on continuous functions mapping X into Y that satisfy P, |
|
3604 though this only affects reflexivity.\<close> |
|
3605 |
|
3606 |
|
3607 proposition homotopic_with_refl: |
|
3608 "homotopic_with P X Y f f \<longleftrightarrow> continuous_on X f \<and> image f X \<subseteq> Y \<and> P f" |
|
3609 apply (rule iffI) |
|
3610 using homotopic_with_imp_continuous homotopic_with_imp_property homotopic_with_imp_subset2 apply blast |
|
3611 apply (simp add: homotopic_with_def) |
|
3612 apply (rule_tac x="f \<circ> snd" in exI) |
|
3613 apply (rule conjI continuous_intros | force)+ |
|
3614 done |
|
3615 |
|
3616 lemma homotopic_with_symD: |
|
3617 fixes X :: "'a::real_normed_vector set" |
|
3618 assumes "homotopic_with P X Y f g" |
|
3619 shows "homotopic_with P X Y g f" |
|
3620 using assms |
|
3621 apply (clarsimp simp add: homotopic_with_def) |
|
3622 apply (rename_tac h) |
|
3623 apply (rule_tac x="h \<circ> (\<lambda>y. (1 - fst y, snd y))" in exI) |
|
3624 apply (rule conjI continuous_intros | erule continuous_on_subset | force simp: image_subset_iff)+ |
|
3625 done |
|
3626 |
|
3627 proposition homotopic_with_sym: |
|
3628 fixes X :: "'a::real_normed_vector set" |
|
3629 shows "homotopic_with P X Y f g \<longleftrightarrow> homotopic_with P X Y g f" |
|
3630 using homotopic_with_symD by blast |
|
3631 |
|
3632 lemma split_01: "{0..1::real} = {0..1/2} \<union> {1/2..1}" |
|
3633 by force |
|
3634 |
|
3635 lemma split_01_prod: "{0..1::real} \<times> X = ({0..1/2} \<times> X) \<union> ({1/2..1} \<times> X)" |
|
3636 by force |
|
3637 |
|
3638 proposition homotopic_with_trans: |
|
3639 fixes X :: "'a::real_normed_vector set" |
|
3640 assumes "homotopic_with P X Y f g" and "homotopic_with P X Y g h" |
|
3641 shows "homotopic_with P X Y f h" |
|
3642 proof - |
|
3643 have clo1: "closedin (subtopology euclidean ({0..1/2} \<times> X \<union> {1/2..1} \<times> X)) ({0..1/2::real} \<times> X)" |
|
3644 apply (simp add: closedin_closed split_01_prod [symmetric]) |
|
3645 apply (rule_tac x="{0..1/2} \<times> UNIV" in exI) |
|
3646 apply (force simp: closed_Times) |
|
3647 done |
|
3648 have clo2: "closedin (subtopology euclidean ({0..1/2} \<times> X \<union> {1/2..1} \<times> X)) ({1/2..1::real} \<times> X)" |
|
3649 apply (simp add: closedin_closed split_01_prod [symmetric]) |
|
3650 apply (rule_tac x="{1/2..1} \<times> UNIV" in exI) |
|
3651 apply (force simp: closed_Times) |
|
3652 done |
|
3653 { fix k1 k2:: "real \<times> 'a \<Rightarrow> 'b" |
|
3654 assume cont: "continuous_on ({0..1} \<times> X) k1" "continuous_on ({0..1} \<times> X) k2" |
|
3655 and Y: "k1 ` ({0..1} \<times> X) \<subseteq> Y" "k2 ` ({0..1} \<times> X) \<subseteq> Y" |
|
3656 and geq: "\<forall>x. k1 (1, x) = g x" "\<forall>x. k2 (0, x) = g x" |
|
3657 and k12: "\<forall>x. k1 (0, x) = f x" "\<forall>x. k2 (1, x) = h x" |
|
3658 and P: "\<forall>t\<in>{0..1}. P (\<lambda>x. k1 (t, x))" "\<forall>t\<in>{0..1}. P (\<lambda>x. k2 (t, x))" |
|
3659 define k where "k y = |
|
3660 (if fst y \<le> 1 / 2 |
|
3661 then (k1 \<circ> (\<lambda>x. (2 *\<^sub>R fst x, snd x))) y |
|
3662 else (k2 \<circ> (\<lambda>x. (2 *\<^sub>R fst x -1, snd x))) y)" for y |
|
3663 have keq: "k1 (2 * u, v) = k2 (2 * u - 1, v)" if "u = 1/2" for u v |
|
3664 by (simp add: geq that) |
|
3665 have "continuous_on ({0..1} \<times> X) k" |
|
3666 using cont |
|
3667 apply (simp add: split_01_prod k_def) |
|
3668 apply (rule clo1 clo2 continuous_on_cases_local continuous_intros | erule continuous_on_subset | simp add: linear image_subset_iff)+ |
|
3669 apply (force simp: keq) |
|
3670 done |
|
3671 moreover have "k ` ({0..1} \<times> X) \<subseteq> Y" |
|
3672 using Y by (force simp: k_def) |
|
3673 moreover have "\<forall>x. k (0, x) = f x" |
|
3674 by (simp add: k_def k12) |
|
3675 moreover have "(\<forall>x. k (1, x) = h x)" |
|
3676 by (simp add: k_def k12) |
|
3677 moreover have "\<forall>t\<in>{0..1}. P (\<lambda>x. k (t, x))" |
|
3678 using P |
|
3679 apply (clarsimp simp add: k_def) |
|
3680 apply (case_tac "t \<le> 1/2", auto) |
|
3681 done |
|
3682 ultimately have *: "\<exists>k :: real \<times> 'a \<Rightarrow> 'b. |
|
3683 continuous_on ({0..1} \<times> X) k \<and> k ` ({0..1} \<times> X) \<subseteq> Y \<and> |
|
3684 (\<forall>x. k (0, x) = f x) \<and> (\<forall>x. k (1, x) = h x) \<and> (\<forall>t\<in>{0..1}. P (\<lambda>x. k (t, x)))" |
|
3685 by blast |
|
3686 } note * = this |
|
3687 show ?thesis |
|
3688 using assms by (auto intro: * simp add: homotopic_with_def) |
|
3689 qed |
|
3690 |
|
3691 proposition homotopic_compose: |
|
3692 fixes s :: "'a::real_normed_vector set" |
|
3693 shows "\<lbrakk>homotopic_with (\<lambda>x. True) s t f f'; homotopic_with (\<lambda>x. True) t u g g'\<rbrakk> |
|
3694 \<Longrightarrow> homotopic_with (\<lambda>x. True) s u (g \<circ> f) (g' \<circ> f')" |
|
3695 apply (rule homotopic_with_trans [where g = "g \<circ> f'"]) |
|
3696 apply (metis homotopic_compose_continuous_left homotopic_with_imp_continuous homotopic_with_imp_subset1) |
|
3697 by (metis homotopic_compose_continuous_right homotopic_with_imp_continuous homotopic_with_imp_subset2) |
|
3698 |
|
3699 |
|
3700 text\<open>Homotopic triviality implicitly incorporates path-connectedness.\<close> |
|
3701 lemma homotopic_triviality: |
|
3702 fixes S :: "'a::real_normed_vector set" |
|
3703 shows "(\<forall>f g. continuous_on S f \<and> f ` S \<subseteq> T \<and> |
|
3704 continuous_on S g \<and> g ` S \<subseteq> T |
|
3705 \<longrightarrow> homotopic_with (\<lambda>x. True) S T f g) \<longleftrightarrow> |
|
3706 (S = {} \<or> path_connected T) \<and> |
|
3707 (\<forall>f. continuous_on S f \<and> f ` S \<subseteq> T \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)))" |
|
3708 (is "?lhs = ?rhs") |
|
3709 proof (cases "S = {} \<or> T = {}") |
|
3710 case True then show ?thesis by auto |
|
3711 next |
|
3712 case False show ?thesis |
|
3713 proof |
|
3714 assume LHS [rule_format]: ?lhs |
|
3715 have pab: "path_component T a b" if "a \<in> T" "b \<in> T" for a b |
|
3716 proof - |
|
3717 have "homotopic_with (\<lambda>x. True) S T (\<lambda>x. a) (\<lambda>x. b)" |
|
3718 by (simp add: LHS continuous_on_const image_subset_iff that) |
|
3719 then show ?thesis |
|
3720 using False homotopic_constant_maps by blast |
|
3721 qed |
|
3722 moreover |
|
3723 have "\<exists>c. homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)" if "continuous_on S f" "f ` S \<subseteq> T" for f |
|
3724 by (metis (full_types) False LHS equals0I homotopic_constant_maps homotopic_with_imp_continuous homotopic_with_imp_subset2 pab that) |
|
3725 ultimately show ?rhs |
|
3726 by (simp add: path_connected_component) |
|
3727 next |
|
3728 assume RHS: ?rhs |
|
3729 with False have T: "path_connected T" |
|
3730 by blast |
|
3731 show ?lhs |
|
3732 proof clarify |
|
3733 fix f g |
|
3734 assume "continuous_on S f" "f ` S \<subseteq> T" "continuous_on S g" "g ` S \<subseteq> T" |
|
3735 obtain c d where c: "homotopic_with (\<lambda>x. True) S T f (\<lambda>x. c)" and d: "homotopic_with (\<lambda>x. True) S T g (\<lambda>x. d)" |
|
3736 using False \<open>continuous_on S f\<close> \<open>f ` S \<subseteq> T\<close> RHS \<open>continuous_on S g\<close> \<open>g ` S \<subseteq> T\<close> by blast |
|
3737 then have "c \<in> T" "d \<in> T" |
|
3738 using False homotopic_with_imp_subset2 by fastforce+ |
|
3739 with T have "path_component T c d" |
|
3740 using path_connected_component by blast |
|
3741 then have "homotopic_with (\<lambda>x. True) S T (\<lambda>x. c) (\<lambda>x. d)" |
|
3742 by (simp add: homotopic_constant_maps) |
|
3743 with c d show "homotopic_with (\<lambda>x. True) S T f g" |
|
3744 by (meson homotopic_with_symD homotopic_with_trans) |
|
3745 qed |
|
3746 qed |
|
3747 qed |
|
3748 |
|
3749 |
|
3750 subsection\<open>Homotopy of paths, maintaining the same endpoints\<close> |
|
3751 |
|
3752 |
|
3753 definition%important homotopic_paths :: "['a set, real \<Rightarrow> 'a, real \<Rightarrow> 'a::topological_space] \<Rightarrow> bool" |
|
3754 where |
|
3755 "homotopic_paths s p q \<equiv> |
|
3756 homotopic_with (\<lambda>r. pathstart r = pathstart p \<and> pathfinish r = pathfinish p) {0..1} s p q" |
|
3757 |
|
3758 lemma homotopic_paths: |
|
3759 "homotopic_paths s p q \<longleftrightarrow> |
|
3760 (\<exists>h. continuous_on ({0..1} \<times> {0..1}) h \<and> |
|
3761 h ` ({0..1} \<times> {0..1}) \<subseteq> s \<and> |
|
3762 (\<forall>x \<in> {0..1}. h(0,x) = p x) \<and> |
|
3763 (\<forall>x \<in> {0..1}. h(1,x) = q x) \<and> |
|
3764 (\<forall>t \<in> {0..1::real}. pathstart(h \<circ> Pair t) = pathstart p \<and> |
|
3765 pathfinish(h \<circ> Pair t) = pathfinish p))" |
|
3766 by (auto simp: homotopic_paths_def homotopic_with pathstart_def pathfinish_def) |
|
3767 |
|
3768 proposition homotopic_paths_imp_pathstart: |
|
3769 "homotopic_paths s p q \<Longrightarrow> pathstart p = pathstart q" |
|
3770 by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) |
|
3771 |
|
3772 proposition homotopic_paths_imp_pathfinish: |
|
3773 "homotopic_paths s p q \<Longrightarrow> pathfinish p = pathfinish q" |
|
3774 by (metis (mono_tags, lifting) homotopic_paths_def homotopic_with_imp_property) |
|
3775 |
|
3776 lemma homotopic_paths_imp_path: |
|
3777 "homotopic_paths s p q \<Longrightarrow> path p \<and> path q" |
|
3778 using homotopic_paths_def homotopic_with_imp_continuous path_def by blast |
|
3779 |
|
3780 lemma homotopic_paths_imp_subset: |
|
3781 "homotopic_paths s p q \<Longrightarrow> path_image p \<subseteq> s \<and> path_image q \<subseteq> s" |
|
3782 by (simp add: homotopic_paths_def homotopic_with_imp_subset1 homotopic_with_imp_subset2 path_image_def) |
|
3783 |
|
3784 proposition homotopic_paths_refl [simp]: "homotopic_paths s p p \<longleftrightarrow> path p \<and> path_image p \<subseteq> s" |
|
3785 by (simp add: homotopic_paths_def homotopic_with_refl path_def path_image_def) |
|
3786 |
|
3787 proposition homotopic_paths_sym: "homotopic_paths s p q \<Longrightarrow> homotopic_paths s q p" |
|
3788 by (metis (mono_tags) homotopic_paths_def homotopic_paths_imp_pathfinish homotopic_paths_imp_pathstart homotopic_with_symD) |
|
3789 |
|
3790 proposition homotopic_paths_sym_eq: "homotopic_paths s p q \<longleftrightarrow> homotopic_paths s q p" |
|
3791 by (metis homotopic_paths_sym) |
|
3792 |
|
3793 proposition homotopic_paths_trans [trans]: |
|
3794 "\<lbrakk>homotopic_paths s p q; homotopic_paths s q r\<rbrakk> \<Longrightarrow> homotopic_paths s p r" |
|
3795 apply (simp add: homotopic_paths_def) |
|
3796 apply (rule homotopic_with_trans, assumption) |
|
3797 by (metis (mono_tags, lifting) homotopic_with_imp_property homotopic_with_mono) |
|
3798 |
|
3799 proposition homotopic_paths_eq: |
|
3800 "\<lbrakk>path p; path_image p \<subseteq> s; \<And>t. t \<in> {0..1} \<Longrightarrow> p t = q t\<rbrakk> \<Longrightarrow> homotopic_paths s p q" |
|
3801 apply (simp add: homotopic_paths_def) |
|
3802 apply (rule homotopic_with_eq) |
|
3803 apply (auto simp: path_def homotopic_with_refl pathstart_def pathfinish_def path_image_def elim: continuous_on_eq) |
|
3804 done |
|
3805 |
|
3806 proposition homotopic_paths_reparametrize: |
|
3807 assumes "path p" |
|
3808 and pips: "path_image p \<subseteq> s" |
|
3809 and contf: "continuous_on {0..1} f" |
|
3810 and f01:"f ` {0..1} \<subseteq> {0..1}" |
|
3811 and [simp]: "f(0) = 0" "f(1) = 1" |
|
3812 and q: "\<And>t. t \<in> {0..1} \<Longrightarrow> q(t) = p(f t)" |
|
3813 shows "homotopic_paths s p q" |
|
3814 proof - |
|
3815 have contp: "continuous_on {0..1} p" |
|
3816 by (metis \<open>path p\<close> path_def) |
|
3817 then have "continuous_on {0..1} (p \<circ> f)" |
|
3818 using contf continuous_on_compose continuous_on_subset f01 by blast |
|
3819 then have "path q" |
|
3820 by (simp add: path_def) (metis q continuous_on_cong) |
|
3821 have piqs: "path_image q \<subseteq> s" |
|
3822 by (metis (no_types, hide_lams) pips f01 image_subset_iff path_image_def q) |
|
3823 have fb0: "\<And>a b. \<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> 0 \<le> (1 - a) * f b + a * b" |
|
3824 using f01 by force |
|
3825 have fb1: "\<lbrakk>0 \<le> a; a \<le> 1; 0 \<le> b; b \<le> 1\<rbrakk> \<Longrightarrow> (1 - a) * f b + a * b \<le> 1" for a b |
|
3826 using f01 [THEN subsetD, of "f b"] by (simp add: convex_bound_le) |
|
3827 have "homotopic_paths s q p" |
|
3828 proof (rule homotopic_paths_trans) |
|
3829 show "homotopic_paths s q (p \<circ> f)" |
|
3830 using q by (force intro: homotopic_paths_eq [OF \<open>path q\<close> piqs]) |
|
3831 next |
|
3832 show "homotopic_paths s (p \<circ> f) p" |
|
3833 apply (simp add: homotopic_paths_def homotopic_with_def) |
|
3834 apply (rule_tac x="p \<circ> (\<lambda>y. (1 - (fst y)) *\<^sub>R ((f \<circ> snd) y) + (fst y) *\<^sub>R snd y)" in exI) |
|
3835 apply (rule conjI contf continuous_intros continuous_on_subset [OF contp] | simp)+ |
|
3836 using pips [unfolded path_image_def] |
|
3837 apply (auto simp: fb0 fb1 pathstart_def pathfinish_def) |
|
3838 done |
|
3839 qed |
|
3840 then show ?thesis |
|
3841 by (simp add: homotopic_paths_sym) |
|
3842 qed |
|
3843 |
|
3844 lemma homotopic_paths_subset: "\<lbrakk>homotopic_paths s p q; s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_paths t p q" |
|
3845 using homotopic_paths_def homotopic_with_subset_right by blast |
|
3846 |
|
3847 |
|
3848 text\<open> A slightly ad-hoc but useful lemma in constructing homotopies.\<close> |
|
3849 lemma homotopic_join_lemma: |
|
3850 fixes q :: "[real,real] \<Rightarrow> 'a::topological_space" |
|
3851 assumes p: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. p (fst y) (snd y))" |
|
3852 and q: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. q (fst y) (snd y))" |
|
3853 and pf: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish(p t) = pathstart(q t)" |
|
3854 shows "continuous_on ({0..1} \<times> {0..1}) (\<lambda>y. (p(fst y) +++ q(fst y)) (snd y))" |
|
3855 proof - |
|
3856 have 1: "(\<lambda>y. p (fst y) (2 * snd y)) = (\<lambda>y. p (fst y) (snd y)) \<circ> (\<lambda>y. (fst y, 2 * snd y))" |
|
3857 by (rule ext) (simp) |
|
3858 have 2: "(\<lambda>y. q (fst y) (2 * snd y - 1)) = (\<lambda>y. q (fst y) (snd y)) \<circ> (\<lambda>y. (fst y, 2 * snd y - 1))" |
|
3859 by (rule ext) (simp) |
|
3860 show ?thesis |
|
3861 apply (simp add: joinpaths_def) |
|
3862 apply (rule continuous_on_cases_le) |
|
3863 apply (simp_all only: 1 2) |
|
3864 apply (rule continuous_intros continuous_on_subset [OF p] continuous_on_subset [OF q] | force)+ |
|
3865 using pf |
|
3866 apply (auto simp: mult.commute pathstart_def pathfinish_def) |
|
3867 done |
|
3868 qed |
|
3869 |
|
3870 text\<open> Congruence properties of homotopy w.r.t. path-combining operations.\<close> |
|
3871 |
|
3872 lemma homotopic_paths_reversepath_D: |
|
3873 assumes "homotopic_paths s p q" |
|
3874 shows "homotopic_paths s (reversepath p) (reversepath q)" |
|
3875 using assms |
|
3876 apply (simp add: homotopic_paths_def homotopic_with_def, clarify) |
|
3877 apply (rule_tac x="h \<circ> (\<lambda>x. (fst x, 1 - snd x))" in exI) |
|
3878 apply (rule conjI continuous_intros)+ |
|
3879 apply (auto simp: reversepath_def pathstart_def pathfinish_def elim!: continuous_on_subset) |
|
3880 done |
|
3881 |
|
3882 proposition homotopic_paths_reversepath: |
|
3883 "homotopic_paths s (reversepath p) (reversepath q) \<longleftrightarrow> homotopic_paths s p q" |
|
3884 using homotopic_paths_reversepath_D by force |
|
3885 |
|
3886 |
|
3887 proposition homotopic_paths_join: |
|
3888 "\<lbrakk>homotopic_paths s p p'; homotopic_paths s q q'; pathfinish p = pathstart q\<rbrakk> \<Longrightarrow> homotopic_paths s (p +++ q) (p' +++ q')" |
|
3889 apply (simp add: homotopic_paths_def homotopic_with_def, clarify) |
|
3890 apply (rename_tac k1 k2) |
|
3891 apply (rule_tac x="(\<lambda>y. ((k1 \<circ> Pair (fst y)) +++ (k2 \<circ> Pair (fst y))) (snd y))" in exI) |
|
3892 apply (rule conjI continuous_intros homotopic_join_lemma)+ |
|
3893 apply (auto simp: joinpaths_def pathstart_def pathfinish_def path_image_def) |
|
3894 done |
|
3895 |
|
3896 proposition homotopic_paths_continuous_image: |
|
3897 "\<lbrakk>homotopic_paths s f g; continuous_on s h; h ` s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_paths t (h \<circ> f) (h \<circ> g)" |
|
3898 unfolding homotopic_paths_def |
|
3899 apply (rule homotopic_with_compose_continuous_left [of _ _ _ s]) |
|
3900 apply (auto simp: pathstart_def pathfinish_def elim!: homotopic_with_mono) |
|
3901 done |
|
3902 |
|
3903 |
|
3904 subsection\<open>Group properties for homotopy of paths\<close> |
|
3905 |
|
3906 text%important\<open>So taking equivalence classes under homotopy would give the fundamental group\<close> |
|
3907 |
|
3908 proposition homotopic_paths_rid: |
|
3909 "\<lbrakk>path p; path_image p \<subseteq> s\<rbrakk> \<Longrightarrow> homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) p" |
|
3910 apply (subst homotopic_paths_sym) |
|
3911 apply (rule homotopic_paths_reparametrize [where f = "\<lambda>t. if t \<le> 1 / 2 then 2 *\<^sub>R t else 1"]) |
|
3912 apply (simp_all del: le_divide_eq_numeral1) |
|
3913 apply (subst split_01) |
|
3914 apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ |
|
3915 done |
|
3916 |
|
3917 proposition homotopic_paths_lid: |
|
3918 "\<lbrakk>path p; path_image p \<subseteq> s\<rbrakk> \<Longrightarrow> homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p) p" |
|
3919 using homotopic_paths_rid [of "reversepath p" s] |
|
3920 by (metis homotopic_paths_reversepath path_image_reversepath path_reversepath pathfinish_linepath |
|
3921 pathfinish_reversepath reversepath_joinpaths reversepath_linepath) |
|
3922 |
|
3923 proposition homotopic_paths_assoc: |
|
3924 "\<lbrakk>path p; path_image p \<subseteq> s; path q; path_image q \<subseteq> s; path r; path_image r \<subseteq> s; pathfinish p = pathstart q; |
|
3925 pathfinish q = pathstart r\<rbrakk> |
|
3926 \<Longrightarrow> homotopic_paths s (p +++ (q +++ r)) ((p +++ q) +++ r)" |
|
3927 apply (subst homotopic_paths_sym) |
|
3928 apply (rule homotopic_paths_reparametrize |
|
3929 [where f = "\<lambda>t. if t \<le> 1 / 2 then inverse 2 *\<^sub>R t |
|
3930 else if t \<le> 3 / 4 then t - (1 / 4) |
|
3931 else 2 *\<^sub>R t - 1"]) |
|
3932 apply (simp_all del: le_divide_eq_numeral1) |
|
3933 apply (simp add: subset_path_image_join) |
|
3934 apply (rule continuous_on_cases_1 continuous_intros)+ |
|
3935 apply (auto simp: joinpaths_def) |
|
3936 done |
|
3937 |
|
3938 proposition homotopic_paths_rinv: |
|
3939 assumes "path p" "path_image p \<subseteq> s" |
|
3940 shows "homotopic_paths s (p +++ reversepath p) (linepath (pathstart p) (pathstart p))" |
|
3941 proof - |
|
3942 have "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. (subpath 0 (fst x) p +++ reversepath (subpath 0 (fst x) p)) (snd x))" |
|
3943 using assms |
|
3944 apply (simp add: joinpaths_def subpath_def reversepath_def path_def del: le_divide_eq_numeral1) |
|
3945 apply (rule continuous_on_cases_le) |
|
3946 apply (rule_tac [2] continuous_on_compose [of _ _ p, unfolded o_def]) |
|
3947 apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) |
|
3948 apply (auto intro!: continuous_intros simp del: eq_divide_eq_numeral1) |
|
3949 apply (force elim!: continuous_on_subset simp add: mult_le_one)+ |
|
3950 done |
|
3951 then show ?thesis |
|
3952 using assms |
|
3953 apply (subst homotopic_paths_sym_eq) |
|
3954 unfolding homotopic_paths_def homotopic_with_def |
|
3955 apply (rule_tac x="(\<lambda>y. (subpath 0 (fst y) p +++ reversepath(subpath 0 (fst y) p)) (snd y))" in exI) |
|
3956 apply (simp add: path_defs joinpaths_def subpath_def reversepath_def) |
|
3957 apply (force simp: mult_le_one) |
|
3958 done |
|
3959 qed |
|
3960 |
|
3961 proposition homotopic_paths_linv: |
|
3962 assumes "path p" "path_image p \<subseteq> s" |
|
3963 shows "homotopic_paths s (reversepath p +++ p) (linepath (pathfinish p) (pathfinish p))" |
|
3964 using homotopic_paths_rinv [of "reversepath p" s] assms by simp |
|
3965 |
|
3966 |
|
3967 subsection\<open>Homotopy of loops without requiring preservation of endpoints\<close> |
|
3968 |
|
3969 definition%important homotopic_loops :: "'a::topological_space set \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> bool" where |
|
3970 "homotopic_loops s p q \<equiv> |
|
3971 homotopic_with (\<lambda>r. pathfinish r = pathstart r) {0..1} s p q" |
|
3972 |
|
3973 lemma homotopic_loops: |
|
3974 "homotopic_loops s p q \<longleftrightarrow> |
|
3975 (\<exists>h. continuous_on ({0..1::real} \<times> {0..1}) h \<and> |
|
3976 image h ({0..1} \<times> {0..1}) \<subseteq> s \<and> |
|
3977 (\<forall>x \<in> {0..1}. h(0,x) = p x) \<and> |
|
3978 (\<forall>x \<in> {0..1}. h(1,x) = q x) \<and> |
|
3979 (\<forall>t \<in> {0..1}. pathfinish(h \<circ> Pair t) = pathstart(h \<circ> Pair t)))" |
|
3980 by (simp add: homotopic_loops_def pathstart_def pathfinish_def homotopic_with) |
|
3981 |
|
3982 proposition homotopic_loops_imp_loop: |
|
3983 "homotopic_loops s p q \<Longrightarrow> pathfinish p = pathstart p \<and> pathfinish q = pathstart q" |
|
3984 using homotopic_with_imp_property homotopic_loops_def by blast |
|
3985 |
|
3986 proposition homotopic_loops_imp_path: |
|
3987 "homotopic_loops s p q \<Longrightarrow> path p \<and> path q" |
|
3988 unfolding homotopic_loops_def path_def |
|
3989 using homotopic_with_imp_continuous by blast |
|
3990 |
|
3991 proposition homotopic_loops_imp_subset: |
|
3992 "homotopic_loops s p q \<Longrightarrow> path_image p \<subseteq> s \<and> path_image q \<subseteq> s" |
|
3993 unfolding homotopic_loops_def path_image_def |
|
3994 by (metis homotopic_with_imp_subset1 homotopic_with_imp_subset2) |
|
3995 |
|
3996 proposition homotopic_loops_refl: |
|
3997 "homotopic_loops s p p \<longleftrightarrow> |
|
3998 path p \<and> path_image p \<subseteq> s \<and> pathfinish p = pathstart p" |
|
3999 by (simp add: homotopic_loops_def homotopic_with_refl path_image_def path_def) |
|
4000 |
|
4001 proposition homotopic_loops_sym: "homotopic_loops s p q \<Longrightarrow> homotopic_loops s q p" |
|
4002 by (simp add: homotopic_loops_def homotopic_with_sym) |
|
4003 |
|
4004 proposition homotopic_loops_sym_eq: "homotopic_loops s p q \<longleftrightarrow> homotopic_loops s q p" |
|
4005 by (metis homotopic_loops_sym) |
|
4006 |
|
4007 proposition homotopic_loops_trans: |
|
4008 "\<lbrakk>homotopic_loops s p q; homotopic_loops s q r\<rbrakk> \<Longrightarrow> homotopic_loops s p r" |
|
4009 unfolding homotopic_loops_def by (blast intro: homotopic_with_trans) |
|
4010 |
|
4011 proposition homotopic_loops_subset: |
|
4012 "\<lbrakk>homotopic_loops s p q; s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_loops t p q" |
|
4013 by (simp add: homotopic_loops_def homotopic_with_subset_right) |
|
4014 |
|
4015 proposition homotopic_loops_eq: |
|
4016 "\<lbrakk>path p; path_image p \<subseteq> s; pathfinish p = pathstart p; \<And>t. t \<in> {0..1} \<Longrightarrow> p(t) = q(t)\<rbrakk> |
|
4017 \<Longrightarrow> homotopic_loops s p q" |
|
4018 unfolding homotopic_loops_def |
|
4019 apply (rule homotopic_with_eq) |
|
4020 apply (rule homotopic_with_refl [where f = p, THEN iffD2]) |
|
4021 apply (simp_all add: path_image_def path_def pathstart_def pathfinish_def) |
|
4022 done |
|
4023 |
|
4024 proposition homotopic_loops_continuous_image: |
|
4025 "\<lbrakk>homotopic_loops s f g; continuous_on s h; h ` s \<subseteq> t\<rbrakk> \<Longrightarrow> homotopic_loops t (h \<circ> f) (h \<circ> g)" |
|
4026 unfolding homotopic_loops_def |
|
4027 apply (rule homotopic_with_compose_continuous_left) |
|
4028 apply (erule homotopic_with_mono) |
|
4029 by (simp add: pathfinish_def pathstart_def) |
|
4030 |
|
4031 |
|
4032 subsection\<open>Relations between the two variants of homotopy\<close> |
|
4033 |
|
4034 proposition homotopic_paths_imp_homotopic_loops: |
|
4035 "\<lbrakk>homotopic_paths s p q; pathfinish p = pathstart p; pathfinish q = pathstart p\<rbrakk> \<Longrightarrow> homotopic_loops s p q" |
|
4036 by (auto simp: homotopic_paths_def homotopic_loops_def intro: homotopic_with_mono) |
|
4037 |
|
4038 proposition homotopic_loops_imp_homotopic_paths_null: |
|
4039 assumes "homotopic_loops s p (linepath a a)" |
|
4040 shows "homotopic_paths s p (linepath (pathstart p) (pathstart p))" |
|
4041 proof - |
|
4042 have "path p" by (metis assms homotopic_loops_imp_path) |
|
4043 have ploop: "pathfinish p = pathstart p" by (metis assms homotopic_loops_imp_loop) |
|
4044 have pip: "path_image p \<subseteq> s" by (metis assms homotopic_loops_imp_subset) |
|
4045 obtain h where conth: "continuous_on ({0..1::real} \<times> {0..1}) h" |
|
4046 and hs: "h ` ({0..1} \<times> {0..1}) \<subseteq> s" |
|
4047 and [simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(0,x) = p x" |
|
4048 and [simp]: "\<And>x. x \<in> {0..1} \<Longrightarrow> h(1,x) = a" |
|
4049 and ends: "\<And>t. t \<in> {0..1} \<Longrightarrow> pathfinish (h \<circ> Pair t) = pathstart (h \<circ> Pair t)" |
|
4050 using assms by (auto simp: homotopic_loops homotopic_with) |
|
4051 have conth0: "path (\<lambda>u. h (u, 0))" |
|
4052 unfolding path_def |
|
4053 apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) |
|
4054 apply (force intro: continuous_intros continuous_on_subset [OF conth])+ |
|
4055 done |
|
4056 have pih0: "path_image (\<lambda>u. h (u, 0)) \<subseteq> s" |
|
4057 using hs by (force simp: path_image_def) |
|
4058 have c1: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. h (fst x * snd x, 0))" |
|
4059 apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) |
|
4060 apply (force simp: mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ |
|
4061 done |
|
4062 have c2: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. h (fst x - fst x * snd x, 0))" |
|
4063 apply (rule continuous_on_compose [of _ _ h, unfolded o_def]) |
|
4064 apply (force simp: mult_left_le mult_le_one intro: continuous_intros continuous_on_subset [OF conth])+ |
|
4065 apply (rule continuous_on_subset [OF conth]) |
|
4066 apply (auto simp: algebra_simps add_increasing2 mult_left_le) |
|
4067 done |
|
4068 have [simp]: "\<And>t. \<lbrakk>0 \<le> t \<and> t \<le> 1\<rbrakk> \<Longrightarrow> h (t, 1) = h (t, 0)" |
|
4069 using ends by (simp add: pathfinish_def pathstart_def) |
|
4070 have adhoc_le: "c * 4 \<le> 1 + c * (d * 4)" if "\<not> d * 4 \<le> 3" "0 \<le> c" "c \<le> 1" for c d::real |
|
4071 proof - |
|
4072 have "c * 3 \<le> c * (d * 4)" using that less_eq_real_def by auto |
|
4073 with \<open>c \<le> 1\<close> show ?thesis by fastforce |
|
4074 qed |
|
4075 have *: "\<And>p x. (path p \<and> path(reversepath p)) \<and> |
|
4076 (path_image p \<subseteq> s \<and> path_image(reversepath p) \<subseteq> s) \<and> |
|
4077 (pathfinish p = pathstart(linepath a a +++ reversepath p) \<and> |
|
4078 pathstart(reversepath p) = a) \<and> pathstart p = x |
|
4079 \<Longrightarrow> homotopic_paths s (p +++ linepath a a +++ reversepath p) (linepath x x)" |
|
4080 by (metis homotopic_paths_lid homotopic_paths_join |
|
4081 homotopic_paths_trans homotopic_paths_sym homotopic_paths_rinv) |
|
4082 have 1: "homotopic_paths s p (p +++ linepath (pathfinish p) (pathfinish p))" |
|
4083 using \<open>path p\<close> homotopic_paths_rid homotopic_paths_sym pip by blast |
|
4084 moreover have "homotopic_paths s (p +++ linepath (pathfinish p) (pathfinish p)) |
|
4085 (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p))" |
|
4086 apply (rule homotopic_paths_sym) |
|
4087 using homotopic_paths_lid [of "p +++ linepath (pathfinish p) (pathfinish p)" s] |
|
4088 by (metis 1 homotopic_paths_imp_path homotopic_paths_imp_pathstart homotopic_paths_imp_subset) |
|
4089 moreover have "homotopic_paths s (linepath (pathstart p) (pathstart p) +++ p +++ linepath (pathfinish p) (pathfinish p)) |
|
4090 ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0)))" |
|
4091 apply (simp add: homotopic_paths_def homotopic_with_def) |
|
4092 apply (rule_tac x="\<lambda>y. (subpath 0 (fst y) (\<lambda>u. h (u, 0)) +++ (\<lambda>u. h (Pair (fst y) u)) +++ subpath (fst y) 0 (\<lambda>u. h (u, 0))) (snd y)" in exI) |
|
4093 apply (simp add: subpath_reversepath) |
|
4094 apply (intro conjI homotopic_join_lemma) |
|
4095 using ploop |
|
4096 apply (simp_all add: path_defs joinpaths_def o_def subpath_def conth c1 c2) |
|
4097 apply (force simp: algebra_simps mult_le_one mult_left_le intro: hs [THEN subsetD] adhoc_le) |
|
4098 done |
|
4099 moreover have "homotopic_paths s ((\<lambda>u. h (u, 0)) +++ linepath a a +++ reversepath (\<lambda>u. h (u, 0))) |
|
4100 (linepath (pathstart p) (pathstart p))" |
|
4101 apply (rule *) |
|
4102 apply (simp add: pih0 pathstart_def pathfinish_def conth0) |
|
4103 apply (simp add: reversepath_def joinpaths_def) |
|
4104 done |
|
4105 ultimately show ?thesis |
|
4106 by (blast intro: homotopic_paths_trans) |
|
4107 qed |
|
4108 |
|
4109 proposition homotopic_loops_conjugate: |
|
4110 fixes s :: "'a::real_normed_vector set" |
|
4111 assumes "path p" "path q" and pip: "path_image p \<subseteq> s" and piq: "path_image q \<subseteq> s" |
|
4112 and papp: "pathfinish p = pathstart q" and qloop: "pathfinish q = pathstart q" |
|
4113 shows "homotopic_loops s (p +++ q +++ reversepath p) q" |
|
4114 proof - |
|
4115 have contp: "continuous_on {0..1} p" using \<open>path p\<close> [unfolded path_def] by blast |
|
4116 have contq: "continuous_on {0..1} q" using \<open>path q\<close> [unfolded path_def] by blast |
|
4117 have c1: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. p ((1 - fst x) * snd x + fst x))" |
|
4118 apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) |
|
4119 apply (force simp: mult_le_one intro!: continuous_intros) |
|
4120 apply (rule continuous_on_subset [OF contp]) |
|
4121 apply (auto simp: algebra_simps add_increasing2 mult_right_le_one_le sum_le_prod1) |
|
4122 done |
|
4123 have c2: "continuous_on ({0..1} \<times> {0..1}) (\<lambda>x. p ((fst x - 1) * snd x + 1))" |
|
4124 apply (rule continuous_on_compose [of _ _ p, unfolded o_def]) |
|
4125 apply (force simp: mult_le_one intro!: continuous_intros) |
|
4126 apply (rule continuous_on_subset [OF contp]) |
|
4127 apply (auto simp: algebra_simps add_increasing2 mult_left_le_one_le) |
|
4128 done |
|
4129 have ps1: "\<And>a b. \<lbrakk>b * 2 \<le> 1; 0 \<le> b; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((1 - a) * (2 * b) + a) \<in> s" |
|
4130 using sum_le_prod1 |
|
4131 by (force simp: algebra_simps add_increasing2 mult_left_le intro: pip [unfolded path_image_def, THEN subsetD]) |
|
4132 have ps2: "\<And>a b. \<lbrakk>\<not> 4 * b \<le> 3; b \<le> 1; 0 \<le> a; a \<le> 1\<rbrakk> \<Longrightarrow> p ((a - 1) * (4 * b - 3) + 1) \<in> s" |
|
4133 apply (rule pip [unfolded path_image_def, THEN subsetD]) |
|
4134 apply (rule image_eqI, blast) |
|
4135 apply (simp add: algebra_simps) |
|
4136 by (metis add_mono_thms_linordered_semiring(1) affine_ineq linear mult.commute mult.left_neutral mult_right_mono not_le |
|
4137 add.commute zero_le_numeral) |
|
4138 have qs: "\<And>a b. \<lbrakk>4 * b \<le> 3; \<not> b * 2 \<le> 1\<rbrakk> \<Longrightarrow> q (4 * b - 2) \<in> s" |
|
4139 using path_image_def piq by fastforce |
|
4140 have "homotopic_loops s (p +++ q +++ reversepath p) |
|
4141 (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q))" |
|
4142 apply (simp add: homotopic_loops_def homotopic_with_def) |
|
4143 apply (rule_tac x="(\<lambda>y. (subpath (fst y) 1 p +++ q +++ subpath 1 (fst y) p) (snd y))" in exI) |
|
4144 apply (simp add: subpath_refl subpath_reversepath) |
|
4145 apply (intro conjI homotopic_join_lemma) |
|
4146 using papp qloop |
|
4147 apply (simp_all add: path_defs joinpaths_def o_def subpath_def c1 c2) |
|
4148 apply (force simp: contq intro: continuous_on_compose [of _ _ q, unfolded o_def] continuous_on_id continuous_on_snd) |
|
4149 apply (auto simp: ps1 ps2 qs) |
|
4150 done |
|
4151 moreover have "homotopic_loops s (linepath (pathstart q) (pathstart q) +++ q +++ linepath (pathstart q) (pathstart q)) q" |
|
4152 proof - |
|
4153 have "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q) q" |
|
4154 using \<open>path q\<close> homotopic_paths_lid qloop piq by auto |
|
4155 hence 1: "\<And>f. homotopic_paths s f q \<or> \<not> homotopic_paths s f (linepath (pathfinish q) (pathfinish q) +++ q)" |
|
4156 using homotopic_paths_trans by blast |
|
4157 hence "homotopic_paths s (linepath (pathfinish q) (pathfinish q) +++ q +++ linepath (pathfinish q) (pathfinish q)) q" |
|
4158 proof - |
|
4159 have "homotopic_paths s (q +++ linepath (pathfinish q) (pathfinish q)) q" |
|
4160 by (simp add: \<open>path q\<close> homotopic_paths_rid piq) |
|
4161 thus ?thesis |
|
4162 by (metis (no_types) 1 \<open>path q\<close> homotopic_paths_join homotopic_paths_rinv homotopic_paths_sym |
|
4163 homotopic_paths_trans qloop pathfinish_linepath piq) |
|
4164 qed |
|
4165 thus ?thesis |
|
4166 by (metis (no_types) qloop homotopic_loops_sym homotopic_paths_imp_homotopic_loops homotopic_paths_imp_pathfinish homotopic_paths_sym) |
|
4167 qed |
|
4168 ultimately show ?thesis |
|
4169 by (blast intro: homotopic_loops_trans) |
|
4170 qed |
|
4171 |
|
4172 lemma homotopic_paths_loop_parts: |
|
4173 assumes loops: "homotopic_loops S (p +++ reversepath q) (linepath a a)" and "path q" |
|
4174 shows "homotopic_paths S p q" |
|
4175 proof - |
|
4176 have paths: "homotopic_paths S (p +++ reversepath q) (linepath (pathstart p) (pathstart p))" |
|
4177 using homotopic_loops_imp_homotopic_paths_null [OF loops] by simp |
|
4178 then have "path p" |
|
4179 using \<open>path q\<close> homotopic_loops_imp_path loops path_join path_join_path_ends path_reversepath by blast |
|
4180 show ?thesis |
|
4181 proof (cases "pathfinish p = pathfinish q") |
|
4182 case True |
|
4183 have pipq: "path_image p \<subseteq> S" "path_image q \<subseteq> S" |
|
4184 by (metis Un_subset_iff paths \<open>path p\<close> \<open>path q\<close> homotopic_loops_imp_subset homotopic_paths_imp_path loops |
|
4185 path_image_join path_image_reversepath path_imp_reversepath path_join_eq)+ |
|
4186 have "homotopic_paths S p (p +++ (linepath (pathfinish p) (pathfinish p)))" |
|
4187 using \<open>path p\<close> \<open>path_image p \<subseteq> S\<close> homotopic_paths_rid homotopic_paths_sym by blast |
|
4188 moreover have "homotopic_paths S (p +++ (linepath (pathfinish p) (pathfinish p))) (p +++ (reversepath q +++ q))" |
|
4189 by (simp add: True \<open>path p\<close> \<open>path q\<close> pipq homotopic_paths_join homotopic_paths_linv homotopic_paths_sym) |
|
4190 moreover have "homotopic_paths S (p +++ (reversepath q +++ q)) ((p +++ reversepath q) +++ q)" |
|
4191 by (simp add: True \<open>path p\<close> \<open>path q\<close> homotopic_paths_assoc pipq) |
|
4192 moreover have "homotopic_paths S ((p +++ reversepath q) +++ q) (linepath (pathstart p) (pathstart p) +++ q)" |
|
4193 by (simp add: \<open>path q\<close> homotopic_paths_join paths pipq) |
|
4194 moreover then have "homotopic_paths S (linepath (pathstart p) (pathstart p) +++ q) q" |
|
4195 by (metis \<open>path q\<close> homotopic_paths_imp_path homotopic_paths_lid linepath_trivial path_join_path_ends pathfinish_def pipq(2)) |
|
4196 ultimately show ?thesis |
|
4197 using homotopic_paths_trans by metis |
|
4198 next |
|
4199 case False |
|
4200 then show ?thesis |
|
4201 using \<open>path q\<close> homotopic_loops_imp_path loops path_join_path_ends by fastforce |
|
4202 qed |
|
4203 qed |
|
4204 |
|
4205 |
|
4206 subsection%unimportant\<open>Homotopy of "nearby" function, paths and loops\<close> |
|
4207 |
|
4208 lemma homotopic_with_linear: |
|
4209 fixes f g :: "_ \<Rightarrow> 'b::real_normed_vector" |
|
4210 assumes contf: "continuous_on s f" |
|
4211 and contg:"continuous_on s g" |
|
4212 and sub: "\<And>x. x \<in> s \<Longrightarrow> closed_segment (f x) (g x) \<subseteq> t" |
|
4213 shows "homotopic_with (\<lambda>z. True) s t f g" |
|
4214 apply (simp add: homotopic_with_def) |
|
4215 apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R f(snd y) + (fst y) *\<^sub>R g(snd y))" in exI) |
|
4216 apply (intro conjI) |
|
4217 apply (rule subset_refl continuous_intros continuous_on_subset [OF contf] continuous_on_compose2 [where g=f] |
|
4218 continuous_on_subset [OF contg] continuous_on_compose2 [where g=g]| simp)+ |
|
4219 using sub closed_segment_def apply fastforce+ |
|
4220 done |
|
4221 |
|
4222 lemma homotopic_paths_linear: |
|
4223 fixes g h :: "real \<Rightarrow> 'a::real_normed_vector" |
|
4224 assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" |
|
4225 "\<And>t. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> s" |
|
4226 shows "homotopic_paths s g h" |
|
4227 using assms |
|
4228 unfolding path_def |
|
4229 apply (simp add: closed_segment_def pathstart_def pathfinish_def homotopic_paths_def homotopic_with_def) |
|
4230 apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R (g \<circ> snd) y + (fst y) *\<^sub>R (h \<circ> snd) y)" in exI) |
|
4231 apply (intro conjI subsetI continuous_intros; force) |
|
4232 done |
|
4233 |
|
4234 lemma homotopic_loops_linear: |
|
4235 fixes g h :: "real \<Rightarrow> 'a::real_normed_vector" |
|
4236 assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" |
|
4237 "\<And>t x. t \<in> {0..1} \<Longrightarrow> closed_segment (g t) (h t) \<subseteq> s" |
|
4238 shows "homotopic_loops s g h" |
|
4239 using assms |
|
4240 unfolding path_def |
|
4241 apply (simp add: pathstart_def pathfinish_def homotopic_loops_def homotopic_with_def) |
|
4242 apply (rule_tac x="\<lambda>y. ((1 - (fst y)) *\<^sub>R g(snd y) + (fst y) *\<^sub>R h(snd y))" in exI) |
|
4243 apply (auto intro!: continuous_intros intro: continuous_on_compose2 [where g=g] continuous_on_compose2 [where g=h]) |
|
4244 apply (force simp: closed_segment_def) |
|
4245 done |
|
4246 |
|
4247 lemma homotopic_paths_nearby_explicit: |
|
4248 assumes "path g" "path h" "pathstart h = pathstart g" "pathfinish h = pathfinish g" |
|
4249 and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> s\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)" |
|
4250 shows "homotopic_paths s g h" |
|
4251 apply (rule homotopic_paths_linear [OF assms(1-4)]) |
|
4252 by (metis no segment_bound(1) subsetI norm_minus_commute not_le) |
|
4253 |
|
4254 lemma homotopic_loops_nearby_explicit: |
|
4255 assumes "path g" "path h" "pathfinish g = pathstart g" "pathfinish h = pathstart h" |
|
4256 and no: "\<And>t x. \<lbrakk>t \<in> {0..1}; x \<notin> s\<rbrakk> \<Longrightarrow> norm(h t - g t) < norm(g t - x)" |
|
4257 shows "homotopic_loops s g h" |
|
4258 apply (rule homotopic_loops_linear [OF assms(1-4)]) |
|
4259 by (metis no segment_bound(1) subsetI norm_minus_commute not_le) |
|
4260 |
|
4261 lemma homotopic_nearby_paths: |
|
4262 fixes g h :: "real \<Rightarrow> 'a::euclidean_space" |
|
4263 assumes "path g" "open s" "path_image g \<subseteq> s" |
|
4264 shows "\<exists>e. 0 < e \<and> |
|
4265 (\<forall>h. path h \<and> |
|
4266 pathstart h = pathstart g \<and> pathfinish h = pathfinish g \<and> |
|
4267 (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_paths s g h)" |
|
4268 proof - |
|
4269 obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - s \<Longrightarrow> e \<le> dist x y" |
|
4270 using separate_compact_closed [of "path_image g" "-s"] assms by force |
|
4271 show ?thesis |
|
4272 apply (intro exI conjI) |
|
4273 using e [unfolded dist_norm] |
|
4274 apply (auto simp: intro!: homotopic_paths_nearby_explicit assms \<open>e > 0\<close>) |
|
4275 by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) |
|
4276 qed |
|
4277 |
|
4278 lemma homotopic_nearby_loops: |
|
4279 fixes g h :: "real \<Rightarrow> 'a::euclidean_space" |
|
4280 assumes "path g" "open s" "path_image g \<subseteq> s" "pathfinish g = pathstart g" |
|
4281 shows "\<exists>e. 0 < e \<and> |
|
4282 (\<forall>h. path h \<and> pathfinish h = pathstart h \<and> |
|
4283 (\<forall>t \<in> {0..1}. norm(h t - g t) < e) \<longrightarrow> homotopic_loops s g h)" |
|
4284 proof - |
|
4285 obtain e where "e > 0" and e: "\<And>x y. x \<in> path_image g \<Longrightarrow> y \<in> - s \<Longrightarrow> e \<le> dist x y" |
|
4286 using separate_compact_closed [of "path_image g" "-s"] assms by force |
|
4287 show ?thesis |
|
4288 apply (intro exI conjI) |
|
4289 using e [unfolded dist_norm] |
|
4290 apply (auto simp: intro!: homotopic_loops_nearby_explicit assms \<open>e > 0\<close>) |
|
4291 by (metis atLeastAtMost_iff imageI le_less_trans not_le path_image_def) |
|
4292 qed |
|
4293 |
|
4294 |
|
4295 subsection\<open> Homotopy and subpaths\<close> |
|
4296 |
|
4297 lemma homotopic_join_subpaths1: |
|
4298 assumes "path g" and pag: "path_image g \<subseteq> s" |
|
4299 and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}" "u \<le> v" "v \<le> w" |
|
4300 shows "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" |
|
4301 proof - |
|
4302 have 1: "t * 2 \<le> 1 \<Longrightarrow> u + t * (v * 2) \<le> v + t * (u * 2)" for t |
|
4303 using affine_ineq \<open>u \<le> v\<close> by fastforce |
|
4304 have 2: "t * 2 > 1 \<Longrightarrow> u + (2*t - 1) * v \<le> v + (2*t - 1) * w" for t |
|
4305 by (metis add_mono_thms_linordered_semiring(1) diff_gt_0_iff_gt less_eq_real_def mult.commute mult_right_mono \<open>u \<le> v\<close> \<open>v \<le> w\<close>) |
|
4306 have t2: "\<And>t::real. t*2 = 1 \<Longrightarrow> t = 1/2" by auto |
|
4307 show ?thesis |
|
4308 apply (rule homotopic_paths_subset [OF _ pag]) |
|
4309 using assms |
|
4310 apply (cases "w = u") |
|
4311 using homotopic_paths_rinv [of "subpath u v g" "path_image g"] |
|
4312 apply (force simp: closed_segment_eq_real_ivl image_mono path_image_def subpath_refl) |
|
4313 apply (rule homotopic_paths_sym) |
|
4314 apply (rule homotopic_paths_reparametrize |
|
4315 [where f = "\<lambda>t. if t \<le> 1 / 2 |
|
4316 then inverse((w - u)) *\<^sub>R (2 * (v - u)) *\<^sub>R t |
|
4317 else inverse((w - u)) *\<^sub>R ((v - u) + (w - v) *\<^sub>R (2 *\<^sub>R t - 1))"]) |
|
4318 using \<open>path g\<close> path_subpath u w apply blast |
|
4319 using \<open>path g\<close> path_image_subpath_subset u w(1) apply blast |
|
4320 apply simp_all |
|
4321 apply (subst split_01) |
|
4322 apply (rule continuous_on_cases continuous_intros | force simp: pathfinish_def joinpaths_def)+ |
|
4323 apply (simp_all add: field_simps not_le) |
|
4324 apply (force dest!: t2) |
|
4325 apply (force simp: algebra_simps mult_left_mono affine_ineq dest!: 1 2) |
|
4326 apply (simp add: joinpaths_def subpath_def) |
|
4327 apply (force simp: algebra_simps) |
|
4328 done |
|
4329 qed |
|
4330 |
|
4331 lemma homotopic_join_subpaths2: |
|
4332 assumes "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" |
|
4333 shows "homotopic_paths s (subpath w v g +++ subpath v u g) (subpath w u g)" |
|
4334 by (metis assms homotopic_paths_reversepath_D pathfinish_subpath pathstart_subpath reversepath_joinpaths reversepath_subpath) |
|
4335 |
|
4336 lemma homotopic_join_subpaths3: |
|
4337 assumes hom: "homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" |
|
4338 and "path g" and pag: "path_image g \<subseteq> s" |
|
4339 and u: "u \<in> {0..1}" and v: "v \<in> {0..1}" and w: "w \<in> {0..1}" |
|
4340 shows "homotopic_paths s (subpath v w g +++ subpath w u g) (subpath v u g)" |
|
4341 proof - |
|
4342 have "homotopic_paths s (subpath u w g +++ subpath w v g) ((subpath u v g +++ subpath v w g) +++ subpath w v g)" |
|
4343 apply (rule homotopic_paths_join) |
|
4344 using hom homotopic_paths_sym_eq apply blast |
|
4345 apply (metis \<open>path g\<close> homotopic_paths_eq pag path_image_subpath_subset path_subpath subset_trans v w, simp) |
|
4346 done |
|
4347 also have "homotopic_paths s ((subpath u v g +++ subpath v w g) +++ subpath w v g) (subpath u v g +++ subpath v w g +++ subpath w v g)" |
|
4348 apply (rule homotopic_paths_sym [OF homotopic_paths_assoc]) |
|
4349 using assms by (simp_all add: path_image_subpath_subset [THEN order_trans]) |
|
4350 also have "homotopic_paths s (subpath u v g +++ subpath v w g +++ subpath w v g) |
|
4351 (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g)))" |
|
4352 apply (rule homotopic_paths_join) |
|
4353 apply (metis \<open>path g\<close> homotopic_paths_eq order.trans pag path_image_subpath_subset path_subpath u v) |
|
4354 apply (metis (no_types, lifting) \<open>path g\<close> homotopic_paths_linv order_trans pag path_image_subpath_subset path_subpath pathfinish_subpath reversepath_subpath v w) |
|
4355 apply simp |
|
4356 done |
|
4357 also have "homotopic_paths s (subpath u v g +++ linepath (pathfinish (subpath u v g)) (pathfinish (subpath u v g))) (subpath u v g)" |
|
4358 apply (rule homotopic_paths_rid) |
|
4359 using \<open>path g\<close> path_subpath u v apply blast |
|
4360 apply (meson \<open>path g\<close> order.trans pag path_image_subpath_subset u v) |
|
4361 done |
|
4362 finally have "homotopic_paths s (subpath u w g +++ subpath w v g) (subpath u v g)" . |
|
4363 then show ?thesis |
|
4364 using homotopic_join_subpaths2 by blast |
|
4365 qed |
|
4366 |
|
4367 proposition homotopic_join_subpaths: |
|
4368 "\<lbrakk>path g; path_image g \<subseteq> s; u \<in> {0..1}; v \<in> {0..1}; w \<in> {0..1}\<rbrakk> |
|
4369 \<Longrightarrow> homotopic_paths s (subpath u v g +++ subpath v w g) (subpath u w g)" |
|
4370 apply (rule le_cases3 [of u v w]) |
|
4371 using homotopic_join_subpaths1 homotopic_join_subpaths2 homotopic_join_subpaths3 by metis+ |
|
4372 |
|
4373 text\<open>Relating homotopy of trivial loops to path-connectedness.\<close> |
|
4374 |
|
4375 lemma path_component_imp_homotopic_points: |
|
4376 "path_component S a b \<Longrightarrow> homotopic_loops S (linepath a a) (linepath b b)" |
|
4377 apply (simp add: path_component_def homotopic_loops_def homotopic_with_def |
|
4378 pathstart_def pathfinish_def path_image_def path_def, clarify) |
|
4379 apply (rule_tac x="g \<circ> fst" in exI) |
|
4380 apply (intro conjI continuous_intros continuous_on_compose)+ |
|
4381 apply (auto elim!: continuous_on_subset) |
|
4382 done |
|
4383 |
|
4384 lemma homotopic_loops_imp_path_component_value: |
|
4385 "\<lbrakk>homotopic_loops S p q; 0 \<le> t; t \<le> 1\<rbrakk> |
|
4386 \<Longrightarrow> path_component S (p t) (q t)" |
|
4387 apply (simp add: path_component_def homotopic_loops_def homotopic_with_def |
|
4388 pathstart_def pathfinish_def path_image_def path_def, clarify) |
|
4389 apply (rule_tac x="h \<circ> (\<lambda>u. (u, t))" in exI) |
|
4390 apply (intro conjI continuous_intros continuous_on_compose)+ |
|
4391 apply (auto elim!: continuous_on_subset) |
|
4392 done |
|
4393 |
|
4394 lemma homotopic_points_eq_path_component: |
|
4395 "homotopic_loops S (linepath a a) (linepath b b) \<longleftrightarrow> |
|
4396 path_component S a b" |
|
4397 by (auto simp: path_component_imp_homotopic_points |
|
4398 dest: homotopic_loops_imp_path_component_value [where t=1]) |
|
4399 |
|
4400 lemma path_connected_eq_homotopic_points: |
|
4401 "path_connected S \<longleftrightarrow> |
|
4402 (\<forall>a b. a \<in> S \<and> b \<in> S \<longrightarrow> homotopic_loops S (linepath a a) (linepath b b))" |
|
4403 by (auto simp: path_connected_def path_component_def homotopic_points_eq_path_component) |
|
4404 |
|
4405 |
|
4406 subsection\<open>Simply connected sets\<close> |
|
4407 |
|
4408 text%important\<open>defined as "all loops are homotopic (as loops)\<close> |
|
4409 |
|
4410 definition%important simply_connected where |
|
4411 "simply_connected S \<equiv> |
|
4412 \<forall>p q. path p \<and> pathfinish p = pathstart p \<and> path_image p \<subseteq> S \<and> |
|
4413 path q \<and> pathfinish q = pathstart q \<and> path_image q \<subseteq> S |
|
4414 \<longrightarrow> homotopic_loops S p q" |
|
4415 |
|
4416 lemma simply_connected_empty [iff]: "simply_connected {}" |
|
4417 by (simp add: simply_connected_def) |
|
4418 |
|
4419 lemma simply_connected_imp_path_connected: |
|
4420 fixes S :: "_::real_normed_vector set" |
|
4421 shows "simply_connected S \<Longrightarrow> path_connected S" |
|
4422 by (simp add: simply_connected_def path_connected_eq_homotopic_points) |
|
4423 |
|
4424 lemma simply_connected_imp_connected: |
|
4425 fixes S :: "_::real_normed_vector set" |
|
4426 shows "simply_connected S \<Longrightarrow> connected S" |
|
4427 by (simp add: path_connected_imp_connected simply_connected_imp_path_connected) |
|
4428 |
|
4429 lemma simply_connected_eq_contractible_loop_any: |
|
4430 fixes S :: "_::real_normed_vector set" |
|
4431 shows "simply_connected S \<longleftrightarrow> |
|
4432 (\<forall>p a. path p \<and> path_image p \<subseteq> S \<and> |
|
4433 pathfinish p = pathstart p \<and> a \<in> S |
|
4434 \<longrightarrow> homotopic_loops S p (linepath a a))" |
|
4435 apply (simp add: simply_connected_def) |
|
4436 apply (rule iffI, force, clarify) |
|
4437 apply (rule_tac q = "linepath (pathstart p) (pathstart p)" in homotopic_loops_trans) |
|
4438 apply (fastforce simp add:) |
|
4439 using homotopic_loops_sym apply blast |
|
4440 done |
|
4441 |
|
4442 lemma simply_connected_eq_contractible_loop_some: |
|
4443 fixes S :: "_::real_normed_vector set" |
|
4444 shows "simply_connected S \<longleftrightarrow> |
|
4445 path_connected S \<and> |
|
4446 (\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p |
|
4447 \<longrightarrow> (\<exists>a. a \<in> S \<and> homotopic_loops S p (linepath a a)))" |
|
4448 apply (rule iffI) |
|
4449 apply (fastforce simp: simply_connected_imp_path_connected simply_connected_eq_contractible_loop_any) |
|
4450 apply (clarsimp simp add: simply_connected_eq_contractible_loop_any) |
|
4451 apply (drule_tac x=p in spec) |
|
4452 using homotopic_loops_trans path_connected_eq_homotopic_points |
|
4453 apply blast |
|
4454 done |
|
4455 |
|
4456 lemma simply_connected_eq_contractible_loop_all: |
|
4457 fixes S :: "_::real_normed_vector set" |
|
4458 shows "simply_connected S \<longleftrightarrow> |
|
4459 S = {} \<or> |
|
4460 (\<exists>a \<in> S. \<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p |
|
4461 \<longrightarrow> homotopic_loops S p (linepath a a))" |
|
4462 (is "?lhs = ?rhs") |
|
4463 proof (cases "S = {}") |
|
4464 case True then show ?thesis by force |
|
4465 next |
|
4466 case False |
|
4467 then obtain a where "a \<in> S" by blast |
|
4468 show ?thesis |
|
4469 proof |
|
4470 assume "simply_connected S" |
|
4471 then show ?rhs |
|
4472 using \<open>a \<in> S\<close> \<open>simply_connected S\<close> simply_connected_eq_contractible_loop_any |
|
4473 by blast |
|
4474 next |
|
4475 assume ?rhs |
|
4476 then show "simply_connected S" |
|
4477 apply (simp add: simply_connected_eq_contractible_loop_any False) |
|
4478 by (meson homotopic_loops_refl homotopic_loops_sym homotopic_loops_trans |
|
4479 path_component_imp_homotopic_points path_component_refl) |
|
4480 qed |
|
4481 qed |
|
4482 |
|
4483 lemma simply_connected_eq_contractible_path: |
|
4484 fixes S :: "_::real_normed_vector set" |
|
4485 shows "simply_connected S \<longleftrightarrow> |
|
4486 path_connected S \<and> |
|
4487 (\<forall>p. path p \<and> path_image p \<subseteq> S \<and> pathfinish p = pathstart p |
|
4488 \<longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p)))" |
|
4489 apply (rule iffI) |
|
4490 apply (simp add: simply_connected_imp_path_connected) |
|
4491 apply (metis simply_connected_eq_contractible_loop_some homotopic_loops_imp_homotopic_paths_null) |
|
4492 by (meson homotopic_paths_imp_homotopic_loops pathfinish_linepath pathstart_in_path_image |
|
4493 simply_connected_eq_contractible_loop_some subset_iff) |
|
4494 |
|
4495 lemma simply_connected_eq_homotopic_paths: |
|
4496 fixes S :: "_::real_normed_vector set" |
|
4497 shows "simply_connected S \<longleftrightarrow> |
|
4498 path_connected S \<and> |
|
4499 (\<forall>p q. path p \<and> path_image p \<subseteq> S \<and> |
|
4500 path q \<and> path_image q \<subseteq> S \<and> |
|
4501 pathstart q = pathstart p \<and> pathfinish q = pathfinish p |
|
4502 \<longrightarrow> homotopic_paths S p q)" |
|
4503 (is "?lhs = ?rhs") |
|
4504 proof |
|
4505 assume ?lhs |
|
4506 then have pc: "path_connected S" |
|
4507 and *: "\<And>p. \<lbrakk>path p; path_image p \<subseteq> S; |
|
4508 pathfinish p = pathstart p\<rbrakk> |
|
4509 \<Longrightarrow> homotopic_paths S p (linepath (pathstart p) (pathstart p))" |
|
4510 by (auto simp: simply_connected_eq_contractible_path) |
|
4511 have "homotopic_paths S p q" |
|
4512 if "path p" "path_image p \<subseteq> S" "path q" |
|
4513 "path_image q \<subseteq> S" "pathstart q = pathstart p" |
|
4514 "pathfinish q = pathfinish p" for p q |
|
4515 proof - |
|
4516 have "homotopic_paths S p (p +++ linepath (pathfinish p) (pathfinish p))" |
|
4517 by (simp add: homotopic_paths_rid homotopic_paths_sym that) |
|
4518 also have "homotopic_paths S (p +++ linepath (pathfinish p) (pathfinish p)) |
|
4519 (p +++ reversepath q +++ q)" |
|
4520 using that |
|
4521 by (metis homotopic_paths_join homotopic_paths_linv homotopic_paths_refl homotopic_paths_sym_eq pathstart_linepath) |
|
4522 also have "homotopic_paths S (p +++ reversepath q +++ q) |
|
4523 ((p +++ reversepath q) +++ q)" |
|
4524 by (simp add: that homotopic_paths_assoc) |
|
4525 also have "homotopic_paths S ((p +++ reversepath q) +++ q) |
|
4526 (linepath (pathstart q) (pathstart q) +++ q)" |
|
4527 using * [of "p +++ reversepath q"] that |
|
4528 by (simp add: homotopic_paths_join path_image_join) |
|
4529 also have "homotopic_paths S (linepath (pathstart q) (pathstart q) +++ q) q" |
|
4530 using that homotopic_paths_lid by blast |
|
4531 finally show ?thesis . |
|
4532 qed |
|
4533 then show ?rhs |
|
4534 by (blast intro: pc *) |
|
4535 next |
|
4536 assume ?rhs |
|
4537 then show ?lhs |
|
4538 by (force simp: simply_connected_eq_contractible_path) |
|
4539 qed |
|
4540 |
|
4541 proposition simply_connected_Times: |
|
4542 fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
4543 assumes S: "simply_connected S" and T: "simply_connected T" |
|
4544 shows "simply_connected(S \<times> T)" |
|
4545 proof - |
|
4546 have "homotopic_loops (S \<times> T) p (linepath (a, b) (a, b))" |
|
4547 if "path p" "path_image p \<subseteq> S \<times> T" "p 1 = p 0" "a \<in> S" "b \<in> T" |
|
4548 for p a b |
|
4549 proof - |
|
4550 have "path (fst \<circ> p)" |
|
4551 apply (rule Path_Connected.path_continuous_image [OF \<open>path p\<close>]) |
|
4552 apply (rule continuous_intros)+ |
|
4553 done |
|
4554 moreover have "path_image (fst \<circ> p) \<subseteq> S" |
|
4555 using that apply (simp add: path_image_def) by force |
|
4556 ultimately have p1: "homotopic_loops S (fst \<circ> p) (linepath a a)" |
|
4557 using S that |
|
4558 apply (simp add: simply_connected_eq_contractible_loop_any) |
|
4559 apply (drule_tac x="fst \<circ> p" in spec) |
|
4560 apply (drule_tac x=a in spec) |
|
4561 apply (auto simp: pathstart_def pathfinish_def) |
|
4562 done |
|
4563 have "path (snd \<circ> p)" |
|
4564 apply (rule Path_Connected.path_continuous_image [OF \<open>path p\<close>]) |
|
4565 apply (rule continuous_intros)+ |
|
4566 done |
|
4567 moreover have "path_image (snd \<circ> p) \<subseteq> T" |
|
4568 using that apply (simp add: path_image_def) by force |
|
4569 ultimately have p2: "homotopic_loops T (snd \<circ> p) (linepath b b)" |
|
4570 using T that |
|
4571 apply (simp add: simply_connected_eq_contractible_loop_any) |
|
4572 apply (drule_tac x="snd \<circ> p" in spec) |
|
4573 apply (drule_tac x=b in spec) |
|
4574 apply (auto simp: pathstart_def pathfinish_def) |
|
4575 done |
|
4576 show ?thesis |
|
4577 using p1 p2 |
|
4578 apply (simp add: homotopic_loops, clarify) |
|
4579 apply (rename_tac h k) |
|
4580 apply (rule_tac x="\<lambda>z. Pair (h z) (k z)" in exI) |
|
4581 apply (intro conjI continuous_intros | assumption)+ |
|
4582 apply (auto simp: pathstart_def pathfinish_def) |
|
4583 done |
|
4584 qed |
|
4585 with assms show ?thesis |
|
4586 by (simp add: simply_connected_eq_contractible_loop_any pathfinish_def pathstart_def) |
|
4587 qed |
|
4588 |
|
4589 |
|
4590 subsection\<open>Contractible sets\<close> |
|
4591 |
|
4592 definition%important contractible where |
|
4593 "contractible S \<equiv> \<exists>a. homotopic_with (\<lambda>x. True) S S id (\<lambda>x. a)" |
|
4594 |
|
4595 proposition contractible_imp_simply_connected: |
|
4596 fixes S :: "_::real_normed_vector set" |
|
4597 assumes "contractible S" shows "simply_connected S" |
|
4598 proof (cases "S = {}") |
|
4599 case True then show ?thesis by force |
|
4600 next |
|
4601 case False |
|
4602 obtain a where a: "homotopic_with (\<lambda>x. True) S S id (\<lambda>x. a)" |
|
4603 using assms by (force simp: contractible_def) |
|
4604 then have "a \<in> S" |
|
4605 by (metis False homotopic_constant_maps homotopic_with_symD homotopic_with_trans path_component_mem(2)) |
|
4606 show ?thesis |
|
4607 apply (simp add: simply_connected_eq_contractible_loop_all False) |
|
4608 apply (rule bexI [OF _ \<open>a \<in> S\<close>]) |
|
4609 using a apply (simp add: homotopic_loops_def homotopic_with_def path_def path_image_def pathfinish_def pathstart_def, clarify) |
|
4610 apply (rule_tac x="(h \<circ> (\<lambda>y. (fst y, (p \<circ> snd) y)))" in exI) |
|
4611 apply (intro conjI continuous_on_compose continuous_intros) |
|
4612 apply (erule continuous_on_subset | force)+ |
|
4613 done |
|
4614 qed |
|
4615 |
|
4616 corollary contractible_imp_connected: |
|
4617 fixes S :: "_::real_normed_vector set" |
|
4618 shows "contractible S \<Longrightarrow> connected S" |
|
4619 by (simp add: contractible_imp_simply_connected simply_connected_imp_connected) |
|
4620 |
|
4621 lemma contractible_imp_path_connected: |
|
4622 fixes S :: "_::real_normed_vector set" |
|
4623 shows "contractible S \<Longrightarrow> path_connected S" |
|
4624 by (simp add: contractible_imp_simply_connected simply_connected_imp_path_connected) |
|
4625 |
|
4626 lemma nullhomotopic_through_contractible: |
|
4627 fixes S :: "_::topological_space set" |
|
4628 assumes f: "continuous_on S f" "f ` S \<subseteq> T" |
|
4629 and g: "continuous_on T g" "g ` T \<subseteq> U" |
|
4630 and T: "contractible T" |
|
4631 obtains c where "homotopic_with (\<lambda>h. True) S U (g \<circ> f) (\<lambda>x. c)" |
|
4632 proof - |
|
4633 obtain b where b: "homotopic_with (\<lambda>x. True) T T id (\<lambda>x. b)" |
|
4634 using assms by (force simp: contractible_def) |
|
4635 have "homotopic_with (\<lambda>f. True) T U (g \<circ> id) (g \<circ> (\<lambda>x. b))" |
|
4636 by (rule homotopic_compose_continuous_left [OF b g]) |
|
4637 then have "homotopic_with (\<lambda>f. True) S U (g \<circ> id \<circ> f) (g \<circ> (\<lambda>x. b) \<circ> f)" |
|
4638 by (rule homotopic_compose_continuous_right [OF _ f]) |
|
4639 then show ?thesis |
|
4640 by (simp add: comp_def that) |
|
4641 qed |
|
4642 |
|
4643 lemma nullhomotopic_into_contractible: |
|
4644 assumes f: "continuous_on S f" "f ` S \<subseteq> T" |
|
4645 and T: "contractible T" |
|
4646 obtains c where "homotopic_with (\<lambda>h. True) S T f (\<lambda>x. c)" |
|
4647 apply (rule nullhomotopic_through_contractible [OF f, of id T]) |
|
4648 using assms |
|
4649 apply (auto simp: continuous_on_id) |
|
4650 done |
|
4651 |
|
4652 lemma nullhomotopic_from_contractible: |
|
4653 assumes f: "continuous_on S f" "f ` S \<subseteq> T" |
|
4654 and S: "contractible S" |
|
4655 obtains c where "homotopic_with (\<lambda>h. True) S T f (\<lambda>x. c)" |
|
4656 apply (rule nullhomotopic_through_contractible [OF continuous_on_id _ f S, of S]) |
|
4657 using assms |
|
4658 apply (auto simp: comp_def) |
|
4659 done |
|
4660 |
|
4661 lemma homotopic_through_contractible: |
|
4662 fixes S :: "_::real_normed_vector set" |
|
4663 assumes "continuous_on S f1" "f1 ` S \<subseteq> T" |
|
4664 "continuous_on T g1" "g1 ` T \<subseteq> U" |
|
4665 "continuous_on S f2" "f2 ` S \<subseteq> T" |
|
4666 "continuous_on T g2" "g2 ` T \<subseteq> U" |
|
4667 "contractible T" "path_connected U" |
|
4668 shows "homotopic_with (\<lambda>h. True) S U (g1 \<circ> f1) (g2 \<circ> f2)" |
|
4669 proof - |
|
4670 obtain c1 where c1: "homotopic_with (\<lambda>h. True) S U (g1 \<circ> f1) (\<lambda>x. c1)" |
|
4671 apply (rule nullhomotopic_through_contractible [of S f1 T g1 U]) |
|
4672 using assms apply auto |
|
4673 done |
|
4674 obtain c2 where c2: "homotopic_with (\<lambda>h. True) S U (g2 \<circ> f2) (\<lambda>x. c2)" |
|
4675 apply (rule nullhomotopic_through_contractible [of S f2 T g2 U]) |
|
4676 using assms apply auto |
|
4677 done |
|
4678 have *: "S = {} \<or> (\<exists>t. path_connected t \<and> t \<subseteq> U \<and> c2 \<in> t \<and> c1 \<in> t)" |
|
4679 proof (cases "S = {}") |
|
4680 case True then show ?thesis by force |
|
4681 next |
|
4682 case False |
|
4683 with c1 c2 have "c1 \<in> U" "c2 \<in> U" |
|
4684 using homotopic_with_imp_subset2 all_not_in_conv image_subset_iff by blast+ |
|
4685 with \<open>path_connected U\<close> show ?thesis by blast |
|
4686 qed |
|
4687 show ?thesis |
|
4688 apply (rule homotopic_with_trans [OF c1]) |
|
4689 apply (rule homotopic_with_symD) |
|
4690 apply (rule homotopic_with_trans [OF c2]) |
|
4691 apply (simp add: path_component homotopic_constant_maps *) |
|
4692 done |
|
4693 qed |
|
4694 |
|
4695 lemma homotopic_into_contractible: |
|
4696 fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" |
|
4697 assumes f: "continuous_on S f" "f ` S \<subseteq> T" |
|
4698 and g: "continuous_on S g" "g ` S \<subseteq> T" |
|
4699 and T: "contractible T" |
|
4700 shows "homotopic_with (\<lambda>h. True) S T f g" |
|
4701 using homotopic_through_contractible [of S f T id T g id] |
|
4702 by (simp add: assms contractible_imp_path_connected continuous_on_id) |
|
4703 |
|
4704 lemma homotopic_from_contractible: |
|
4705 fixes S :: "'a::real_normed_vector set" and T:: "'b::real_normed_vector set" |
|
4706 assumes f: "continuous_on S f" "f ` S \<subseteq> T" |
|
4707 and g: "continuous_on S g" "g ` S \<subseteq> T" |
|
4708 and "contractible S" "path_connected T" |
|
4709 shows "homotopic_with (\<lambda>h. True) S T f g" |
|
4710 using homotopic_through_contractible [of S id S f T id g] |
|
4711 by (simp add: assms contractible_imp_path_connected continuous_on_id) |
|
4712 |
|
4713 lemma starlike_imp_contractible_gen: |
|
4714 fixes S :: "'a::real_normed_vector set" |
|
4715 assumes S: "starlike S" |
|
4716 and P: "\<And>a T. \<lbrakk>a \<in> S; 0 \<le> T; T \<le> 1\<rbrakk> \<Longrightarrow> P(\<lambda>x. (1 - T) *\<^sub>R x + T *\<^sub>R a)" |
|
4717 obtains a where "homotopic_with P S S (\<lambda>x. x) (\<lambda>x. a)" |
|
4718 proof - |
|
4719 obtain a where "a \<in> S" and a: "\<And>x. x \<in> S \<Longrightarrow> closed_segment a x \<subseteq> S" |
|
4720 using S by (auto simp: starlike_def) |
|
4721 have "(\<lambda>y. (1 - fst y) *\<^sub>R snd y + fst y *\<^sub>R a) ` ({0..1} \<times> S) \<subseteq> S" |
|
4722 apply clarify |
|
4723 apply (erule a [unfolded closed_segment_def, THEN subsetD], simp) |
|
4724 apply (metis add_diff_cancel_right' diff_ge_0_iff_ge le_add_diff_inverse pth_c(1)) |
|
4725 done |
|
4726 then show ?thesis |
|
4727 apply (rule_tac a=a in that) |
|
4728 using \<open>a \<in> S\<close> |
|
4729 apply (simp add: homotopic_with_def) |
|
4730 apply (rule_tac x="\<lambda>y. (1 - (fst y)) *\<^sub>R snd y + (fst y) *\<^sub>R a" in exI) |
|
4731 apply (intro conjI ballI continuous_on_compose continuous_intros) |
|
4732 apply (simp_all add: P) |
|
4733 done |
|
4734 qed |
|
4735 |
|
4736 lemma starlike_imp_contractible: |
|
4737 fixes S :: "'a::real_normed_vector set" |
|
4738 shows "starlike S \<Longrightarrow> contractible S" |
|
4739 using starlike_imp_contractible_gen contractible_def by (fastforce simp: id_def) |
|
4740 |
|
4741 lemma contractible_UNIV [simp]: "contractible (UNIV :: 'a::real_normed_vector set)" |
|
4742 by (simp add: starlike_imp_contractible) |
|
4743 |
|
4744 lemma starlike_imp_simply_connected: |
|
4745 fixes S :: "'a::real_normed_vector set" |
|
4746 shows "starlike S \<Longrightarrow> simply_connected S" |
|
4747 by (simp add: contractible_imp_simply_connected starlike_imp_contractible) |
|
4748 |
|
4749 lemma convex_imp_simply_connected: |
|
4750 fixes S :: "'a::real_normed_vector set" |
|
4751 shows "convex S \<Longrightarrow> simply_connected S" |
|
4752 using convex_imp_starlike starlike_imp_simply_connected by blast |
|
4753 |
|
4754 lemma starlike_imp_path_connected: |
|
4755 fixes S :: "'a::real_normed_vector set" |
|
4756 shows "starlike S \<Longrightarrow> path_connected S" |
|
4757 by (simp add: simply_connected_imp_path_connected starlike_imp_simply_connected) |
|
4758 |
|
4759 lemma starlike_imp_connected: |
|
4760 fixes S :: "'a::real_normed_vector set" |
|
4761 shows "starlike S \<Longrightarrow> connected S" |
|
4762 by (simp add: path_connected_imp_connected starlike_imp_path_connected) |
|
4763 |
|
4764 lemma is_interval_simply_connected_1: |
|
4765 fixes S :: "real set" |
|
4766 shows "is_interval S \<longleftrightarrow> simply_connected S" |
|
4767 using convex_imp_simply_connected is_interval_convex_1 is_interval_path_connected_1 simply_connected_imp_path_connected by auto |
|
4768 |
|
4769 lemma contractible_empty [simp]: "contractible {}" |
|
4770 by (simp add: contractible_def homotopic_with) |
|
4771 |
|
4772 lemma contractible_convex_tweak_boundary_points: |
|
4773 fixes S :: "'a::euclidean_space set" |
|
4774 assumes "convex S" and TS: "rel_interior S \<subseteq> T" "T \<subseteq> closure S" |
|
4775 shows "contractible T" |
|
4776 proof (cases "S = {}") |
|
4777 case True |
|
4778 with assms show ?thesis |
|
4779 by (simp add: subsetCE) |
|
4780 next |
|
4781 case False |
|
4782 show ?thesis |
|
4783 apply (rule starlike_imp_contractible) |
|
4784 apply (rule starlike_convex_tweak_boundary_points [OF \<open>convex S\<close> False TS]) |
|
4785 done |
|
4786 qed |
|
4787 |
|
4788 lemma convex_imp_contractible: |
|
4789 fixes S :: "'a::real_normed_vector set" |
|
4790 shows "convex S \<Longrightarrow> contractible S" |
|
4791 using contractible_empty convex_imp_starlike starlike_imp_contractible by blast |
|
4792 |
|
4793 lemma contractible_sing [simp]: |
|
4794 fixes a :: "'a::real_normed_vector" |
|
4795 shows "contractible {a}" |
|
4796 by (rule convex_imp_contractible [OF convex_singleton]) |
|
4797 |
|
4798 lemma is_interval_contractible_1: |
|
4799 fixes S :: "real set" |
|
4800 shows "is_interval S \<longleftrightarrow> contractible S" |
|
4801 using contractible_imp_simply_connected convex_imp_contractible is_interval_convex_1 |
|
4802 is_interval_simply_connected_1 by auto |
|
4803 |
|
4804 lemma contractible_Times: |
|
4805 fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" |
|
4806 assumes S: "contractible S" and T: "contractible T" |
|
4807 shows "contractible (S \<times> T)" |
|
4808 proof - |
|
4809 obtain a h where conth: "continuous_on ({0..1} \<times> S) h" |
|
4810 and hsub: "h ` ({0..1} \<times> S) \<subseteq> S" |
|
4811 and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (0, x) = x" |
|
4812 and [simp]: "\<And>x. x \<in> S \<Longrightarrow> h (1::real, x) = a" |
|
4813 using S by (auto simp: contractible_def homotopic_with) |
|
4814 obtain b k where contk: "continuous_on ({0..1} \<times> T) k" |
|
4815 and ksub: "k ` ({0..1} \<times> T) \<subseteq> T" |
|
4816 and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (0, x) = x" |
|
4817 and [simp]: "\<And>x. x \<in> T \<Longrightarrow> k (1::real, x) = b" |
|
4818 using T by (auto simp: contractible_def homotopic_with) |
|
4819 show ?thesis |
|
4820 apply (simp add: contractible_def homotopic_with) |
|
4821 apply (rule exI [where x=a]) |
|
4822 apply (rule exI [where x=b]) |
|
4823 apply (rule exI [where x = "\<lambda>z. (h (fst z, fst(snd z)), k (fst z, snd(snd z)))"]) |
|
4824 apply (intro conjI ballI continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF contk]) |
|
4825 using hsub ksub |
|
4826 apply auto |
|
4827 done |
|
4828 qed |
|
4829 |
|
4830 lemma homotopy_dominated_contractibility: |
|
4831 fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
4832 assumes S: "contractible S" |
|
4833 and f: "continuous_on S f" "image f S \<subseteq> T" |
|
4834 and g: "continuous_on T g" "image g T \<subseteq> S" |
|
4835 and hom: "homotopic_with (\<lambda>x. True) T T (f \<circ> g) id" |
|
4836 shows "contractible T" |
|
4837 proof - |
|
4838 obtain b where "homotopic_with (\<lambda>h. True) S T f (\<lambda>x. b)" |
|
4839 using nullhomotopic_from_contractible [OF f S] . |
|
4840 then have homg: "homotopic_with (\<lambda>x. True) T T ((\<lambda>x. b) \<circ> g) (f \<circ> g)" |
|
4841 by (rule homotopic_with_compose_continuous_right [OF homotopic_with_symD g]) |
|
4842 show ?thesis |
|
4843 apply (simp add: contractible_def) |
|
4844 apply (rule exI [where x = b]) |
|
4845 apply (rule homotopic_with_symD) |
|
4846 apply (rule homotopic_with_trans [OF _ hom]) |
|
4847 using homg apply (simp add: o_def) |
|
4848 done |
|
4849 qed |
|
4850 |
|
4851 |
|
4852 subsection\<open>Local versions of topological properties in general\<close> |
|
4853 |
|
4854 definition%important locally :: "('a::topological_space set \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool" |
|
4855 where |
|
4856 "locally P S \<equiv> |
|
4857 \<forall>w x. openin (subtopology euclidean S) w \<and> x \<in> w |
|
4858 \<longrightarrow> (\<exists>u v. openin (subtopology euclidean S) u \<and> P v \<and> |
|
4859 x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w)" |
|
4860 |
|
4861 lemma locallyI: |
|
4862 assumes "\<And>w x. \<lbrakk>openin (subtopology euclidean S) w; x \<in> w\<rbrakk> |
|
4863 \<Longrightarrow> \<exists>u v. openin (subtopology euclidean S) u \<and> P v \<and> |
|
4864 x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w" |
|
4865 shows "locally P S" |
|
4866 using assms by (force simp: locally_def) |
|
4867 |
|
4868 lemma locallyE: |
|
4869 assumes "locally P S" "openin (subtopology euclidean S) w" "x \<in> w" |
|
4870 obtains u v where "openin (subtopology euclidean S) u" |
|
4871 "P v" "x \<in> u" "u \<subseteq> v" "v \<subseteq> w" |
|
4872 using assms unfolding locally_def by meson |
|
4873 |
|
4874 lemma locally_mono: |
|
4875 assumes "locally P S" "\<And>t. P t \<Longrightarrow> Q t" |
|
4876 shows "locally Q S" |
|
4877 by (metis assms locally_def) |
|
4878 |
|
4879 lemma locally_open_subset: |
|
4880 assumes "locally P S" "openin (subtopology euclidean S) t" |
|
4881 shows "locally P t" |
|
4882 using assms |
|
4883 apply (simp add: locally_def) |
|
4884 apply (erule all_forward)+ |
|
4885 apply (rule impI) |
|
4886 apply (erule impCE) |
|
4887 using openin_trans apply blast |
|
4888 apply (erule ex_forward) |
|
4889 by (metis (no_types, hide_lams) Int_absorb1 Int_lower1 Int_subset_iff openin_open openin_subtopology_Int_subset) |
|
4890 |
|
4891 lemma locally_diff_closed: |
|
4892 "\<lbrakk>locally P S; closedin (subtopology euclidean S) t\<rbrakk> \<Longrightarrow> locally P (S - t)" |
|
4893 using locally_open_subset closedin_def by fastforce |
|
4894 |
|
4895 lemma locally_empty [iff]: "locally P {}" |
|
4896 by (simp add: locally_def openin_subtopology) |
|
4897 |
|
4898 lemma locally_singleton [iff]: |
|
4899 fixes a :: "'a::metric_space" |
|
4900 shows "locally P {a} \<longleftrightarrow> P {a}" |
|
4901 apply (simp add: locally_def openin_euclidean_subtopology_iff subset_singleton_iff conj_disj_distribR cong: conj_cong) |
|
4902 using zero_less_one by blast |
|
4903 |
|
4904 lemma locally_iff: |
|
4905 "locally P S \<longleftrightarrow> |
|
4906 (\<forall>T x. open T \<and> x \<in> S \<inter> T \<longrightarrow> (\<exists>U. open U \<and> (\<exists>v. P v \<and> x \<in> S \<inter> U \<and> S \<inter> U \<subseteq> v \<and> v \<subseteq> S \<inter> T)))" |
|
4907 apply (simp add: le_inf_iff locally_def openin_open, safe) |
|
4908 apply (metis IntE IntI le_inf_iff) |
|
4909 apply (metis IntI Int_subset_iff) |
|
4910 done |
|
4911 |
|
4912 lemma locally_Int: |
|
4913 assumes S: "locally P S" and t: "locally P t" |
|
4914 and P: "\<And>S t. P S \<and> P t \<Longrightarrow> P(S \<inter> t)" |
|
4915 shows "locally P (S \<inter> t)" |
|
4916 using S t unfolding locally_iff |
|
4917 apply clarify |
|
4918 apply (drule_tac x=T in spec)+ |
|
4919 apply (drule_tac x=x in spec)+ |
|
4920 apply clarsimp |
|
4921 apply (rename_tac U1 U2 V1 V2) |
|
4922 apply (rule_tac x="U1 \<inter> U2" in exI) |
|
4923 apply (simp add: open_Int) |
|
4924 apply (rule_tac x="V1 \<inter> V2" in exI) |
|
4925 apply (auto intro: P) |
|
4926 done |
|
4927 |
|
4928 lemma locally_Times: |
|
4929 fixes S :: "('a::metric_space) set" and T :: "('b::metric_space) set" |
|
4930 assumes PS: "locally P S" and QT: "locally Q T" and R: "\<And>S T. P S \<and> Q T \<Longrightarrow> R(S \<times> T)" |
|
4931 shows "locally R (S \<times> T)" |
|
4932 unfolding locally_def |
|
4933 proof (clarify) |
|
4934 fix W x y |
|
4935 assume W: "openin (subtopology euclidean (S \<times> T)) W" and xy: "(x, y) \<in> W" |
|
4936 then obtain U V where "openin (subtopology euclidean S) U" "x \<in> U" |
|
4937 "openin (subtopology euclidean T) V" "y \<in> V" "U \<times> V \<subseteq> W" |
|
4938 using Times_in_interior_subtopology by metis |
|
4939 then obtain U1 U2 V1 V2 |
|
4940 where opeS: "openin (subtopology euclidean S) U1 \<and> P U2 \<and> x \<in> U1 \<and> U1 \<subseteq> U2 \<and> U2 \<subseteq> U" |
|
4941 and opeT: "openin (subtopology euclidean T) V1 \<and> Q V2 \<and> y \<in> V1 \<and> V1 \<subseteq> V2 \<and> V2 \<subseteq> V" |
|
4942 by (meson PS QT locallyE) |
|
4943 with \<open>U \<times> V \<subseteq> W\<close> show "\<exists>u v. openin (subtopology euclidean (S \<times> T)) u \<and> R v \<and> (x,y) \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> W" |
|
4944 apply (rule_tac x="U1 \<times> V1" in exI) |
|
4945 apply (rule_tac x="U2 \<times> V2" in exI) |
|
4946 apply (auto simp: openin_Times R) |
|
4947 done |
|
4948 qed |
|
4949 |
|
4950 |
|
4951 proposition homeomorphism_locally_imp: |
|
4952 fixes S :: "'a::metric_space set" and t :: "'b::t2_space set" |
|
4953 assumes S: "locally P S" and hom: "homeomorphism S t f g" |
|
4954 and Q: "\<And>S S'. \<lbrakk>P S; homeomorphism S S' f g\<rbrakk> \<Longrightarrow> Q S'" |
|
4955 shows "locally Q t" |
|
4956 proof (clarsimp simp: locally_def) |
|
4957 fix W y |
|
4958 assume "y \<in> W" and "openin (subtopology euclidean t) W" |
|
4959 then obtain T where T: "open T" "W = t \<inter> T" |
|
4960 by (force simp: openin_open) |
|
4961 then have "W \<subseteq> t" by auto |
|
4962 have f: "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "f ` S = t" "continuous_on S f" |
|
4963 and g: "\<And>y. y \<in> t \<Longrightarrow> f(g y) = y" "g ` t = S" "continuous_on t g" |
|
4964 using hom by (auto simp: homeomorphism_def) |
|
4965 have gw: "g ` W = S \<inter> f -` W" |
|
4966 using \<open>W \<subseteq> t\<close> |
|
4967 apply auto |
|
4968 using \<open>g ` t = S\<close> \<open>W \<subseteq> t\<close> apply blast |
|
4969 using g \<open>W \<subseteq> t\<close> apply auto[1] |
|
4970 by (simp add: f rev_image_eqI) |
|
4971 have \<circ>: "openin (subtopology euclidean S) (g ` W)" |
|
4972 proof - |
|
4973 have "continuous_on S f" |
|
4974 using f(3) by blast |
|
4975 then show "openin (subtopology euclidean S) (g ` W)" |
|
4976 by (simp add: gw Collect_conj_eq \<open>openin (subtopology euclidean t) W\<close> continuous_on_open f(2)) |
|
4977 qed |
|
4978 then obtain u v |
|
4979 where osu: "openin (subtopology euclidean S) u" and uv: "P v" "g y \<in> u" "u \<subseteq> v" "v \<subseteq> g ` W" |
|
4980 using S [unfolded locally_def, rule_format, of "g ` W" "g y"] \<open>y \<in> W\<close> by force |
|
4981 have "v \<subseteq> S" using uv by (simp add: gw) |
|
4982 have fv: "f ` v = t \<inter> {x. g x \<in> v}" |
|
4983 using \<open>f ` S = t\<close> f \<open>v \<subseteq> S\<close> by auto |
|
4984 have "f ` v \<subseteq> W" |
|
4985 using uv using Int_lower2 gw image_subsetI mem_Collect_eq subset_iff by auto |
|
4986 have contvf: "continuous_on v f" |
|
4987 using \<open>v \<subseteq> S\<close> continuous_on_subset f(3) by blast |
|
4988 have contvg: "continuous_on (f ` v) g" |
|
4989 using \<open>f ` v \<subseteq> W\<close> \<open>W \<subseteq> t\<close> continuous_on_subset [OF g(3)] by blast |
|
4990 have homv: "homeomorphism v (f ` v) f g" |
|
4991 using \<open>v \<subseteq> S\<close> \<open>W \<subseteq> t\<close> f |
|
4992 apply (simp add: homeomorphism_def contvf contvg, auto) |
|
4993 by (metis f(1) rev_image_eqI rev_subsetD) |
|
4994 have 1: "openin (subtopology euclidean t) (t \<inter> g -` u)" |
|
4995 apply (rule continuous_on_open [THEN iffD1, rule_format]) |
|
4996 apply (rule \<open>continuous_on t g\<close>) |
|
4997 using \<open>g ` t = S\<close> apply (simp add: osu) |
|
4998 done |
|
4999 have 2: "\<exists>V. Q V \<and> y \<in> (t \<inter> g -` u) \<and> (t \<inter> g -` u) \<subseteq> V \<and> V \<subseteq> W" |
|
5000 apply (rule_tac x="f ` v" in exI) |
|
5001 apply (intro conjI Q [OF \<open>P v\<close> homv]) |
|
5002 using \<open>W \<subseteq> t\<close> \<open>y \<in> W\<close> \<open>f ` v \<subseteq> W\<close> uv apply (auto simp: fv) |
|
5003 done |
|
5004 show "\<exists>U. openin (subtopology euclidean t) U \<and> (\<exists>v. Q v \<and> y \<in> U \<and> U \<subseteq> v \<and> v \<subseteq> W)" |
|
5005 by (meson 1 2) |
|
5006 qed |
|
5007 |
|
5008 lemma homeomorphism_locally: |
|
5009 fixes f:: "'a::metric_space \<Rightarrow> 'b::metric_space" |
|
5010 assumes hom: "homeomorphism S t f g" |
|
5011 and eq: "\<And>S t. homeomorphism S t f g \<Longrightarrow> (P S \<longleftrightarrow> Q t)" |
|
5012 shows "locally P S \<longleftrightarrow> locally Q t" |
|
5013 apply (rule iffI) |
|
5014 apply (erule homeomorphism_locally_imp [OF _ hom]) |
|
5015 apply (simp add: eq) |
|
5016 apply (erule homeomorphism_locally_imp) |
|
5017 using eq homeomorphism_sym homeomorphism_symD [OF hom] apply blast+ |
|
5018 done |
|
5019 |
|
5020 lemma homeomorphic_locally: |
|
5021 fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" |
|
5022 assumes hom: "S homeomorphic T" |
|
5023 and iff: "\<And>X Y. X homeomorphic Y \<Longrightarrow> (P X \<longleftrightarrow> Q Y)" |
|
5024 shows "locally P S \<longleftrightarrow> locally Q T" |
|
5025 proof - |
|
5026 obtain f g where hom: "homeomorphism S T f g" |
|
5027 using assms by (force simp: homeomorphic_def) |
|
5028 then show ?thesis |
|
5029 using homeomorphic_def local.iff |
|
5030 by (blast intro!: homeomorphism_locally) |
|
5031 qed |
|
5032 |
|
5033 lemma homeomorphic_local_compactness: |
|
5034 fixes S:: "'a::metric_space set" and T:: "'b::metric_space set" |
|
5035 shows "S homeomorphic T \<Longrightarrow> locally compact S \<longleftrightarrow> locally compact T" |
|
5036 by (simp add: homeomorphic_compactness homeomorphic_locally) |
|
5037 |
|
5038 lemma locally_translation: |
|
5039 fixes P :: "'a :: real_normed_vector set \<Rightarrow> bool" |
|
5040 shows |
|
5041 "(\<And>S. P (image (\<lambda>x. a + x) S) \<longleftrightarrow> P S) |
|
5042 \<Longrightarrow> locally P (image (\<lambda>x. a + x) S) \<longleftrightarrow> locally P S" |
|
5043 apply (rule homeomorphism_locally [OF homeomorphism_translation]) |
|
5044 apply (simp add: homeomorphism_def) |
|
5045 by metis |
|
5046 |
|
5047 lemma locally_injective_linear_image: |
|
5048 fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
5049 assumes f: "linear f" "inj f" and iff: "\<And>S. P (f ` S) \<longleftrightarrow> Q S" |
|
5050 shows "locally P (f ` S) \<longleftrightarrow> locally Q S" |
|
5051 apply (rule linear_homeomorphism_image [OF f]) |
|
5052 apply (rule_tac f=g and g = f in homeomorphism_locally, assumption) |
|
5053 by (metis iff homeomorphism_def) |
|
5054 |
|
5055 lemma locally_open_map_image: |
|
5056 fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_vector" |
|
5057 assumes P: "locally P S" |
|
5058 and f: "continuous_on S f" |
|
5059 and oo: "\<And>t. openin (subtopology euclidean S) t |
|
5060 \<Longrightarrow> openin (subtopology euclidean (f ` S)) (f ` t)" |
|
5061 and Q: "\<And>t. \<lbrakk>t \<subseteq> S; P t\<rbrakk> \<Longrightarrow> Q(f ` t)" |
|
5062 shows "locally Q (f ` S)" |
|
5063 proof (clarsimp simp add: locally_def) |
|
5064 fix W y |
|
5065 assume oiw: "openin (subtopology euclidean (f ` S)) W" and "y \<in> W" |
|
5066 then have "W \<subseteq> f ` S" by (simp add: openin_euclidean_subtopology_iff) |
|
5067 have oivf: "openin (subtopology euclidean S) (S \<inter> f -` W)" |
|
5068 by (rule continuous_on_open [THEN iffD1, rule_format, OF f oiw]) |
|
5069 then obtain x where "x \<in> S" "f x = y" |
|
5070 using \<open>W \<subseteq> f ` S\<close> \<open>y \<in> W\<close> by blast |
|
5071 then obtain U V |
|
5072 where "openin (subtopology euclidean S) U" "P V" "x \<in> U" "U \<subseteq> V" "V \<subseteq> S \<inter> f -` W" |
|
5073 using P [unfolded locally_def, rule_format, of "(S \<inter> f -` W)" x] oivf \<open>y \<in> W\<close> |
|
5074 by auto |
|
5075 then show "\<exists>X. openin (subtopology euclidean (f ` S)) X \<and> (\<exists>Y. Q Y \<and> y \<in> X \<and> X \<subseteq> Y \<and> Y \<subseteq> W)" |
|
5076 apply (rule_tac x="f ` U" in exI) |
|
5077 apply (rule conjI, blast intro!: oo) |
|
5078 apply (rule_tac x="f ` V" in exI) |
|
5079 apply (force simp: \<open>f x = y\<close> rev_image_eqI intro: Q) |
|
5080 done |
|
5081 qed |
|
5082 |
|
5083 |
|
5084 subsection\<open>An induction principle for connected sets\<close> |
|
5085 |
|
5086 proposition connected_induction: |
|
5087 assumes "connected S" |
|
5088 and opD: "\<And>T a. \<lbrakk>openin (subtopology euclidean S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z" |
|
5089 and opI: "\<And>a. a \<in> S |
|
5090 \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and> |
|
5091 (\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<and> Q x \<longrightarrow> Q y)" |
|
5092 and etc: "a \<in> S" "b \<in> S" "P a" "P b" "Q a" |
|
5093 shows "Q b" |
|
5094 proof - |
|
5095 have 1: "openin (subtopology euclidean S) |
|
5096 {b. \<exists>T. openin (subtopology euclidean S) T \<and> |
|
5097 b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> Q x)}" |
|
5098 apply (subst openin_subopen, clarify) |
|
5099 apply (rule_tac x=T in exI, auto) |
|
5100 done |
|
5101 have 2: "openin (subtopology euclidean S) |
|
5102 {b. \<exists>T. openin (subtopology euclidean S) T \<and> |
|
5103 b \<in> T \<and> (\<forall>x\<in>T. P x \<longrightarrow> \<not> Q x)}" |
|
5104 apply (subst openin_subopen, clarify) |
|
5105 apply (rule_tac x=T in exI, auto) |
|
5106 done |
|
5107 show ?thesis |
|
5108 using \<open>connected S\<close> |
|
5109 apply (simp only: connected_openin HOL.not_ex HOL.de_Morgan_conj) |
|
5110 apply (elim disjE allE) |
|
5111 apply (blast intro: 1) |
|
5112 apply (blast intro: 2, simp_all) |
|
5113 apply clarify apply (metis opI) |
|
5114 using opD apply (blast intro: etc elim: dest:) |
|
5115 using opI etc apply meson+ |
|
5116 done |
|
5117 qed |
|
5118 |
|
5119 lemma connected_equivalence_relation_gen: |
|
5120 assumes "connected S" |
|
5121 and etc: "a \<in> S" "b \<in> S" "P a" "P b" |
|
5122 and trans: "\<And>x y z. \<lbrakk>R x y; R y z\<rbrakk> \<Longrightarrow> R x z" |
|
5123 and opD: "\<And>T a. \<lbrakk>openin (subtopology euclidean S) T; a \<in> T\<rbrakk> \<Longrightarrow> \<exists>z. z \<in> T \<and> P z" |
|
5124 and opI: "\<And>a. a \<in> S |
|
5125 \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and> |
|
5126 (\<forall>x \<in> T. \<forall>y \<in> T. P x \<and> P y \<longrightarrow> R x y)" |
|
5127 shows "R a b" |
|
5128 proof - |
|
5129 have "\<And>a b c. \<lbrakk>a \<in> S; P a; b \<in> S; c \<in> S; P b; P c; R a b\<rbrakk> \<Longrightarrow> R a c" |
|
5130 apply (rule connected_induction [OF \<open>connected S\<close> opD], simp_all) |
|
5131 by (meson trans opI) |
|
5132 then show ?thesis by (metis etc opI) |
|
5133 qed |
|
5134 |
|
5135 lemma connected_induction_simple: |
|
5136 assumes "connected S" |
|
5137 and etc: "a \<in> S" "b \<in> S" "P a" |
|
5138 and opI: "\<And>a. a \<in> S |
|
5139 \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and> |
|
5140 (\<forall>x \<in> T. \<forall>y \<in> T. P x \<longrightarrow> P y)" |
|
5141 shows "P b" |
|
5142 apply (rule connected_induction [OF \<open>connected S\<close> _, where P = "\<lambda>x. True"], blast) |
|
5143 apply (frule opI) |
|
5144 using etc apply simp_all |
|
5145 done |
|
5146 |
|
5147 lemma connected_equivalence_relation: |
|
5148 assumes "connected S" |
|
5149 and etc: "a \<in> S" "b \<in> S" |
|
5150 and sym: "\<And>x y. \<lbrakk>R x y; x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> R y x" |
|
5151 and trans: "\<And>x y z. \<lbrakk>R x y; R y z; x \<in> S; y \<in> S; z \<in> S\<rbrakk> \<Longrightarrow> R x z" |
|
5152 and opI: "\<And>a. a \<in> S \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. R a x)" |
|
5153 shows "R a b" |
|
5154 proof - |
|
5155 have "\<And>a b c. \<lbrakk>a \<in> S; b \<in> S; c \<in> S; R a b\<rbrakk> \<Longrightarrow> R a c" |
|
5156 apply (rule connected_induction_simple [OF \<open>connected S\<close>], simp_all) |
|
5157 by (meson local.sym local.trans opI openin_imp_subset subsetCE) |
|
5158 then show ?thesis by (metis etc opI) |
|
5159 qed |
|
5160 |
|
5161 lemma locally_constant_imp_constant: |
|
5162 assumes "connected S" |
|
5163 and opI: "\<And>a. a \<in> S |
|
5164 \<Longrightarrow> \<exists>T. openin (subtopology euclidean S) T \<and> a \<in> T \<and> (\<forall>x \<in> T. f x = f a)" |
|
5165 shows "f constant_on S" |
|
5166 proof - |
|
5167 have "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> f x = f y" |
|
5168 apply (rule connected_equivalence_relation [OF \<open>connected S\<close>], simp_all) |
|
5169 by (metis opI) |
|
5170 then show ?thesis |
|
5171 by (metis constant_on_def) |
|
5172 qed |
|
5173 |
|
5174 lemma locally_constant: |
|
5175 "connected S \<Longrightarrow> locally (\<lambda>U. f constant_on U) S \<longleftrightarrow> f constant_on S" |
|
5176 apply (simp add: locally_def) |
|
5177 apply (rule iffI) |
|
5178 apply (rule locally_constant_imp_constant, assumption) |
|
5179 apply (metis (mono_tags, hide_lams) constant_on_def constant_on_subset openin_subtopology_self) |
|
5180 by (meson constant_on_subset openin_imp_subset order_refl) |
|
5181 |
|
5182 |
|
5183 subsection\<open>Basic properties of local compactness\<close> |
|
5184 |
|
5185 proposition locally_compact: |
|
5186 fixes s :: "'a :: metric_space set" |
|
5187 shows |
|
5188 "locally compact s \<longleftrightarrow> |
|
5189 (\<forall>x \<in> s. \<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> |
|
5190 openin (subtopology euclidean s) u \<and> compact v)" |
|
5191 (is "?lhs = ?rhs") |
|
5192 proof |
|
5193 assume ?lhs |
|
5194 then show ?rhs |
|
5195 apply clarify |
|
5196 apply (erule_tac w = "s \<inter> ball x 1" in locallyE) |
|
5197 by auto |
|
5198 next |
|
5199 assume r [rule_format]: ?rhs |
|
5200 have *: "\<exists>u v. |
|
5201 openin (subtopology euclidean s) u \<and> |
|
5202 compact v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<inter> T" |
|
5203 if "open T" "x \<in> s" "x \<in> T" for x T |
|
5204 proof - |
|
5205 obtain u v where uv: "x \<in> u" "u \<subseteq> v" "v \<subseteq> s" "compact v" "openin (subtopology euclidean s) u" |
|
5206 using r [OF \<open>x \<in> s\<close>] by auto |
|
5207 obtain e where "e>0" and e: "cball x e \<subseteq> T" |
|
5208 using open_contains_cball \<open>open T\<close> \<open>x \<in> T\<close> by blast |
|
5209 show ?thesis |
|
5210 apply (rule_tac x="(s \<inter> ball x e) \<inter> u" in exI) |
|
5211 apply (rule_tac x="cball x e \<inter> v" in exI) |
|
5212 using that \<open>e > 0\<close> e uv |
|
5213 apply auto |
|
5214 done |
|
5215 qed |
|
5216 show ?lhs |
|
5217 apply (rule locallyI) |
|
5218 apply (subst (asm) openin_open) |
|
5219 apply (blast intro: *) |
|
5220 done |
|
5221 qed |
|
5222 |
|
5223 lemma locally_compactE: |
|
5224 fixes s :: "'a :: metric_space set" |
|
5225 assumes "locally compact s" |
|
5226 obtains u v where "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and> |
|
5227 openin (subtopology euclidean s) (u x) \<and> compact (v x)" |
|
5228 using assms |
|
5229 unfolding locally_compact by metis |
|
5230 |
|
5231 lemma locally_compact_alt: |
|
5232 fixes s :: "'a :: heine_borel set" |
|
5233 shows "locally compact s \<longleftrightarrow> |
|
5234 (\<forall>x \<in> s. \<exists>u. x \<in> u \<and> |
|
5235 openin (subtopology euclidean s) u \<and> compact(closure u) \<and> closure u \<subseteq> s)" |
|
5236 apply (simp add: locally_compact) |
|
5237 apply (intro ball_cong ex_cong refl iffI) |
|
5238 apply (metis bounded_subset closure_eq closure_mono compact_eq_bounded_closed dual_order.trans) |
|
5239 by (meson closure_subset compact_closure) |
|
5240 |
|
5241 lemma locally_compact_Int_cball: |
|
5242 fixes s :: "'a :: heine_borel set" |
|
5243 shows "locally compact s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>e. 0 < e \<and> closed(cball x e \<inter> s))" |
|
5244 (is "?lhs = ?rhs") |
|
5245 proof |
|
5246 assume ?lhs |
|
5247 then show ?rhs |
|
5248 apply (simp add: locally_compact openin_contains_cball) |
|
5249 apply (clarify | assumption | drule bspec)+ |
|
5250 by (metis (no_types, lifting) compact_cball compact_imp_closed compact_Int inf.absorb_iff2 inf.orderE inf_sup_aci(2)) |
|
5251 next |
|
5252 assume ?rhs |
|
5253 then show ?lhs |
|
5254 apply (simp add: locally_compact openin_contains_cball) |
|
5255 apply (clarify | assumption | drule bspec)+ |
|
5256 apply (rule_tac x="ball x e \<inter> s" in exI, simp) |
|
5257 apply (rule_tac x="cball x e \<inter> s" in exI) |
|
5258 using compact_eq_bounded_closed |
|
5259 apply auto |
|
5260 apply (metis open_ball le_infI1 mem_ball open_contains_cball_eq) |
|
5261 done |
|
5262 qed |
|
5263 |
|
5264 lemma locally_compact_compact: |
|
5265 fixes s :: "'a :: heine_borel set" |
|
5266 shows "locally compact s \<longleftrightarrow> |
|
5267 (\<forall>k. k \<subseteq> s \<and> compact k |
|
5268 \<longrightarrow> (\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> |
|
5269 openin (subtopology euclidean s) u \<and> compact v))" |
|
5270 (is "?lhs = ?rhs") |
|
5271 proof |
|
5272 assume ?lhs |
|
5273 then obtain u v where |
|
5274 uv: "\<And>x. x \<in> s \<Longrightarrow> x \<in> u x \<and> u x \<subseteq> v x \<and> v x \<subseteq> s \<and> |
|
5275 openin (subtopology euclidean s) (u x) \<and> compact (v x)" |
|
5276 by (metis locally_compactE) |
|
5277 have *: "\<exists>u v. k \<subseteq> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (subtopology euclidean s) u \<and> compact v" |
|
5278 if "k \<subseteq> s" "compact k" for k |
|
5279 proof - |
|
5280 have "\<And>C. (\<forall>c\<in>C. openin (subtopology euclidean k) c) \<and> k \<subseteq> \<Union>C \<Longrightarrow> |
|
5281 \<exists>D\<subseteq>C. finite D \<and> k \<subseteq> \<Union>D" |
|
5282 using that by (simp add: compact_eq_openin_cover) |
|
5283 moreover have "\<forall>c \<in> (\<lambda>x. k \<inter> u x) ` k. openin (subtopology euclidean k) c" |
|
5284 using that by clarify (metis subsetD inf.absorb_iff2 openin_subset openin_subtopology_Int_subset topspace_euclidean_subtopology uv) |
|
5285 moreover have "k \<subseteq> \<Union>((\<lambda>x. k \<inter> u x) ` k)" |
|
5286 using that by clarsimp (meson subsetCE uv) |
|
5287 ultimately obtain D where "D \<subseteq> (\<lambda>x. k \<inter> u x) ` k" "finite D" "k \<subseteq> \<Union>D" |
|
5288 by metis |
|
5289 then obtain T where T: "T \<subseteq> k" "finite T" "k \<subseteq> \<Union>((\<lambda>x. k \<inter> u x) ` T)" |
|
5290 by (metis finite_subset_image) |
|
5291 have Tuv: "\<Union>(u ` T) \<subseteq> \<Union>(v ` T)" |
|
5292 using T that by (force simp: dest!: uv) |
|
5293 show ?thesis |
|
5294 apply (rule_tac x="\<Union>(u ` T)" in exI) |
|
5295 apply (rule_tac x="\<Union>(v ` T)" in exI) |
|
5296 apply (simp add: Tuv) |
|
5297 using T that |
|
5298 apply (auto simp: dest!: uv) |
|
5299 done |
|
5300 qed |
|
5301 show ?rhs |
|
5302 by (blast intro: *) |
|
5303 next |
|
5304 assume ?rhs |
|
5305 then show ?lhs |
|
5306 apply (clarsimp simp add: locally_compact) |
|
5307 apply (drule_tac x="{x}" in spec, simp) |
|
5308 done |
|
5309 qed |
|
5310 |
|
5311 lemma open_imp_locally_compact: |
|
5312 fixes s :: "'a :: heine_borel set" |
|
5313 assumes "open s" |
|
5314 shows "locally compact s" |
|
5315 proof - |
|
5316 have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> openin (subtopology euclidean s) u \<and> compact v" |
|
5317 if "x \<in> s" for x |
|
5318 proof - |
|
5319 obtain e where "e>0" and e: "cball x e \<subseteq> s" |
|
5320 using open_contains_cball assms \<open>x \<in> s\<close> by blast |
|
5321 have ope: "openin (subtopology euclidean s) (ball x e)" |
|
5322 by (meson e open_ball ball_subset_cball dual_order.trans open_subset) |
|
5323 show ?thesis |
|
5324 apply (rule_tac x="ball x e" in exI) |
|
5325 apply (rule_tac x="cball x e" in exI) |
|
5326 using \<open>e > 0\<close> e apply (auto simp: ope) |
|
5327 done |
|
5328 qed |
|
5329 show ?thesis |
|
5330 unfolding locally_compact |
|
5331 by (blast intro: *) |
|
5332 qed |
|
5333 |
|
5334 lemma closed_imp_locally_compact: |
|
5335 fixes s :: "'a :: heine_borel set" |
|
5336 assumes "closed s" |
|
5337 shows "locally compact s" |
|
5338 proof - |
|
5339 have *: "\<exists>u v. x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> s \<and> |
|
5340 openin (subtopology euclidean s) u \<and> compact v" |
|
5341 if "x \<in> s" for x |
|
5342 proof - |
|
5343 show ?thesis |
|
5344 apply (rule_tac x = "s \<inter> ball x 1" in exI) |
|
5345 apply (rule_tac x = "s \<inter> cball x 1" in exI) |
|
5346 using \<open>x \<in> s\<close> assms apply auto |
|
5347 done |
|
5348 qed |
|
5349 show ?thesis |
|
5350 unfolding locally_compact |
|
5351 by (blast intro: *) |
|
5352 qed |
|
5353 |
|
5354 lemma locally_compact_UNIV: "locally compact (UNIV :: 'a :: heine_borel set)" |
|
5355 by (simp add: closed_imp_locally_compact) |
|
5356 |
|
5357 lemma locally_compact_Int: |
|
5358 fixes s :: "'a :: t2_space set" |
|
5359 shows "\<lbrakk>locally compact s; locally compact t\<rbrakk> \<Longrightarrow> locally compact (s \<inter> t)" |
|
5360 by (simp add: compact_Int locally_Int) |
|
5361 |
|
5362 lemma locally_compact_closedin: |
|
5363 fixes s :: "'a :: heine_borel set" |
|
5364 shows "\<lbrakk>closedin (subtopology euclidean s) t; locally compact s\<rbrakk> |
|
5365 \<Longrightarrow> locally compact t" |
|
5366 unfolding closedin_closed |
|
5367 using closed_imp_locally_compact locally_compact_Int by blast |
|
5368 |
|
5369 lemma locally_compact_delete: |
|
5370 fixes s :: "'a :: t1_space set" |
|
5371 shows "locally compact s \<Longrightarrow> locally compact (s - {a})" |
|
5372 by (auto simp: openin_delete locally_open_subset) |
|
5373 |
|
5374 lemma locally_closed: |
|
5375 fixes s :: "'a :: heine_borel set" |
|
5376 shows "locally closed s \<longleftrightarrow> locally compact s" |
|
5377 (is "?lhs = ?rhs") |
|
5378 proof |
|
5379 assume ?lhs |
|
5380 then show ?rhs |
|
5381 apply (simp only: locally_def) |
|
5382 apply (erule all_forward imp_forward asm_rl exE)+ |
|
5383 apply (rule_tac x = "u \<inter> ball x 1" in exI) |
|
5384 apply (rule_tac x = "v \<inter> cball x 1" in exI) |
|
5385 apply (force intro: openin_trans) |
|
5386 done |
|
5387 next |
|
5388 assume ?rhs then show ?lhs |
|
5389 using compact_eq_bounded_closed locally_mono by blast |
|
5390 qed |
|
5391 |
|
5392 lemma locally_compact_openin_Un: |
|
5393 fixes S :: "'a::euclidean_space set" |
|
5394 assumes LCS: "locally compact S" and LCT:"locally compact T" |
|
5395 and opS: "openin (subtopology euclidean (S \<union> T)) S" |
|
5396 and opT: "openin (subtopology euclidean (S \<union> T)) T" |
|
5397 shows "locally compact (S \<union> T)" |
|
5398 proof - |
|
5399 have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" for x |
|
5400 proof - |
|
5401 obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" |
|
5402 using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast |
|
5403 moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> S" |
|
5404 by (meson \<open>x \<in> S\<close> opS openin_contains_cball) |
|
5405 then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> S" |
|
5406 by force |
|
5407 ultimately show ?thesis |
|
5408 apply (rule_tac x="min e1 e2" in exI) |
|
5409 apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int) |
|
5410 by (metis closed_Int closed_cball inf_left_commute) |
|
5411 qed |
|
5412 moreover have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> T" for x |
|
5413 proof - |
|
5414 obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)" |
|
5415 using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast |
|
5416 moreover obtain e2 where "e2 > 0" and e2: "cball x e2 \<inter> (S \<union> T) \<subseteq> T" |
|
5417 by (meson \<open>x \<in> T\<close> opT openin_contains_cball) |
|
5418 then have "cball x e2 \<inter> (S \<union> T) = cball x e2 \<inter> T" |
|
5419 by force |
|
5420 ultimately show ?thesis |
|
5421 apply (rule_tac x="min e1 e2" in exI) |
|
5422 apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int) |
|
5423 by (metis closed_Int closed_cball inf_left_commute) |
|
5424 qed |
|
5425 ultimately show ?thesis |
|
5426 by (force simp: locally_compact_Int_cball) |
|
5427 qed |
|
5428 |
|
5429 lemma locally_compact_closedin_Un: |
|
5430 fixes S :: "'a::euclidean_space set" |
|
5431 assumes LCS: "locally compact S" and LCT:"locally compact T" |
|
5432 and clS: "closedin (subtopology euclidean (S \<union> T)) S" |
|
5433 and clT: "closedin (subtopology euclidean (S \<union> T)) T" |
|
5434 shows "locally compact (S \<union> T)" |
|
5435 proof - |
|
5436 have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if "x \<in> S" "x \<in> T" for x |
|
5437 proof - |
|
5438 obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" |
|
5439 using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast |
|
5440 moreover |
|
5441 obtain e2 where "e2 > 0" and e2: "closed (cball x e2 \<inter> T)" |
|
5442 using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast |
|
5443 ultimately show ?thesis |
|
5444 apply (rule_tac x="min e1 e2" in exI) |
|
5445 apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib) |
|
5446 by (metis closed_Int closed_Un closed_cball inf_left_commute) |
|
5447 qed |
|
5448 moreover |
|
5449 have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<in> S" "x \<notin> T" for x |
|
5450 proof - |
|
5451 obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> S)" |
|
5452 using LCS \<open>x \<in> S\<close> unfolding locally_compact_Int_cball by blast |
|
5453 moreover |
|
5454 obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S - T" |
|
5455 using clT x by (fastforce simp: openin_contains_cball closedin_def) |
|
5456 then have "closed (cball x e2 \<inter> T)" |
|
5457 proof - |
|
5458 have "{} = T - (T - cball x e2)" |
|
5459 using Diff_subset Int_Diff \<open>cball x e2 \<inter> (S \<union> T) \<subseteq> S - T\<close> by auto |
|
5460 then show ?thesis |
|
5461 by (simp add: Diff_Diff_Int inf_commute) |
|
5462 qed |
|
5463 ultimately show ?thesis |
|
5464 apply (rule_tac x="min e1 e2" in exI) |
|
5465 apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib) |
|
5466 by (metis closed_Int closed_Un closed_cball inf_left_commute) |
|
5467 qed |
|
5468 moreover |
|
5469 have "\<exists>e>0. closed (cball x e \<inter> (S \<union> T))" if x: "x \<notin> S" "x \<in> T" for x |
|
5470 proof - |
|
5471 obtain e1 where "e1 > 0" and e1: "closed (cball x e1 \<inter> T)" |
|
5472 using LCT \<open>x \<in> T\<close> unfolding locally_compact_Int_cball by blast |
|
5473 moreover |
|
5474 obtain e2 where "e2>0" and "cball x e2 \<inter> (S \<union> T) \<subseteq> S \<union> T - S" |
|
5475 using clS x by (fastforce simp: openin_contains_cball closedin_def) |
|
5476 then have "closed (cball x e2 \<inter> S)" |
|
5477 by (metis Diff_disjoint Int_empty_right closed_empty inf.left_commute inf.orderE inf_sup_absorb) |
|
5478 ultimately show ?thesis |
|
5479 apply (rule_tac x="min e1 e2" in exI) |
|
5480 apply (auto simp: cball_min_Int \<open>e2 > 0\<close> inf_assoc closed_Int Int_Un_distrib) |
|
5481 by (metis closed_Int closed_Un closed_cball inf_left_commute) |
|
5482 qed |
|
5483 ultimately show ?thesis |
|
5484 by (auto simp: locally_compact_Int_cball) |
|
5485 qed |
|
5486 |
|
5487 lemma locally_compact_Times: |
|
5488 fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" |
|
5489 shows "\<lbrakk>locally compact S; locally compact T\<rbrakk> \<Longrightarrow> locally compact (S \<times> T)" |
|
5490 by (auto simp: compact_Times locally_Times) |
|
5491 |
|
5492 lemma locally_compact_compact_subopen: |
|
5493 fixes S :: "'a :: heine_borel set" |
|
5494 shows |
|
5495 "locally compact S \<longleftrightarrow> |
|
5496 (\<forall>K T. K \<subseteq> S \<and> compact K \<and> open T \<and> K \<subseteq> T |
|
5497 \<longrightarrow> (\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and> |
|
5498 openin (subtopology euclidean S) U \<and> compact V))" |
|
5499 (is "?lhs = ?rhs") |
|
5500 proof |
|
5501 assume L: ?lhs |
|
5502 show ?rhs |
|
5503 proof clarify |
|
5504 fix K :: "'a set" and T :: "'a set" |
|
5505 assume "K \<subseteq> S" and "compact K" and "open T" and "K \<subseteq> T" |
|
5506 obtain U V where "K \<subseteq> U" "U \<subseteq> V" "V \<subseteq> S" "compact V" |
|
5507 and ope: "openin (subtopology euclidean S) U" |
|
5508 using L unfolding locally_compact_compact by (meson \<open>K \<subseteq> S\<close> \<open>compact K\<close>) |
|
5509 show "\<exists>U V. K \<subseteq> U \<and> U \<subseteq> V \<and> U \<subseteq> T \<and> V \<subseteq> S \<and> |
|
5510 openin (subtopology euclidean S) U \<and> compact V" |
|
5511 proof (intro exI conjI) |
|
5512 show "K \<subseteq> U \<inter> T" |
|
5513 by (simp add: \<open>K \<subseteq> T\<close> \<open>K \<subseteq> U\<close>) |
|
5514 show "U \<inter> T \<subseteq> closure(U \<inter> T)" |
|
5515 by (rule closure_subset) |
|
5516 show "closure (U \<inter> T) \<subseteq> S" |
|
5517 by (metis \<open>U \<subseteq> V\<close> \<open>V \<subseteq> S\<close> \<open>compact V\<close> closure_closed closure_mono compact_imp_closed inf.cobounded1 subset_trans) |
|
5518 show "openin (subtopology euclidean S) (U \<inter> T)" |
|
5519 by (simp add: \<open>open T\<close> ope openin_Int_open) |
|
5520 show "compact (closure (U \<inter> T))" |
|
5521 by (meson Int_lower1 \<open>U \<subseteq> V\<close> \<open>compact V\<close> bounded_subset compact_closure compact_eq_bounded_closed) |
|
5522 qed auto |
|
5523 qed |
|
5524 next |
|
5525 assume ?rhs then show ?lhs |
|
5526 unfolding locally_compact_compact |
|
5527 by (metis open_openin openin_topspace subtopology_superset top.extremum topspace_euclidean_subtopology) |
|
5528 qed |
|
5529 |
|
5530 |
|
5531 subsection\<open>Sura-Bura's results about compact components of sets\<close> |
|
5532 |
|
5533 proposition Sura_Bura_compact: |
|
5534 fixes S :: "'a::euclidean_space set" |
|
5535 assumes "compact S" and C: "C \<in> components S" |
|
5536 shows "C = \<Inter>{T. C \<subseteq> T \<and> openin (subtopology euclidean S) T \<and> |
|
5537 closedin (subtopology euclidean S) T}" |
|
5538 (is "C = \<Inter>?\<T>") |
|
5539 proof |
|
5540 obtain x where x: "C = connected_component_set S x" and "x \<in> S" |
|
5541 using C by (auto simp: components_def) |
|
5542 have "C \<subseteq> S" |
|
5543 by (simp add: C in_components_subset) |
|
5544 have "\<Inter>?\<T> \<subseteq> connected_component_set S x" |
|
5545 proof (rule connected_component_maximal) |
|
5546 have "x \<in> C" |
|
5547 by (simp add: \<open>x \<in> S\<close> x) |
|
5548 then show "x \<in> \<Inter>?\<T>" |
|
5549 by blast |
|
5550 have clo: "closed (\<Inter>?\<T>)" |
|
5551 by (simp add: \<open>compact S\<close> closed_Inter closedin_compact_eq compact_imp_closed) |
|
5552 have False |
|
5553 if K1: "closedin (subtopology euclidean (\<Inter>?\<T>)) K1" and |
|
5554 K2: "closedin (subtopology euclidean (\<Inter>?\<T>)) K2" and |
|
5555 K12_Int: "K1 \<inter> K2 = {}" and K12_Un: "K1 \<union> K2 = \<Inter>?\<T>" and "K1 \<noteq> {}" "K2 \<noteq> {}" |
|
5556 for K1 K2 |
|
5557 proof - |
|
5558 have "closed K1" "closed K2" |
|
5559 using closedin_closed_trans clo K1 K2 by blast+ |
|
5560 then obtain V1 V2 where "open V1" "open V2" "K1 \<subseteq> V1" "K2 \<subseteq> V2" and V12: "V1 \<inter> V2 = {}" |
|
5561 using separation_normal \<open>K1 \<inter> K2 = {}\<close> by metis |
|
5562 have SV12_ne: "(S - (V1 \<union> V2)) \<inter> (\<Inter>?\<T>) \<noteq> {}" |
|
5563 proof (rule compact_imp_fip) |
|
5564 show "compact (S - (V1 \<union> V2))" |
|
5565 by (simp add: \<open>open V1\<close> \<open>open V2\<close> \<open>compact S\<close> compact_diff open_Un) |
|
5566 show clo\<T>: "closed T" if "T \<in> ?\<T>" for T |
|
5567 using that \<open>compact S\<close> |
|
5568 by (force intro: closedin_closed_trans simp add: compact_imp_closed) |
|
5569 show "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> \<noteq> {}" if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F> |
|
5570 proof |
|
5571 assume djo: "(S - (V1 \<union> V2)) \<inter> \<Inter>\<F> = {}" |
|
5572 obtain D where opeD: "openin (subtopology euclidean S) D" |
|
5573 and cloD: "closedin (subtopology euclidean S) D" |
|
5574 and "C \<subseteq> D" and DV12: "D \<subseteq> V1 \<union> V2" |
|
5575 proof (cases "\<F> = {}") |
|
5576 case True |
|
5577 with \<open>C \<subseteq> S\<close> djo that show ?thesis |
|
5578 by force |
|
5579 next |
|
5580 case False show ?thesis |
|
5581 proof |
|
5582 show ope: "openin (subtopology euclidean S) (\<Inter>\<F>)" |
|
5583 using openin_Inter \<open>finite \<F>\<close> False \<F> by blast |
|
5584 then show "closedin (subtopology euclidean S) (\<Inter>\<F>)" |
|
5585 by (meson clo\<T> \<F> closed_Inter closed_subset openin_imp_subset subset_eq) |
|
5586 show "C \<subseteq> \<Inter>\<F>" |
|
5587 using \<F> by auto |
|
5588 show "\<Inter>\<F> \<subseteq> V1 \<union> V2" |
|
5589 using ope djo openin_imp_subset by fastforce |
|
5590 qed |
|
5591 qed |
|
5592 have "connected C" |
|
5593 by (simp add: x) |
|
5594 have "closed D" |
|
5595 using \<open>compact S\<close> cloD closedin_closed_trans compact_imp_closed by blast |
|
5596 have cloV1: "closedin (subtopology euclidean D) (D \<inter> closure V1)" |
|
5597 and cloV2: "closedin (subtopology euclidean D) (D \<inter> closure V2)" |
|
5598 by (simp_all add: closedin_closed_Int) |
|
5599 moreover have "D \<inter> closure V1 = D \<inter> V1" "D \<inter> closure V2 = D \<inter> V2" |
|
5600 apply safe |
|
5601 using \<open>D \<subseteq> V1 \<union> V2\<close> \<open>open V1\<close> \<open>open V2\<close> V12 |
|
5602 apply (simp_all add: closure_subset [THEN subsetD] closure_iff_nhds_not_empty, blast+) |
|
5603 done |
|
5604 ultimately have cloDV1: "closedin (subtopology euclidean D) (D \<inter> V1)" |
|
5605 and cloDV2: "closedin (subtopology euclidean D) (D \<inter> V2)" |
|
5606 by metis+ |
|
5607 then obtain U1 U2 where "closed U1" "closed U2" |
|
5608 and D1: "D \<inter> V1 = D \<inter> U1" and D2: "D \<inter> V2 = D \<inter> U2" |
|
5609 by (auto simp: closedin_closed) |
|
5610 have "D \<inter> U1 \<inter> C \<noteq> {}" |
|
5611 proof |
|
5612 assume "D \<inter> U1 \<inter> C = {}" |
|
5613 then have *: "C \<subseteq> D \<inter> V2" |
|
5614 using D1 DV12 \<open>C \<subseteq> D\<close> by auto |
|
5615 have "\<Inter>?\<T> \<subseteq> D \<inter> V2" |
|
5616 apply (rule Inter_lower) |
|
5617 using * apply simp |
|
5618 by (meson cloDV2 \<open>open V2\<close> cloD closedin_trans le_inf_iff opeD openin_Int_open) |
|
5619 then show False |
|
5620 using K1 V12 \<open>K1 \<noteq> {}\<close> \<open>K1 \<subseteq> V1\<close> closedin_imp_subset by blast |
|
5621 qed |
|
5622 moreover have "D \<inter> U2 \<inter> C \<noteq> {}" |
|
5623 proof |
|
5624 assume "D \<inter> U2 \<inter> C = {}" |
|
5625 then have *: "C \<subseteq> D \<inter> V1" |
|
5626 using D2 DV12 \<open>C \<subseteq> D\<close> by auto |
|
5627 have "\<Inter>?\<T> \<subseteq> D \<inter> V1" |
|
5628 apply (rule Inter_lower) |
|
5629 using * apply simp |
|
5630 by (meson cloDV1 \<open>open V1\<close> cloD closedin_trans le_inf_iff opeD openin_Int_open) |
|
5631 then show False |
|
5632 using K2 V12 \<open>K2 \<noteq> {}\<close> \<open>K2 \<subseteq> V2\<close> closedin_imp_subset by blast |
|
5633 qed |
|
5634 ultimately show False |
|
5635 using \<open>connected C\<close> unfolding connected_closed |
|
5636 apply (simp only: not_ex) |
|
5637 apply (drule_tac x="D \<inter> U1" in spec) |
|
5638 apply (drule_tac x="D \<inter> U2" in spec) |
|
5639 using \<open>C \<subseteq> D\<close> D1 D2 V12 DV12 \<open>closed U1\<close> \<open>closed U2\<close> \<open>closed D\<close> |
|
5640 by blast |
|
5641 qed |
|
5642 qed |
|
5643 show False |
|
5644 by (metis (full_types) DiffE UnE Un_upper2 SV12_ne \<open>K1 \<subseteq> V1\<close> \<open>K2 \<subseteq> V2\<close> disjoint_iff_not_equal subsetCE sup_ge1 K12_Un) |
|
5645 qed |
|
5646 then show "connected (\<Inter>?\<T>)" |
|
5647 by (auto simp: connected_closedin_eq) |
|
5648 show "\<Inter>?\<T> \<subseteq> S" |
|
5649 by (fastforce simp: C in_components_subset) |
|
5650 qed |
|
5651 with x show "\<Inter>?\<T> \<subseteq> C" by simp |
|
5652 qed auto |
|
5653 |
|
5654 |
|
5655 corollary Sura_Bura_clopen_subset: |
|
5656 fixes S :: "'a::euclidean_space set" |
|
5657 assumes S: "locally compact S" and C: "C \<in> components S" and "compact C" |
|
5658 and U: "open U" "C \<subseteq> U" |
|
5659 obtains K where "openin (subtopology euclidean S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U" |
|
5660 proof (rule ccontr) |
|
5661 assume "\<not> thesis" |
|
5662 with that have neg: "\<nexists>K. openin (subtopology euclidean S) K \<and> compact K \<and> C \<subseteq> K \<and> K \<subseteq> U" |
|
5663 by metis |
|
5664 obtain V K where "C \<subseteq> V" "V \<subseteq> U" "V \<subseteq> K" "K \<subseteq> S" "compact K" |
|
5665 and opeSV: "openin (subtopology euclidean S) V" |
|
5666 using S U \<open>compact C\<close> |
|
5667 apply (simp add: locally_compact_compact_subopen) |
|
5668 by (meson C in_components_subset) |
|
5669 let ?\<T> = "{T. C \<subseteq> T \<and> openin (subtopology euclidean K) T \<and> compact T \<and> T \<subseteq> K}" |
|
5670 have CK: "C \<in> components K" |
|
5671 by (meson C \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> components_intermediate_subset subset_trans) |
|
5672 with \<open>compact K\<close> |
|
5673 have "C = \<Inter>{T. C \<subseteq> T \<and> openin (subtopology euclidean K) T \<and> closedin (subtopology euclidean K) T}" |
|
5674 by (simp add: Sura_Bura_compact) |
|
5675 then have Ceq: "C = \<Inter>?\<T>" |
|
5676 by (simp add: closedin_compact_eq \<open>compact K\<close>) |
|
5677 obtain W where "open W" and W: "V = S \<inter> W" |
|
5678 using opeSV by (auto simp: openin_open) |
|
5679 have "-(U \<inter> W) \<inter> \<Inter>?\<T> \<noteq> {}" |
|
5680 proof (rule closed_imp_fip_compact) |
|
5681 show "- (U \<inter> W) \<inter> \<Inter>\<F> \<noteq> {}" |
|
5682 if "finite \<F>" and \<F>: "\<F> \<subseteq> ?\<T>" for \<F> |
|
5683 proof (cases "\<F> = {}") |
|
5684 case True |
|
5685 have False if "U = UNIV" "W = UNIV" |
|
5686 proof - |
|
5687 have "V = S" |
|
5688 by (simp add: W \<open>W = UNIV\<close>) |
|
5689 with neg show False |
|
5690 using \<open>C \<subseteq> V\<close> \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> \<open>V \<subseteq> U\<close> \<open>compact K\<close> by auto |
|
5691 qed |
|
5692 with True show ?thesis |
|
5693 by auto |
|
5694 next |
|
5695 case False |
|
5696 show ?thesis |
|
5697 proof |
|
5698 assume "- (U \<inter> W) \<inter> \<Inter>\<F> = {}" |
|
5699 then have FUW: "\<Inter>\<F> \<subseteq> U \<inter> W" |
|
5700 by blast |
|
5701 have "C \<subseteq> \<Inter>\<F>" |
|
5702 using \<F> by auto |
|
5703 moreover have "compact (\<Inter>\<F>)" |
|
5704 by (metis (no_types, lifting) compact_Inter False mem_Collect_eq subsetCE \<F>) |
|
5705 moreover have "\<Inter>\<F> \<subseteq> K" |
|
5706 using False that(2) by fastforce |
|
5707 moreover have opeKF: "openin (subtopology euclidean K) (\<Inter>\<F>)" |
|
5708 using False \<F> \<open>finite \<F>\<close> by blast |
|
5709 then have opeVF: "openin (subtopology euclidean V) (\<Inter>\<F>)" |
|
5710 using W \<open>K \<subseteq> S\<close> \<open>V \<subseteq> K\<close> opeKF \<open>\<Inter>\<F> \<subseteq> K\<close> FUW openin_subset_trans by fastforce |
|
5711 then have "openin (subtopology euclidean S) (\<Inter>\<F>)" |
|
5712 by (metis opeSV openin_trans) |
|
5713 moreover have "\<Inter>\<F> \<subseteq> U" |
|
5714 by (meson \<open>V \<subseteq> U\<close> opeVF dual_order.trans openin_imp_subset) |
|
5715 ultimately show False |
|
5716 using neg by blast |
|
5717 qed |
|
5718 qed |
|
5719 qed (use \<open>open W\<close> \<open>open U\<close> in auto) |
|
5720 with W Ceq \<open>C \<subseteq> V\<close> \<open>C \<subseteq> U\<close> show False |
|
5721 by auto |
|
5722 qed |
|
5723 |
|
5724 |
|
5725 corollary Sura_Bura_clopen_subset_alt: |
|
5726 fixes S :: "'a::euclidean_space set" |
|
5727 assumes S: "locally compact S" and C: "C \<in> components S" and "compact C" |
|
5728 and opeSU: "openin (subtopology euclidean S) U" and "C \<subseteq> U" |
|
5729 obtains K where "openin (subtopology euclidean S) K" "compact K" "C \<subseteq> K" "K \<subseteq> U" |
|
5730 proof - |
|
5731 obtain V where "open V" "U = S \<inter> V" |
|
5732 using opeSU by (auto simp: openin_open) |
|
5733 with \<open>C \<subseteq> U\<close> have "C \<subseteq> V" |
|
5734 by auto |
|
5735 then show ?thesis |
|
5736 using Sura_Bura_clopen_subset [OF S C \<open>compact C\<close> \<open>open V\<close>] |
|
5737 by (metis \<open>U = S \<inter> V\<close> inf.bounded_iff openin_imp_subset that) |
|
5738 qed |
|
5739 |
|
5740 corollary Sura_Bura: |
|
5741 fixes S :: "'a::euclidean_space set" |
|
5742 assumes "locally compact S" "C \<in> components S" "compact C" |
|
5743 shows "C = \<Inter> {K. C \<subseteq> K \<and> compact K \<and> openin (subtopology euclidean S) K}" |
|
5744 (is "C = ?rhs") |
|
5745 proof |
|
5746 show "?rhs \<subseteq> C" |
|
5747 proof (clarsimp, rule ccontr) |
|
5748 fix x |
|
5749 assume *: "\<forall>X. C \<subseteq> X \<and> compact X \<and> openin (subtopology euclidean S) X \<longrightarrow> x \<in> X" |
|
5750 and "x \<notin> C" |
|
5751 obtain U V where "open U" "open V" "{x} \<subseteq> U" "C \<subseteq> V" "U \<inter> V = {}" |
|
5752 using separation_normal [of "{x}" C] |
|
5753 by (metis Int_empty_left \<open>x \<notin> C\<close> \<open>compact C\<close> closed_empty closed_insert compact_imp_closed insert_disjoint(1)) |
|
5754 have "x \<notin> V" |
|
5755 using \<open>U \<inter> V = {}\<close> \<open>{x} \<subseteq> U\<close> by blast |
|
5756 then show False |
|
5757 by (meson "*" Sura_Bura_clopen_subset \<open>C \<subseteq> V\<close> \<open>open V\<close> assms(1) assms(2) assms(3) subsetCE) |
|
5758 qed |
|
5759 qed blast |
|
5760 |
|
5761 |
|
5762 subsection\<open>Special cases of local connectedness and path connectedness\<close> |
|
5763 |
|
5764 lemma locally_connected_1: |
|
5765 assumes |
|
5766 "\<And>v x. \<lbrakk>openin (subtopology euclidean S) v; x \<in> v\<rbrakk> |
|
5767 \<Longrightarrow> \<exists>u. openin (subtopology euclidean S) u \<and> |
|
5768 connected u \<and> x \<in> u \<and> u \<subseteq> v" |
|
5769 shows "locally connected S" |
|
5770 apply (clarsimp simp add: locally_def) |
|
5771 apply (drule assms; blast) |
|
5772 done |
|
5773 |
|
5774 lemma locally_connected_2: |
|
5775 assumes "locally connected S" |
|
5776 "openin (subtopology euclidean S) t" |
|
5777 "x \<in> t" |
|
5778 shows "openin (subtopology euclidean S) (connected_component_set t x)" |
|
5779 proof - |
|
5780 { fix y :: 'a |
|
5781 let ?SS = "subtopology euclidean S" |
|
5782 assume 1: "openin ?SS t" |
|
5783 "\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))" |
|
5784 and "connected_component t x y" |
|
5785 then have "y \<in> t" and y: "y \<in> connected_component_set t x" |
|
5786 using connected_component_subset by blast+ |
|
5787 obtain F where |
|
5788 "\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))" |
|
5789 by moura |
|
5790 then obtain G where |
|
5791 "\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)" |
|
5792 by moura |
|
5793 then have *: "openin ?SS (F y t) \<and> connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t" |
|
5794 using 1 \<open>y \<in> t\<close> by presburger |
|
5795 have "G y t \<subseteq> connected_component_set t y" |
|
5796 by (metis (no_types) * connected_component_eq_self connected_component_mono contra_subsetD) |
|
5797 then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> connected_component_set t x" |
|
5798 by (metis (no_types) * connected_component_eq dual_order.trans y) |
|
5799 } |
|
5800 then show ?thesis |
|
5801 using assms openin_subopen by (force simp: locally_def) |
|
5802 qed |
|
5803 |
|
5804 lemma locally_connected_3: |
|
5805 assumes "\<And>t x. \<lbrakk>openin (subtopology euclidean S) t; x \<in> t\<rbrakk> |
|
5806 \<Longrightarrow> openin (subtopology euclidean S) |
|
5807 (connected_component_set t x)" |
|
5808 "openin (subtopology euclidean S) v" "x \<in> v" |
|
5809 shows "\<exists>u. openin (subtopology euclidean S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v" |
|
5810 using assms connected_component_subset by fastforce |
|
5811 |
|
5812 lemma locally_connected: |
|
5813 "locally connected S \<longleftrightarrow> |
|
5814 (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v |
|
5815 \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and> connected u \<and> x \<in> u \<and> u \<subseteq> v))" |
|
5816 by (metis locally_connected_1 locally_connected_2 locally_connected_3) |
|
5817 |
|
5818 lemma locally_connected_open_connected_component: |
|
5819 "locally connected S \<longleftrightarrow> |
|
5820 (\<forall>t x. openin (subtopology euclidean S) t \<and> x \<in> t |
|
5821 \<longrightarrow> openin (subtopology euclidean S) (connected_component_set t x))" |
|
5822 by (metis locally_connected_1 locally_connected_2 locally_connected_3) |
|
5823 |
|
5824 lemma locally_path_connected_1: |
|
5825 assumes |
|
5826 "\<And>v x. \<lbrakk>openin (subtopology euclidean S) v; x \<in> v\<rbrakk> |
|
5827 \<Longrightarrow> \<exists>u. openin (subtopology euclidean S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v" |
|
5828 shows "locally path_connected S" |
|
5829 apply (clarsimp simp add: locally_def) |
|
5830 apply (drule assms; blast) |
|
5831 done |
|
5832 |
|
5833 lemma locally_path_connected_2: |
|
5834 assumes "locally path_connected S" |
|
5835 "openin (subtopology euclidean S) t" |
|
5836 "x \<in> t" |
|
5837 shows "openin (subtopology euclidean S) (path_component_set t x)" |
|
5838 proof - |
|
5839 { fix y :: 'a |
|
5840 let ?SS = "subtopology euclidean S" |
|
5841 assume 1: "openin ?SS t" |
|
5842 "\<forall>w x. openin ?SS w \<and> x \<in> w \<longrightarrow> (\<exists>u. openin ?SS u \<and> (\<exists>v. path_connected v \<and> x \<in> u \<and> u \<subseteq> v \<and> v \<subseteq> w))" |
|
5843 and "path_component t x y" |
|
5844 then have "y \<in> t" and y: "y \<in> path_component_set t x" |
|
5845 using path_component_mem(2) by blast+ |
|
5846 obtain F where |
|
5847 "\<forall>x y. (\<exists>w. openin ?SS w \<and> (\<exists>u. path_connected u \<and> x \<in> w \<and> w \<subseteq> u \<and> u \<subseteq> y)) = (openin ?SS (F x y) \<and> (\<exists>u. path_connected u \<and> x \<in> F x y \<and> F x y \<subseteq> u \<and> u \<subseteq> y))" |
|
5848 by moura |
|
5849 then obtain G where |
|
5850 "\<forall>a A. (\<exists>U. openin ?SS U \<and> (\<exists>V. path_connected V \<and> a \<in> U \<and> U \<subseteq> V \<and> V \<subseteq> A)) = (openin ?SS (F a A) \<and> path_connected (G a A) \<and> a \<in> F a A \<and> F a A \<subseteq> G a A \<and> G a A \<subseteq> A)" |
|
5851 by moura |
|
5852 then have *: "openin ?SS (F y t) \<and> path_connected (G y t) \<and> y \<in> F y t \<and> F y t \<subseteq> G y t \<and> G y t \<subseteq> t" |
|
5853 using 1 \<open>y \<in> t\<close> by presburger |
|
5854 have "G y t \<subseteq> path_component_set t y" |
|
5855 using * path_component_maximal set_rev_mp by blast |
|
5856 then have "\<exists>A. openin ?SS A \<and> y \<in> A \<and> A \<subseteq> path_component_set t x" |
|
5857 by (metis "*" \<open>G y t \<subseteq> path_component_set t y\<close> dual_order.trans path_component_eq y) |
|
5858 } |
|
5859 then show ?thesis |
|
5860 using assms openin_subopen by (force simp: locally_def) |
|
5861 qed |
|
5862 |
|
5863 lemma locally_path_connected_3: |
|
5864 assumes "\<And>t x. \<lbrakk>openin (subtopology euclidean S) t; x \<in> t\<rbrakk> |
|
5865 \<Longrightarrow> openin (subtopology euclidean S) (path_component_set t x)" |
|
5866 "openin (subtopology euclidean S) v" "x \<in> v" |
|
5867 shows "\<exists>u. openin (subtopology euclidean S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v" |
|
5868 proof - |
|
5869 have "path_component v x x" |
|
5870 by (meson assms(3) path_component_refl) |
|
5871 then show ?thesis |
|
5872 by (metis assms(1) assms(2) assms(3) mem_Collect_eq path_component_subset path_connected_path_component) |
|
5873 qed |
|
5874 |
|
5875 proposition locally_path_connected: |
|
5876 "locally path_connected S \<longleftrightarrow> |
|
5877 (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v |
|
5878 \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and> path_connected u \<and> x \<in> u \<and> u \<subseteq> v))" |
|
5879 by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) |
|
5880 |
|
5881 proposition locally_path_connected_open_path_component: |
|
5882 "locally path_connected S \<longleftrightarrow> |
|
5883 (\<forall>t x. openin (subtopology euclidean S) t \<and> x \<in> t |
|
5884 \<longrightarrow> openin (subtopology euclidean S) (path_component_set t x))" |
|
5885 by (metis locally_path_connected_1 locally_path_connected_2 locally_path_connected_3) |
|
5886 |
|
5887 lemma locally_connected_open_component: |
|
5888 "locally connected S \<longleftrightarrow> |
|
5889 (\<forall>t c. openin (subtopology euclidean S) t \<and> c \<in> components t |
|
5890 \<longrightarrow> openin (subtopology euclidean S) c)" |
|
5891 by (metis components_iff locally_connected_open_connected_component) |
|
5892 |
|
5893 proposition locally_connected_im_kleinen: |
|
5894 "locally connected S \<longleftrightarrow> |
|
5895 (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v |
|
5896 \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and> |
|
5897 x \<in> u \<and> u \<subseteq> v \<and> |
|
5898 (\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> v \<and> x \<in> c \<and> y \<in> c))))" |
|
5899 (is "?lhs = ?rhs") |
|
5900 proof |
|
5901 assume ?lhs |
|
5902 then show ?rhs |
|
5903 by (fastforce simp add: locally_connected) |
|
5904 next |
|
5905 assume ?rhs |
|
5906 have *: "\<exists>T. openin (subtopology euclidean S) T \<and> x \<in> T \<and> T \<subseteq> c" |
|
5907 if "openin (subtopology euclidean S) t" and c: "c \<in> components t" and "x \<in> c" for t c x |
|
5908 proof - |
|
5909 from that \<open>?rhs\<close> [rule_format, of t x] |
|
5910 obtain u where u: |
|
5911 "openin (subtopology euclidean S) u \<and> x \<in> u \<and> u \<subseteq> t \<and> |
|
5912 (\<forall>y. y \<in> u \<longrightarrow> (\<exists>c. connected c \<and> c \<subseteq> t \<and> x \<in> c \<and> y \<in> c))" |
|
5913 using in_components_subset by auto |
|
5914 obtain F :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where |
|
5915 "\<forall>x y. (\<exists>z. z \<in> x \<and> y = connected_component_set x z) = (F x y \<in> x \<and> y = connected_component_set x (F x y))" |
|
5916 by moura |
|
5917 then have F: "F t c \<in> t \<and> c = connected_component_set t (F t c)" |
|
5918 by (meson components_iff c) |
|
5919 obtain G :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a" where |
|
5920 G: "\<forall>x y. (\<exists>z. z \<in> y \<and> z \<notin> x) = (G x y \<in> y \<and> G x y \<notin> x)" |
|
5921 by moura |
|
5922 have "G c u \<notin> u \<or> G c u \<in> c" |
|
5923 using F by (metis (full_types) u connected_componentI connected_component_eq mem_Collect_eq that(3)) |
|
5924 then show ?thesis |
|
5925 using G u by auto |
|
5926 qed |
|
5927 show ?lhs |
|
5928 apply (clarsimp simp add: locally_connected_open_component) |
|
5929 apply (subst openin_subopen) |
|
5930 apply (blast intro: *) |
|
5931 done |
|
5932 qed |
|
5933 |
|
5934 proposition locally_path_connected_im_kleinen: |
|
5935 "locally path_connected S \<longleftrightarrow> |
|
5936 (\<forall>v x. openin (subtopology euclidean S) v \<and> x \<in> v |
|
5937 \<longrightarrow> (\<exists>u. openin (subtopology euclidean S) u \<and> |
|
5938 x \<in> u \<and> u \<subseteq> v \<and> |
|
5939 (\<forall>y. y \<in> u \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> v \<and> |
|
5940 pathstart p = x \<and> pathfinish p = y))))" |
|
5941 (is "?lhs = ?rhs") |
|
5942 proof |
|
5943 assume ?lhs |
|
5944 then show ?rhs |
|
5945 apply (simp add: locally_path_connected path_connected_def) |
|
5946 apply (erule all_forward ex_forward imp_forward conjE | simp)+ |
|
5947 by (meson dual_order.trans) |
|
5948 next |
|
5949 assume ?rhs |
|
5950 have *: "\<exists>T. openin (subtopology euclidean S) T \<and> |
|
5951 x \<in> T \<and> T \<subseteq> path_component_set u z" |
|
5952 if "openin (subtopology euclidean S) u" and "z \<in> u" and c: "path_component u z x" for u z x |
|
5953 proof - |
|
5954 have "x \<in> u" |
|
5955 by (meson c path_component_mem(2)) |
|
5956 with that \<open>?rhs\<close> [rule_format, of u x] |
|
5957 obtain U where U: |
|
5958 "openin (subtopology euclidean S) U \<and> x \<in> U \<and> U \<subseteq> u \<and> |
|
5959 (\<forall>y. y \<in> U \<longrightarrow> (\<exists>p. path p \<and> path_image p \<subseteq> u \<and> pathstart p = x \<and> pathfinish p = y))" |
|
5960 by blast |
|
5961 show ?thesis |
|
5962 apply (rule_tac x=U in exI) |
|
5963 apply (auto simp: U) |
|
5964 apply (metis U c path_component_trans path_component_def) |
|
5965 done |
|
5966 qed |
|
5967 show ?lhs |
|
5968 apply (clarsimp simp add: locally_path_connected_open_path_component) |
|
5969 apply (subst openin_subopen) |
|
5970 apply (blast intro: *) |
|
5971 done |
|
5972 qed |
|
5973 |
|
5974 lemma locally_path_connected_imp_locally_connected: |
|
5975 "locally path_connected S \<Longrightarrow> locally connected S" |
|
5976 using locally_mono path_connected_imp_connected by blast |
|
5977 |
|
5978 lemma locally_connected_components: |
|
5979 "\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally connected c" |
|
5980 by (meson locally_connected_open_component locally_open_subset openin_subtopology_self) |
|
5981 |
|
5982 lemma locally_path_connected_components: |
|
5983 "\<lbrakk>locally path_connected S; c \<in> components S\<rbrakk> \<Longrightarrow> locally path_connected c" |
|
5984 by (meson locally_connected_open_component locally_open_subset locally_path_connected_imp_locally_connected openin_subtopology_self) |
|
5985 |
|
5986 lemma locally_path_connected_connected_component: |
|
5987 "locally path_connected S \<Longrightarrow> locally path_connected (connected_component_set S x)" |
|
5988 by (metis components_iff connected_component_eq_empty locally_empty locally_path_connected_components) |
|
5989 |
|
5990 lemma open_imp_locally_path_connected: |
|
5991 fixes S :: "'a :: real_normed_vector set" |
|
5992 shows "open S \<Longrightarrow> locally path_connected S" |
|
5993 apply (rule locally_mono [of convex]) |
|
5994 apply (simp_all add: locally_def openin_open_eq convex_imp_path_connected) |
|
5995 apply (meson open_ball centre_in_ball convex_ball openE order_trans) |
|
5996 done |
|
5997 |
|
5998 lemma open_imp_locally_connected: |
|
5999 fixes S :: "'a :: real_normed_vector set" |
|
6000 shows "open S \<Longrightarrow> locally connected S" |
|
6001 by (simp add: locally_path_connected_imp_locally_connected open_imp_locally_path_connected) |
|
6002 |
|
6003 lemma locally_path_connected_UNIV: "locally path_connected (UNIV::'a :: real_normed_vector set)" |
|
6004 by (simp add: open_imp_locally_path_connected) |
|
6005 |
|
6006 lemma locally_connected_UNIV: "locally connected (UNIV::'a :: real_normed_vector set)" |
|
6007 by (simp add: open_imp_locally_connected) |
|
6008 |
|
6009 lemma openin_connected_component_locally_connected: |
|
6010 "locally connected S |
|
6011 \<Longrightarrow> openin (subtopology euclidean S) (connected_component_set S x)" |
|
6012 apply (simp add: locally_connected_open_connected_component) |
|
6013 by (metis connected_component_eq_empty connected_component_subset open_empty open_subset openin_subtopology_self) |
|
6014 |
|
6015 lemma openin_components_locally_connected: |
|
6016 "\<lbrakk>locally connected S; c \<in> components S\<rbrakk> \<Longrightarrow> openin (subtopology euclidean S) c" |
|
6017 using locally_connected_open_component openin_subtopology_self by blast |
|
6018 |
|
6019 lemma openin_path_component_locally_path_connected: |
|
6020 "locally path_connected S |
|
6021 \<Longrightarrow> openin (subtopology euclidean S) (path_component_set S x)" |
|
6022 by (metis (no_types) empty_iff locally_path_connected_2 openin_subopen openin_subtopology_self path_component_eq_empty) |
|
6023 |
|
6024 lemma closedin_path_component_locally_path_connected: |
|
6025 "locally path_connected S |
|
6026 \<Longrightarrow> closedin (subtopology euclidean S) (path_component_set S x)" |
|
6027 apply (simp add: closedin_def path_component_subset complement_path_component_Union) |
|
6028 apply (rule openin_Union) |
|
6029 using openin_path_component_locally_path_connected by auto |
|
6030 |
|
6031 lemma convex_imp_locally_path_connected: |
|
6032 fixes S :: "'a:: real_normed_vector set" |
|
6033 shows "convex S \<Longrightarrow> locally path_connected S" |
|
6034 apply (clarsimp simp add: locally_path_connected) |
|
6035 apply (subst (asm) openin_open) |
|
6036 apply clarify |
|
6037 apply (erule (1) openE) |
|
6038 apply (rule_tac x = "S \<inter> ball x e" in exI) |
|
6039 apply (force simp: convex_Int convex_imp_path_connected) |
|
6040 done |
|
6041 |
|
6042 lemma convex_imp_locally_connected: |
|
6043 fixes S :: "'a:: real_normed_vector set" |
|
6044 shows "convex S \<Longrightarrow> locally connected S" |
|
6045 by (simp add: locally_path_connected_imp_locally_connected convex_imp_locally_path_connected) |
|
6046 |
|
6047 |
|
6048 subsection\<open>Relations between components and path components\<close> |
|
6049 |
|
6050 lemma path_component_eq_connected_component: |
|
6051 assumes "locally path_connected S" |
|
6052 shows "(path_component S x = connected_component S x)" |
|
6053 proof (cases "x \<in> S") |
|
6054 case True |
|
6055 have "openin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" |
|
6056 apply (rule openin_subset_trans [of S]) |
|
6057 apply (intro conjI openin_path_component_locally_path_connected [OF assms]) |
|
6058 using path_component_subset_connected_component apply (auto simp: connected_component_subset) |
|
6059 done |
|
6060 moreover have "closedin (subtopology euclidean (connected_component_set S x)) (path_component_set S x)" |
|
6061 apply (rule closedin_subset_trans [of S]) |
|
6062 apply (intro conjI closedin_path_component_locally_path_connected [OF assms]) |
|
6063 using path_component_subset_connected_component apply (auto simp: connected_component_subset) |
|
6064 done |
|
6065 ultimately have *: "path_component_set S x = connected_component_set S x" |
|
6066 by (metis connected_connected_component connected_clopen True path_component_eq_empty) |
|
6067 then show ?thesis |
|
6068 by blast |
|
6069 next |
|
6070 case False then show ?thesis |
|
6071 by (metis Collect_empty_eq_bot connected_component_eq_empty path_component_eq_empty) |
|
6072 qed |
|
6073 |
|
6074 lemma path_component_eq_connected_component_set: |
|
6075 "locally path_connected S \<Longrightarrow> (path_component_set S x = connected_component_set S x)" |
|
6076 by (simp add: path_component_eq_connected_component) |
|
6077 |
|
6078 lemma locally_path_connected_path_component: |
|
6079 "locally path_connected S \<Longrightarrow> locally path_connected (path_component_set S x)" |
|
6080 using locally_path_connected_connected_component path_component_eq_connected_component by fastforce |
|
6081 |
|
6082 lemma open_path_connected_component: |
|
6083 fixes S :: "'a :: real_normed_vector set" |
|
6084 shows "open S \<Longrightarrow> path_component S x = connected_component S x" |
|
6085 by (simp add: path_component_eq_connected_component open_imp_locally_path_connected) |
|
6086 |
|
6087 lemma open_path_connected_component_set: |
|
6088 fixes S :: "'a :: real_normed_vector set" |
|
6089 shows "open S \<Longrightarrow> path_component_set S x = connected_component_set S x" |
|
6090 by (simp add: open_path_connected_component) |
|
6091 |
|
6092 proposition locally_connected_quotient_image: |
|
6093 assumes lcS: "locally connected S" |
|
6094 and oo: "\<And>T. T \<subseteq> f ` S |
|
6095 \<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow> |
|
6096 openin (subtopology euclidean (f ` S)) T" |
|
6097 shows "locally connected (f ` S)" |
|
6098 proof (clarsimp simp: locally_connected_open_component) |
|
6099 fix U C |
|
6100 assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "C \<in> components U" |
|
6101 then have "C \<subseteq> U" "U \<subseteq> f ` S" |
|
6102 by (meson in_components_subset openin_imp_subset)+ |
|
6103 then have "openin (subtopology euclidean (f ` S)) C \<longleftrightarrow> |
|
6104 openin (subtopology euclidean S) (S \<inter> f -` C)" |
|
6105 by (auto simp: oo) |
|
6106 moreover have "openin (subtopology euclidean S) (S \<inter> f -` C)" |
|
6107 proof (subst openin_subopen, clarify) |
|
6108 fix x |
|
6109 assume "x \<in> S" "f x \<in> C" |
|
6110 show "\<exists>T. openin (subtopology euclidean S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` C)" |
|
6111 proof (intro conjI exI) |
|
6112 show "openin (subtopology euclidean S) (connected_component_set (S \<inter> f -` U) x)" |
|
6113 proof (rule ccontr) |
|
6114 assume **: "\<not> openin (subtopology euclidean S) (connected_component_set (S \<inter> f -` U) x)" |
|
6115 then have "x \<notin> (S \<inter> f -` U)" |
|
6116 using \<open>U \<subseteq> f ` S\<close> opefSU lcS locally_connected_2 oo by blast |
|
6117 with ** show False |
|
6118 by (metis (no_types) connected_component_eq_empty empty_iff openin_subopen) |
|
6119 qed |
|
6120 next |
|
6121 show "x \<in> connected_component_set (S \<inter> f -` U) x" |
|
6122 using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by auto |
|
6123 next |
|
6124 have contf: "continuous_on S f" |
|
6125 by (simp add: continuous_on_open oo openin_imp_subset) |
|
6126 then have "continuous_on (connected_component_set (S \<inter> f -` U) x) f" |
|
6127 apply (rule continuous_on_subset) |
|
6128 using connected_component_subset apply blast |
|
6129 done |
|
6130 then have "connected (f ` connected_component_set (S \<inter> f -` U) x)" |
|
6131 by (rule connected_continuous_image [OF _ connected_connected_component]) |
|
6132 moreover have "f ` connected_component_set (S \<inter> f -` U) x \<subseteq> U" |
|
6133 using connected_component_in by blast |
|
6134 moreover have "C \<inter> f ` connected_component_set (S \<inter> f -` U) x \<noteq> {}" |
|
6135 using \<open>C \<subseteq> U\<close> \<open>f x \<in> C\<close> \<open>x \<in> S\<close> by fastforce |
|
6136 ultimately have fC: "f ` (connected_component_set (S \<inter> f -` U) x) \<subseteq> C" |
|
6137 by (rule components_maximal [OF \<open>C \<in> components U\<close>]) |
|
6138 have cUC: "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)" |
|
6139 using connected_component_subset fC by blast |
|
6140 have "connected_component_set (S \<inter> f -` U) x \<subseteq> connected_component_set (S \<inter> f -` C) x" |
|
6141 proof - |
|
6142 { assume "x \<in> connected_component_set (S \<inter> f -` U) x" |
|
6143 then have ?thesis |
|
6144 using cUC connected_component_idemp connected_component_mono by blast } |
|
6145 then show ?thesis |
|
6146 using connected_component_eq_empty by auto |
|
6147 qed |
|
6148 also have "\<dots> \<subseteq> (S \<inter> f -` C)" |
|
6149 by (rule connected_component_subset) |
|
6150 finally show "connected_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` C)" . |
|
6151 qed |
|
6152 qed |
|
6153 ultimately show "openin (subtopology euclidean (f ` S)) C" |
|
6154 by metis |
|
6155 qed |
|
6156 |
|
6157 text\<open>The proof resembles that above but is not identical!\<close> |
|
6158 proposition locally_path_connected_quotient_image: |
|
6159 assumes lcS: "locally path_connected S" |
|
6160 and oo: "\<And>T. T \<subseteq> f ` S |
|
6161 \<Longrightarrow> openin (subtopology euclidean S) (S \<inter> f -` T) \<longleftrightarrow> openin (subtopology euclidean (f ` S)) T" |
|
6162 shows "locally path_connected (f ` S)" |
|
6163 proof (clarsimp simp: locally_path_connected_open_path_component) |
|
6164 fix U y |
|
6165 assume opefSU: "openin (subtopology euclidean (f ` S)) U" and "y \<in> U" |
|
6166 then have "path_component_set U y \<subseteq> U" "U \<subseteq> f ` S" |
|
6167 by (meson path_component_subset openin_imp_subset)+ |
|
6168 then have "openin (subtopology euclidean (f ` S)) (path_component_set U y) \<longleftrightarrow> |
|
6169 openin (subtopology euclidean S) (S \<inter> f -` path_component_set U y)" |
|
6170 proof - |
|
6171 have "path_component_set U y \<subseteq> f ` S" |
|
6172 using \<open>U \<subseteq> f ` S\<close> \<open>path_component_set U y \<subseteq> U\<close> by blast |
|
6173 then show ?thesis |
|
6174 using oo by blast |
|
6175 qed |
|
6176 moreover have "openin (subtopology euclidean S) (S \<inter> f -` path_component_set U y)" |
|
6177 proof (subst openin_subopen, clarify) |
|
6178 fix x |
|
6179 assume "x \<in> S" and Uyfx: "path_component U y (f x)" |
|
6180 then have "f x \<in> U" |
|
6181 using path_component_mem by blast |
|
6182 show "\<exists>T. openin (subtopology euclidean S) T \<and> x \<in> T \<and> T \<subseteq> (S \<inter> f -` path_component_set U y)" |
|
6183 proof (intro conjI exI) |
|
6184 show "openin (subtopology euclidean S) (path_component_set (S \<inter> f -` U) x)" |
|
6185 proof (rule ccontr) |
|
6186 assume **: "\<not> openin (subtopology euclidean S) (path_component_set (S \<inter> f -` U) x)" |
|
6187 then have "x \<notin> (S \<inter> f -` U)" |
|
6188 by (metis (no_types, lifting) \<open>U \<subseteq> f ` S\<close> opefSU lcS oo locally_path_connected_open_path_component) |
|
6189 then show False |
|
6190 using ** \<open>path_component_set U y \<subseteq> U\<close> \<open>x \<in> S\<close> \<open>path_component U y (f x)\<close> by blast |
|
6191 qed |
|
6192 next |
|
6193 show "x \<in> path_component_set (S \<inter> f -` U) x" |
|
6194 by (simp add: \<open>f x \<in> U\<close> \<open>x \<in> S\<close> path_component_refl) |
|
6195 next |
|
6196 have contf: "continuous_on S f" |
|
6197 by (simp add: continuous_on_open oo openin_imp_subset) |
|
6198 then have "continuous_on (path_component_set (S \<inter> f -` U) x) f" |
|
6199 apply (rule continuous_on_subset) |
|
6200 using path_component_subset apply blast |
|
6201 done |
|
6202 then have "path_connected (f ` path_component_set (S \<inter> f -` U) x)" |
|
6203 by (simp add: path_connected_continuous_image) |
|
6204 moreover have "f ` path_component_set (S \<inter> f -` U) x \<subseteq> U" |
|
6205 using path_component_mem by fastforce |
|
6206 moreover have "f x \<in> f ` path_component_set (S \<inter> f -` U) x" |
|
6207 by (force simp: \<open>x \<in> S\<close> \<open>f x \<in> U\<close> path_component_refl_eq) |
|
6208 ultimately have "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U (f x)" |
|
6209 by (meson path_component_maximal) |
|
6210 also have "\<dots> \<subseteq> path_component_set U y" |
|
6211 by (simp add: Uyfx path_component_maximal path_component_subset path_component_sym) |
|
6212 finally have fC: "f ` (path_component_set (S \<inter> f -` U) x) \<subseteq> path_component_set U y" . |
|
6213 have cUC: "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)" |
|
6214 using path_component_subset fC by blast |
|
6215 have "path_component_set (S \<inter> f -` U) x \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) x" |
|
6216 proof - |
|
6217 have "\<And>a. path_component_set (path_component_set (S \<inter> f -` U) x) a \<subseteq> path_component_set (S \<inter> f -` path_component_set U y) a" |
|
6218 using cUC path_component_mono by blast |
|
6219 then show ?thesis |
|
6220 using path_component_path_component by blast |
|
6221 qed |
|
6222 also have "\<dots> \<subseteq> (S \<inter> f -` path_component_set U y)" |
|
6223 by (rule path_component_subset) |
|
6224 finally show "path_component_set (S \<inter> f -` U) x \<subseteq> (S \<inter> f -` path_component_set U y)" . |
|
6225 qed |
|
6226 qed |
|
6227 ultimately show "openin (subtopology euclidean (f ` S)) (path_component_set U y)" |
|
6228 by metis |
|
6229 qed |
|
6230 |
|
6231 subsection%unimportant\<open>Components, continuity, openin, closedin\<close> |
|
6232 |
|
6233 lemma continuous_on_components_gen: |
|
6234 fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
6235 assumes "\<And>c. c \<in> components S \<Longrightarrow> |
|
6236 openin (subtopology euclidean S) c \<and> continuous_on c f" |
|
6237 shows "continuous_on S f" |
|
6238 proof (clarsimp simp: continuous_openin_preimage_eq) |
|
6239 fix t :: "'b set" |
|
6240 assume "open t" |
|
6241 have *: "S \<inter> f -` t = (\<Union>c \<in> components S. c \<inter> f -` t)" |
|
6242 by auto |
|
6243 show "openin (subtopology euclidean S) (S \<inter> f -` t)" |
|
6244 unfolding * using \<open>open t\<close> assms continuous_openin_preimage_gen openin_trans openin_Union by blast |
|
6245 qed |
|
6246 |
|
6247 lemma continuous_on_components: |
|
6248 fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space" |
|
6249 assumes "locally connected S " |
|
6250 "\<And>c. c \<in> components S \<Longrightarrow> continuous_on c f" |
|
6251 shows "continuous_on S f" |
|
6252 apply (rule continuous_on_components_gen) |
|
6253 apply (auto simp: assms intro: openin_components_locally_connected) |
|
6254 done |
|
6255 |
|
6256 lemma continuous_on_components_eq: |
|
6257 "locally connected S |
|
6258 \<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))" |
|
6259 by (meson continuous_on_components continuous_on_subset in_components_subset) |
|
6260 |
|
6261 lemma continuous_on_components_open: |
|
6262 fixes S :: "'a::real_normed_vector set" |
|
6263 assumes "open S " |
|
6264 "\<And>c. c \<in> components S \<Longrightarrow> continuous_on c f" |
|
6265 shows "continuous_on S f" |
|
6266 using continuous_on_components open_imp_locally_connected assms by blast |
|
6267 |
|
6268 lemma continuous_on_components_open_eq: |
|
6269 fixes S :: "'a::real_normed_vector set" |
|
6270 shows "open S \<Longrightarrow> (continuous_on S f \<longleftrightarrow> (\<forall>c \<in> components S. continuous_on c f))" |
|
6271 using continuous_on_subset in_components_subset |
|
6272 by (blast intro: continuous_on_components_open) |
|
6273 |
|
6274 lemma closedin_union_complement_components: |
|
6275 assumes u: "locally connected u" |
|
6276 and S: "closedin (subtopology euclidean u) S" |
|
6277 and cuS: "c \<subseteq> components(u - S)" |
|
6278 shows "closedin (subtopology euclidean u) (S \<union> \<Union>c)" |
|
6279 proof - |
|
6280 have di: "(\<And>S t. S \<in> c \<and> t \<in> c' \<Longrightarrow> disjnt S t) \<Longrightarrow> disjnt (\<Union> c) (\<Union> c')" for c' |
|
6281 by (simp add: disjnt_def) blast |
|
6282 have "S \<subseteq> u" |
|
6283 using S closedin_imp_subset by blast |
|
6284 moreover have "u - S = \<Union>c \<union> \<Union>(components (u - S) - c)" |
|
6285 by (metis Diff_partition Union_components Union_Un_distrib assms(3)) |
|
6286 moreover have "disjnt (\<Union>c) (\<Union>(components (u - S) - c))" |
|
6287 apply (rule di) |
|
6288 by (metis DiffD1 DiffD2 assms(3) components_nonoverlap disjnt_def subsetCE) |
|
6289 ultimately have eq: "S \<union> \<Union>c = u - (\<Union>(components(u - S) - c))" |
|
6290 by (auto simp: disjnt_def) |
|
6291 have *: "openin (subtopology euclidean u) (\<Union>(components (u - S) - c))" |
|
6292 apply (rule openin_Union) |
|
6293 apply (rule openin_trans [of "u - S"]) |
|
6294 apply (simp add: u S locally_diff_closed openin_components_locally_connected) |
|
6295 apply (simp add: openin_diff S) |
|
6296 done |
|
6297 have "openin (subtopology euclidean u) (u - (u - \<Union>(components (u - S) - c)))" |
|
6298 apply (rule openin_diff, simp) |
|
6299 apply (metis closedin_diff closedin_topspace topspace_euclidean_subtopology *) |
|
6300 done |
|
6301 then show ?thesis |
|
6302 by (force simp: eq closedin_def) |
|
6303 qed |
|
6304 |
|
6305 lemma closed_union_complement_components: |
|
6306 fixes S :: "'a::real_normed_vector set" |
|
6307 assumes S: "closed S" and c: "c \<subseteq> components(- S)" |
|
6308 shows "closed(S \<union> \<Union> c)" |
|
6309 proof - |
|
6310 have "closedin (subtopology euclidean UNIV) (S \<union> \<Union>c)" |
|
6311 apply (rule closedin_union_complement_components [OF locally_connected_UNIV]) |
|
6312 using S c apply (simp_all add: Compl_eq_Diff_UNIV) |
|
6313 done |
|
6314 then show ?thesis by simp |
|
6315 qed |
|
6316 |
|
6317 lemma closedin_Un_complement_component: |
|
6318 fixes S :: "'a::real_normed_vector set" |
|
6319 assumes u: "locally connected u" |
|
6320 and S: "closedin (subtopology euclidean u) S" |
|
6321 and c: " c \<in> components(u - S)" |
|
6322 shows "closedin (subtopology euclidean u) (S \<union> c)" |
|
6323 proof - |
|
6324 have "closedin (subtopology euclidean u) (S \<union> \<Union>{c})" |
|
6325 using c by (blast intro: closedin_union_complement_components [OF u S]) |
|
6326 then show ?thesis |
|
6327 by simp |
|
6328 qed |
|
6329 |
|
6330 lemma closed_Un_complement_component: |
|
6331 fixes S :: "'a::real_normed_vector set" |
|
6332 assumes S: "closed S" and c: " c \<in> components(-S)" |
|
6333 shows "closed (S \<union> c)" |
|
6334 by (metis Compl_eq_Diff_UNIV S c closed_closedin closedin_Un_complement_component |
|
6335 locally_connected_UNIV subtopology_UNIV) |
|
6336 |
|
6337 |
|
6338 subsection\<open>Existence of isometry between subspaces of same dimension\<close> |
|
6339 |
|
6340 lemma isometry_subset_subspace: |
|
6341 fixes S :: "'a::euclidean_space set" |
|
6342 and T :: "'b::euclidean_space set" |
|
6343 assumes S: "subspace S" |
|
6344 and T: "subspace T" |
|
6345 and d: "dim S \<le> dim T" |
|
6346 obtains f where "linear f" "f ` S \<subseteq> T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" |
|
6347 proof - |
|
6348 obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B" |
|
6349 and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" |
|
6350 and "independent B" "finite B" "card B = dim S" "span B = S" |
|
6351 by (metis orthonormal_basis_subspace [OF S] independent_finite) |
|
6352 obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C" |
|
6353 and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1" |
|
6354 and "independent C" "finite C" "card C = dim T" "span C = T" |
|
6355 by (metis orthonormal_basis_subspace [OF T] independent_finite) |
|
6356 obtain fb where "fb ` B \<subseteq> C" "inj_on fb B" |
|
6357 by (metis \<open>card B = dim S\<close> \<open>card C = dim T\<close> \<open>finite B\<close> \<open>finite C\<close> card_le_inj d) |
|
6358 then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B" |
|
6359 using Corth |
|
6360 apply (auto simp: pairwise_def orthogonal_clauses) |
|
6361 by (meson subsetD image_eqI inj_on_def) |
|
6362 obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x" |
|
6363 using linear_independent_extend \<open>independent B\<close> by fastforce |
|
6364 have "span (f ` B) \<subseteq> span C" |
|
6365 by (metis \<open>fb ` B \<subseteq> C\<close> ffb image_cong span_mono) |
|
6366 then have "f ` S \<subseteq> T" |
|
6367 unfolding \<open>span B = S\<close> \<open>span C = T\<close> span_linear_image[OF \<open>linear f\<close>] . |
|
6368 have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x" |
|
6369 using B1 C1 \<open>fb ` B \<subseteq> C\<close> by auto |
|
6370 have "norm (f x) = norm x" if "x \<in> S" for x |
|
6371 proof - |
|
6372 interpret linear f by fact |
|
6373 obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)" |
|
6374 using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce |
|
6375 have "norm (f x)^2 = norm (\<Sum>v\<in>B. a v *\<^sub>R fb v)^2" by (simp add: sum scale ffb x) |
|
6376 also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)" |
|
6377 apply (rule norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
|
6378 apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) |
|
6379 done |
|
6380 also have "\<dots> = norm x ^2" |
|
6381 by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
|
6382 finally show ?thesis |
|
6383 by (simp add: norm_eq_sqrt_inner) |
|
6384 qed |
|
6385 then show ?thesis |
|
6386 by (rule that [OF \<open>linear f\<close> \<open>f ` S \<subseteq> T\<close>]) |
|
6387 qed |
|
6388 |
|
6389 proposition isometries_subspaces: |
|
6390 fixes S :: "'a::euclidean_space set" |
|
6391 and T :: "'b::euclidean_space set" |
|
6392 assumes S: "subspace S" |
|
6393 and T: "subspace T" |
|
6394 and d: "dim S = dim T" |
|
6395 obtains f g where "linear f" "linear g" "f ` S = T" "g ` T = S" |
|
6396 "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" |
|
6397 "\<And>x. x \<in> T \<Longrightarrow> norm(g x) = norm x" |
|
6398 "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" |
|
6399 "\<And>x. x \<in> T \<Longrightarrow> f(g x) = x" |
|
6400 proof - |
|
6401 obtain B where "B \<subseteq> S" and Borth: "pairwise orthogonal B" |
|
6402 and B1: "\<And>x. x \<in> B \<Longrightarrow> norm x = 1" |
|
6403 and "independent B" "finite B" "card B = dim S" "span B = S" |
|
6404 by (metis orthonormal_basis_subspace [OF S] independent_finite) |
|
6405 obtain C where "C \<subseteq> T" and Corth: "pairwise orthogonal C" |
|
6406 and C1:"\<And>x. x \<in> C \<Longrightarrow> norm x = 1" |
|
6407 and "independent C" "finite C" "card C = dim T" "span C = T" |
|
6408 by (metis orthonormal_basis_subspace [OF T] independent_finite) |
|
6409 obtain fb where "bij_betw fb B C" |
|
6410 by (metis \<open>finite B\<close> \<open>finite C\<close> bij_betw_iff_card \<open>card B = dim S\<close> \<open>card C = dim T\<close> d) |
|
6411 then have pairwise_orth_fb: "pairwise (\<lambda>v j. orthogonal (fb v) (fb j)) B" |
|
6412 using Corth |
|
6413 apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) |
|
6414 by (meson subsetD image_eqI inj_on_def) |
|
6415 obtain f where "linear f" and ffb: "\<And>x. x \<in> B \<Longrightarrow> f x = fb x" |
|
6416 using linear_independent_extend \<open>independent B\<close> by fastforce |
|
6417 interpret f: linear f by fact |
|
6418 define gb where "gb \<equiv> inv_into B fb" |
|
6419 then have pairwise_orth_gb: "pairwise (\<lambda>v j. orthogonal (gb v) (gb j)) C" |
|
6420 using Borth |
|
6421 apply (auto simp: pairwise_def orthogonal_clauses bij_betw_def) |
|
6422 by (metis \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on bij_betw_inv_into_right inv_into_into) |
|
6423 obtain g where "linear g" and ggb: "\<And>x. x \<in> C \<Longrightarrow> g x = gb x" |
|
6424 using linear_independent_extend \<open>independent C\<close> by fastforce |
|
6425 interpret g: linear g by fact |
|
6426 have "span (f ` B) \<subseteq> span C" |
|
6427 by (metis \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on eq_iff ffb image_cong) |
|
6428 then have "f ` S \<subseteq> T" |
|
6429 unfolding \<open>span B = S\<close> \<open>span C = T\<close> |
|
6430 span_linear_image[OF \<open>linear f\<close>] . |
|
6431 have [simp]: "\<And>x. x \<in> B \<Longrightarrow> norm (fb x) = norm x" |
|
6432 using B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on by fastforce |
|
6433 have f [simp]: "norm (f x) = norm x" "g (f x) = x" if "x \<in> S" for x |
|
6434 proof - |
|
6435 obtain a where x: "x = (\<Sum>v \<in> B. a v *\<^sub>R v)" |
|
6436 using \<open>finite B\<close> \<open>span B = S\<close> \<open>x \<in> S\<close> span_finite by fastforce |
|
6437 have "f x = (\<Sum>v \<in> B. f (a v *\<^sub>R v))" |
|
6438 using linear_sum [OF \<open>linear f\<close>] x by auto |
|
6439 also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R f v)" |
|
6440 by (simp add: f.sum f.scale) |
|
6441 also have "\<dots> = (\<Sum>v \<in> B. a v *\<^sub>R fb v)" |
|
6442 by (simp add: ffb cong: sum.cong) |
|
6443 finally have *: "f x = (\<Sum>v\<in>B. a v *\<^sub>R fb v)" . |
|
6444 then have "(norm (f x))\<^sup>2 = (norm (\<Sum>v\<in>B. a v *\<^sub>R fb v))\<^sup>2" by simp |
|
6445 also have "\<dots> = (\<Sum>v\<in>B. norm ((a v *\<^sub>R fb v))^2)" |
|
6446 apply (rule norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
|
6447 apply (rule pairwise_ortho_scaleR [OF pairwise_orth_fb]) |
|
6448 done |
|
6449 also have "\<dots> = (norm x)\<^sup>2" |
|
6450 by (simp add: x pairwise_ortho_scaleR Borth norm_sum_Pythagorean [OF \<open>finite B\<close>]) |
|
6451 finally show "norm (f x) = norm x" |
|
6452 by (simp add: norm_eq_sqrt_inner) |
|
6453 have "g (f x) = g (\<Sum>v\<in>B. a v *\<^sub>R fb v)" by (simp add: *) |
|
6454 also have "\<dots> = (\<Sum>v\<in>B. g (a v *\<^sub>R fb v))" |
|
6455 by (simp add: g.sum g.scale) |
|
6456 also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R g (fb v))" |
|
6457 by (simp add: g.scale) |
|
6458 also have "\<dots> = (\<Sum>v\<in>B. a v *\<^sub>R v)" |
|
6459 apply (rule sum.cong [OF refl]) |
|
6460 using \<open>bij_betw fb B C\<close> gb_def bij_betwE bij_betw_inv_into_left gb_def ggb by fastforce |
|
6461 also have "\<dots> = x" |
|
6462 using x by blast |
|
6463 finally show "g (f x) = x" . |
|
6464 qed |
|
6465 have [simp]: "\<And>x. x \<in> C \<Longrightarrow> norm (gb x) = norm x" |
|
6466 by (metis B1 C1 \<open>bij_betw fb B C\<close> bij_betw_imp_surj_on gb_def inv_into_into) |
|
6467 have g [simp]: "f (g x) = x" if "x \<in> T" for x |
|
6468 proof - |
|
6469 obtain a where x: "x = (\<Sum>v \<in> C. a v *\<^sub>R v)" |
|
6470 using \<open>finite C\<close> \<open>span C = T\<close> \<open>x \<in> T\<close> span_finite by fastforce |
|
6471 have "g x = (\<Sum>v \<in> C. g (a v *\<^sub>R v))" |
|
6472 by (simp add: x g.sum) |
|
6473 also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R g v)" |
|
6474 by (simp add: g.scale) |
|
6475 also have "\<dots> = (\<Sum>v \<in> C. a v *\<^sub>R gb v)" |
|
6476 by (simp add: ggb cong: sum.cong) |
|
6477 finally have "f (g x) = f (\<Sum>v\<in>C. a v *\<^sub>R gb v)" by simp |
|
6478 also have "\<dots> = (\<Sum>v\<in>C. f (a v *\<^sub>R gb v))" |
|
6479 by (simp add: f.scale f.sum) |
|
6480 also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R f (gb v))" |
|
6481 by (simp add: f.scale f.sum) |
|
6482 also have "\<dots> = (\<Sum>v\<in>C. a v *\<^sub>R v)" |
|
6483 using \<open>bij_betw fb B C\<close> |
|
6484 by (simp add: bij_betw_def gb_def bij_betw_inv_into_right ffb inv_into_into) |
|
6485 also have "\<dots> = x" |
|
6486 using x by blast |
|
6487 finally show "f (g x) = x" . |
|
6488 qed |
|
6489 have gim: "g ` T = S" |
|
6490 by (metis (full_types) S T \<open>f ` S \<subseteq> T\<close> d dim_eq_span dim_image_le f(2) g.linear_axioms |
|
6491 image_iff linear_subspace_image span_eq_iff subset_iff) |
|
6492 have fim: "f ` S = T" |
|
6493 using \<open>g ` T = S\<close> image_iff by fastforce |
|
6494 have [simp]: "norm (g x) = norm x" if "x \<in> T" for x |
|
6495 using fim that by auto |
|
6496 show ?thesis |
|
6497 apply (rule that [OF \<open>linear f\<close> \<open>linear g\<close>]) |
|
6498 apply (simp_all add: fim gim) |
|
6499 done |
|
6500 qed |
|
6501 |
|
6502 corollary isometry_subspaces: |
|
6503 fixes S :: "'a::euclidean_space set" |
|
6504 and T :: "'b::euclidean_space set" |
|
6505 assumes S: "subspace S" |
|
6506 and T: "subspace T" |
|
6507 and d: "dim S = dim T" |
|
6508 obtains f where "linear f" "f ` S = T" "\<And>x. x \<in> S \<Longrightarrow> norm(f x) = norm x" |
|
6509 using isometries_subspaces [OF assms] |
|
6510 by metis |
|
6511 |
|
6512 corollary isomorphisms_UNIV_UNIV: |
|
6513 assumes "DIM('M) = DIM('N)" |
|
6514 obtains f::"'M::euclidean_space \<Rightarrow>'N::euclidean_space" and g |
|
6515 where "linear f" "linear g" |
|
6516 "\<And>x. norm(f x) = norm x" "\<And>y. norm(g y) = norm y" |
|
6517 "\<And>x. g (f x) = x" "\<And>y. f(g y) = y" |
|
6518 using assms by (auto intro: isometries_subspaces [of "UNIV::'M set" "UNIV::'N set"]) |
|
6519 |
|
6520 lemma homeomorphic_subspaces: |
|
6521 fixes S :: "'a::euclidean_space set" |
|
6522 and T :: "'b::euclidean_space set" |
|
6523 assumes S: "subspace S" |
|
6524 and T: "subspace T" |
|
6525 and d: "dim S = dim T" |
|
6526 shows "S homeomorphic T" |
|
6527 proof - |
|
6528 obtain f g where "linear f" "linear g" "f ` S = T" "g ` T = S" |
|
6529 "\<And>x. x \<in> S \<Longrightarrow> g(f x) = x" "\<And>x. x \<in> T \<Longrightarrow> f(g x) = x" |
|
6530 by (blast intro: isometries_subspaces [OF assms]) |
|
6531 then show ?thesis |
|
6532 apply (simp add: homeomorphic_def homeomorphism_def) |
|
6533 apply (rule_tac x=f in exI) |
|
6534 apply (rule_tac x=g in exI) |
|
6535 apply (auto simp: linear_continuous_on linear_conv_bounded_linear) |
|
6536 done |
|
6537 qed |
|
6538 |
|
6539 lemma homeomorphic_affine_sets: |
|
6540 assumes "affine S" "affine T" "aff_dim S = aff_dim T" |
|
6541 shows "S homeomorphic T" |
|
6542 proof (cases "S = {} \<or> T = {}") |
|
6543 case True with assms aff_dim_empty homeomorphic_empty show ?thesis |
|
6544 by metis |
|
6545 next |
|
6546 case False |
|
6547 then obtain a b where ab: "a \<in> S" "b \<in> T" by auto |
|
6548 then have ss: "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" |
|
6549 using affine_diffs_subspace assms by blast+ |
|
6550 have dd: "dim ((+) (- a) ` S) = dim ((+) (- b) ` T)" |
|
6551 using assms ab by (simp add: aff_dim_eq_dim [OF hull_inc] image_def) |
|
6552 have "S homeomorphic ((+) (- a) ` S)" |
|
6553 by (simp add: homeomorphic_translation) |
|
6554 also have "\<dots> homeomorphic ((+) (- b) ` T)" |
|
6555 by (rule homeomorphic_subspaces [OF ss dd]) |
|
6556 also have "\<dots> homeomorphic T" |
|
6557 using homeomorphic_sym homeomorphic_translation by auto |
|
6558 finally show ?thesis . |
|
6559 qed |
|
6560 |
|
6561 |
|
6562 subsection\<open>Retracts, in a general sense, preserve (co)homotopic triviality)\<close> |
|
6563 |
|
6564 locale%important Retracts = |
|
6565 fixes s h t k |
|
6566 assumes conth: "continuous_on s h" |
|
6567 and imh: "h ` s = t" |
|
6568 and contk: "continuous_on t k" |
|
6569 and imk: "k ` t \<subseteq> s" |
|
6570 and idhk: "\<And>y. y \<in> t \<Longrightarrow> h(k y) = y" |
|
6571 |
|
6572 begin |
|
6573 |
|
6574 lemma homotopically_trivial_retraction_gen: |
|
6575 assumes P: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)" |
|
6576 and Q: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)" |
|
6577 and Qeq: "\<And>h k. (\<And>x. x \<in> u \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
|
6578 and hom: "\<And>f g. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f; |
|
6579 continuous_on u g; g ` u \<subseteq> s; P g\<rbrakk> |
|
6580 \<Longrightarrow> homotopic_with P u s f g" |
|
6581 and contf: "continuous_on u f" and imf: "f ` u \<subseteq> t" and Qf: "Q f" |
|
6582 and contg: "continuous_on u g" and img: "g ` u \<subseteq> t" and Qg: "Q g" |
|
6583 shows "homotopic_with Q u t f g" |
|
6584 proof - |
|
6585 have feq: "\<And>x. x \<in> u \<Longrightarrow> (h \<circ> (k \<circ> f)) x = f x" using idhk imf by auto |
|
6586 have geq: "\<And>x. x \<in> u \<Longrightarrow> (h \<circ> (k \<circ> g)) x = g x" using idhk img by auto |
|
6587 have "continuous_on u (k \<circ> f)" |
|
6588 using contf continuous_on_compose continuous_on_subset contk imf by blast |
|
6589 moreover have "(k \<circ> f) ` u \<subseteq> s" |
|
6590 using imf imk by fastforce |
|
6591 moreover have "P (k \<circ> f)" |
|
6592 by (simp add: P Qf contf imf) |
|
6593 moreover have "continuous_on u (k \<circ> g)" |
|
6594 using contg continuous_on_compose continuous_on_subset contk img by blast |
|
6595 moreover have "(k \<circ> g) ` u \<subseteq> s" |
|
6596 using img imk by fastforce |
|
6597 moreover have "P (k \<circ> g)" |
|
6598 by (simp add: P Qg contg img) |
|
6599 ultimately have "homotopic_with P u s (k \<circ> f) (k \<circ> g)" |
|
6600 by (rule hom) |
|
6601 then have "homotopic_with Q u t (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))" |
|
6602 apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) |
|
6603 using Q by (auto simp: conth imh) |
|
6604 then show ?thesis |
|
6605 apply (rule homotopic_with_eq) |
|
6606 apply (metis feq) |
|
6607 apply (metis geq) |
|
6608 apply (metis Qeq) |
|
6609 done |
|
6610 qed |
|
6611 |
|
6612 lemma homotopically_trivial_retraction_null_gen: |
|
6613 assumes P: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> t; Q f\<rbrakk> \<Longrightarrow> P(k \<circ> f)" |
|
6614 and Q: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f\<rbrakk> \<Longrightarrow> Q(h \<circ> f)" |
|
6615 and Qeq: "\<And>h k. (\<And>x. x \<in> u \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
|
6616 and hom: "\<And>f. \<lbrakk>continuous_on u f; f ` u \<subseteq> s; P f\<rbrakk> |
|
6617 \<Longrightarrow> \<exists>c. homotopic_with P u s f (\<lambda>x. c)" |
|
6618 and contf: "continuous_on u f" and imf:"f ` u \<subseteq> t" and Qf: "Q f" |
|
6619 obtains c where "homotopic_with Q u t f (\<lambda>x. c)" |
|
6620 proof - |
|
6621 have feq: "\<And>x. x \<in> u \<Longrightarrow> (h \<circ> (k \<circ> f)) x = f x" using idhk imf by auto |
|
6622 have "continuous_on u (k \<circ> f)" |
|
6623 using contf continuous_on_compose continuous_on_subset contk imf by blast |
|
6624 moreover have "(k \<circ> f) ` u \<subseteq> s" |
|
6625 using imf imk by fastforce |
|
6626 moreover have "P (k \<circ> f)" |
|
6627 by (simp add: P Qf contf imf) |
|
6628 ultimately obtain c where "homotopic_with P u s (k \<circ> f) (\<lambda>x. c)" |
|
6629 by (metis hom) |
|
6630 then have "homotopic_with Q u t (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))" |
|
6631 apply (rule homotopic_with_compose_continuous_left [OF homotopic_with_mono]) |
|
6632 using Q by (auto simp: conth imh) |
|
6633 then show ?thesis |
|
6634 apply (rule_tac c = "h c" in that) |
|
6635 apply (erule homotopic_with_eq) |
|
6636 apply (metis feq, simp) |
|
6637 apply (metis Qeq) |
|
6638 done |
|
6639 qed |
|
6640 |
|
6641 lemma cohomotopically_trivial_retraction_gen: |
|
6642 assumes P: "\<And>f. \<lbrakk>continuous_on t f; f ` t \<subseteq> u; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)" |
|
6643 and Q: "\<And>f. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)" |
|
6644 and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
|
6645 and hom: "\<And>f g. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f; |
|
6646 continuous_on s g; g ` s \<subseteq> u; P g\<rbrakk> |
|
6647 \<Longrightarrow> homotopic_with P s u f g" |
|
6648 and contf: "continuous_on t f" and imf: "f ` t \<subseteq> u" and Qf: "Q f" |
|
6649 and contg: "continuous_on t g" and img: "g ` t \<subseteq> u" and Qg: "Q g" |
|
6650 shows "homotopic_with Q t u f g" |
|
6651 proof - |
|
6652 have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto |
|
6653 have geq: "\<And>x. x \<in> t \<Longrightarrow> (g \<circ> h \<circ> k) x = g x" using idhk img by auto |
|
6654 have "continuous_on s (f \<circ> h)" |
|
6655 using contf conth continuous_on_compose imh by blast |
|
6656 moreover have "(f \<circ> h) ` s \<subseteq> u" |
|
6657 using imf imh by fastforce |
|
6658 moreover have "P (f \<circ> h)" |
|
6659 by (simp add: P Qf contf imf) |
|
6660 moreover have "continuous_on s (g \<circ> h)" |
|
6661 using contg continuous_on_compose continuous_on_subset conth imh by blast |
|
6662 moreover have "(g \<circ> h) ` s \<subseteq> u" |
|
6663 using img imh by fastforce |
|
6664 moreover have "P (g \<circ> h)" |
|
6665 by (simp add: P Qg contg img) |
|
6666 ultimately have "homotopic_with P s u (f \<circ> h) (g \<circ> h)" |
|
6667 by (rule hom) |
|
6668 then have "homotopic_with Q t u (f \<circ> h \<circ> k) (g \<circ> h \<circ> k)" |
|
6669 apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) |
|
6670 using Q by (auto simp: contk imk) |
|
6671 then show ?thesis |
|
6672 apply (rule homotopic_with_eq) |
|
6673 apply (metis feq) |
|
6674 apply (metis geq) |
|
6675 apply (metis Qeq) |
|
6676 done |
|
6677 qed |
|
6678 |
|
6679 lemma cohomotopically_trivial_retraction_null_gen: |
|
6680 assumes P: "\<And>f. \<lbrakk>continuous_on t f; f ` t \<subseteq> u; Q f\<rbrakk> \<Longrightarrow> P(f \<circ> h)" |
|
6681 and Q: "\<And>f. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f\<rbrakk> \<Longrightarrow> Q(f \<circ> k)" |
|
6682 and Qeq: "\<And>h k. (\<And>x. x \<in> t \<Longrightarrow> h x = k x) \<Longrightarrow> Q h = Q k" |
|
6683 and hom: "\<And>f g. \<lbrakk>continuous_on s f; f ` s \<subseteq> u; P f\<rbrakk> |
|
6684 \<Longrightarrow> \<exists>c. homotopic_with P s u f (\<lambda>x. c)" |
|
6685 and contf: "continuous_on t f" and imf: "f ` t \<subseteq> u" and Qf: "Q f" |
|
6686 obtains c where "homotopic_with Q t u f (\<lambda>x. c)" |
|
6687 proof - |
|
6688 have feq: "\<And>x. x \<in> t \<Longrightarrow> (f \<circ> h \<circ> k) x = f x" using idhk imf by auto |
|
6689 have "continuous_on s (f \<circ> h)" |
|
6690 using contf conth continuous_on_compose imh by blast |
|
6691 moreover have "(f \<circ> h) ` s \<subseteq> u" |
|
6692 using imf imh by fastforce |
|
6693 moreover have "P (f \<circ> h)" |
|
6694 by (simp add: P Qf contf imf) |
|
6695 ultimately obtain c where "homotopic_with P s u (f \<circ> h) (\<lambda>x. c)" |
|
6696 by (metis hom) |
|
6697 then have "homotopic_with Q t u (f \<circ> h \<circ> k) ((\<lambda>x. c) \<circ> k)" |
|
6698 apply (rule homotopic_with_compose_continuous_right [OF homotopic_with_mono]) |
|
6699 using Q by (auto simp: contk imk) |
|
6700 then show ?thesis |
|
6701 apply (rule_tac c = c in that) |
|
6702 apply (erule homotopic_with_eq) |
|
6703 apply (metis feq, simp) |
|
6704 apply (metis Qeq) |
|
6705 done |
|
6706 qed |
|
6707 |
|
6708 end |
3387 end |
6709 |
|
6710 lemma simply_connected_retraction_gen: |
|
6711 shows "\<lbrakk>simply_connected S; continuous_on S h; h ` S = T; |
|
6712 continuous_on T k; k ` T \<subseteq> S; \<And>y. y \<in> T \<Longrightarrow> h(k y) = y\<rbrakk> |
|
6713 \<Longrightarrow> simply_connected T" |
|
6714 apply (simp add: simply_connected_def path_def path_image_def homotopic_loops_def, clarify) |
|
6715 apply (rule Retracts.homotopically_trivial_retraction_gen |
|
6716 [of S h _ k _ "\<lambda>p. pathfinish p = pathstart p" "\<lambda>p. pathfinish p = pathstart p"]) |
|
6717 apply (simp_all add: Retracts_def pathfinish_def pathstart_def) |
|
6718 done |
|
6719 |
|
6720 lemma homeomorphic_simply_connected: |
|
6721 "\<lbrakk>S homeomorphic T; simply_connected S\<rbrakk> \<Longrightarrow> simply_connected T" |
|
6722 by (auto simp: homeomorphic_def homeomorphism_def intro: simply_connected_retraction_gen) |
|
6723 |
|
6724 lemma homeomorphic_simply_connected_eq: |
|
6725 "S homeomorphic T \<Longrightarrow> (simply_connected S \<longleftrightarrow> simply_connected T)" |
|
6726 by (metis homeomorphic_simply_connected homeomorphic_sym) |
|
6727 |
|
6728 |
|
6729 subsection\<open>Homotopy equivalence\<close> |
|
6730 |
|
6731 definition%important homotopy_eqv :: "'a::topological_space set \<Rightarrow> 'b::topological_space set \<Rightarrow> bool" |
|
6732 (infix "homotopy'_eqv" 50) |
|
6733 where "S homotopy_eqv T \<equiv> |
|
6734 \<exists>f g. continuous_on S f \<and> f ` S \<subseteq> T \<and> |
|
6735 continuous_on T g \<and> g ` T \<subseteq> S \<and> |
|
6736 homotopic_with (\<lambda>x. True) S S (g \<circ> f) id \<and> |
|
6737 homotopic_with (\<lambda>x. True) T T (f \<circ> g) id" |
|
6738 |
|
6739 lemma homeomorphic_imp_homotopy_eqv: "S homeomorphic T \<Longrightarrow> S homotopy_eqv T" |
|
6740 unfolding homeomorphic_def homotopy_eqv_def homeomorphism_def |
|
6741 by (fastforce intro!: homotopic_with_equal continuous_on_compose) |
|
6742 |
|
6743 lemma homotopy_eqv_refl: "S homotopy_eqv S" |
|
6744 by (rule homeomorphic_imp_homotopy_eqv homeomorphic_refl)+ |
|
6745 |
|
6746 lemma homotopy_eqv_sym: "S homotopy_eqv T \<longleftrightarrow> T homotopy_eqv S" |
|
6747 by (auto simp: homotopy_eqv_def) |
|
6748 |
|
6749 lemma homotopy_eqv_trans [trans]: |
|
6750 fixes S :: "'a::real_normed_vector set" and U :: "'c::real_normed_vector set" |
|
6751 assumes ST: "S homotopy_eqv T" and TU: "T homotopy_eqv U" |
|
6752 shows "S homotopy_eqv U" |
|
6753 proof - |
|
6754 obtain f1 g1 where f1: "continuous_on S f1" "f1 ` S \<subseteq> T" |
|
6755 and g1: "continuous_on T g1" "g1 ` T \<subseteq> S" |
|
6756 and hom1: "homotopic_with (\<lambda>x. True) S S (g1 \<circ> f1) id" |
|
6757 "homotopic_with (\<lambda>x. True) T T (f1 \<circ> g1) id" |
|
6758 using ST by (auto simp: homotopy_eqv_def) |
|
6759 obtain f2 g2 where f2: "continuous_on T f2" "f2 ` T \<subseteq> U" |
|
6760 and g2: "continuous_on U g2" "g2 ` U \<subseteq> T" |
|
6761 and hom2: "homotopic_with (\<lambda>x. True) T T (g2 \<circ> f2) id" |
|
6762 "homotopic_with (\<lambda>x. True) U U (f2 \<circ> g2) id" |
|
6763 using TU by (auto simp: homotopy_eqv_def) |
|
6764 have "homotopic_with (\<lambda>f. True) S T (g2 \<circ> f2 \<circ> f1) (id \<circ> f1)" |
|
6765 by (rule homotopic_with_compose_continuous_right hom2 f1)+ |
|
6766 then have "homotopic_with (\<lambda>f. True) S T (g2 \<circ> (f2 \<circ> f1)) (id \<circ> f1)" |
|
6767 by (simp add: o_assoc) |
|
6768 then have "homotopic_with (\<lambda>x. True) S S |
|
6769 (g1 \<circ> (g2 \<circ> (f2 \<circ> f1))) (g1 \<circ> (id \<circ> f1))" |
|
6770 by (simp add: g1 homotopic_with_compose_continuous_left) |
|
6771 moreover have "homotopic_with (\<lambda>x. True) S S (g1 \<circ> id \<circ> f1) id" |
|
6772 using hom1 by simp |
|
6773 ultimately have SS: "homotopic_with (\<lambda>x. True) S S (g1 \<circ> g2 \<circ> (f2 \<circ> f1)) id" |
|
6774 apply (simp add: o_assoc) |
|
6775 apply (blast intro: homotopic_with_trans) |
|
6776 done |
|
6777 have "homotopic_with (\<lambda>f. True) U T (f1 \<circ> g1 \<circ> g2) (id \<circ> g2)" |
|
6778 by (rule homotopic_with_compose_continuous_right hom1 g2)+ |
|
6779 then have "homotopic_with (\<lambda>f. True) U T (f1 \<circ> (g1 \<circ> g2)) (id \<circ> g2)" |
|
6780 by (simp add: o_assoc) |
|
6781 then have "homotopic_with (\<lambda>x. True) U U |
|
6782 (f2 \<circ> (f1 \<circ> (g1 \<circ> g2))) (f2 \<circ> (id \<circ> g2))" |
|
6783 by (simp add: f2 homotopic_with_compose_continuous_left) |
|
6784 moreover have "homotopic_with (\<lambda>x. True) U U (f2 \<circ> id \<circ> g2) id" |
|
6785 using hom2 by simp |
|
6786 ultimately have UU: "homotopic_with (\<lambda>x. True) U U (f2 \<circ> f1 \<circ> (g1 \<circ> g2)) id" |
|
6787 apply (simp add: o_assoc) |
|
6788 apply (blast intro: homotopic_with_trans) |
|
6789 done |
|
6790 show ?thesis |
|
6791 unfolding homotopy_eqv_def |
|
6792 apply (rule_tac x = "f2 \<circ> f1" in exI) |
|
6793 apply (rule_tac x = "g1 \<circ> g2" in exI) |
|
6794 apply (intro conjI continuous_on_compose SS UU) |
|
6795 using f1 f2 g1 g2 apply (force simp: elim!: continuous_on_subset)+ |
|
6796 done |
|
6797 qed |
|
6798 |
|
6799 lemma homotopy_eqv_inj_linear_image: |
|
6800 fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" |
|
6801 assumes "linear f" "inj f" |
|
6802 shows "(f ` S) homotopy_eqv S" |
|
6803 apply (rule homeomorphic_imp_homotopy_eqv) |
|
6804 using assms homeomorphic_sym linear_homeomorphic_image by auto |
|
6805 |
|
6806 lemma homotopy_eqv_translation: |
|
6807 fixes S :: "'a::real_normed_vector set" |
|
6808 shows "(+) a ` S homotopy_eqv S" |
|
6809 apply (rule homeomorphic_imp_homotopy_eqv) |
|
6810 using homeomorphic_translation homeomorphic_sym by blast |
|
6811 |
|
6812 lemma homotopy_eqv_homotopic_triviality_imp: |
|
6813 fixes S :: "'a::real_normed_vector set" |
|
6814 and T :: "'b::real_normed_vector set" |
|
6815 and U :: "'c::real_normed_vector set" |
|
6816 assumes "S homotopy_eqv T" |
|
6817 and f: "continuous_on U f" "f ` U \<subseteq> T" |
|
6818 and g: "continuous_on U g" "g ` U \<subseteq> T" |
|
6819 and homUS: "\<And>f g. \<lbrakk>continuous_on U f; f ` U \<subseteq> S; |
|
6820 continuous_on U g; g ` U \<subseteq> S\<rbrakk> |
|
6821 \<Longrightarrow> homotopic_with (\<lambda>x. True) U S f g" |
|
6822 shows "homotopic_with (\<lambda>x. True) U T f g" |
|
6823 proof - |
|
6824 obtain h k where h: "continuous_on S h" "h ` S \<subseteq> T" |
|
6825 and k: "continuous_on T k" "k ` T \<subseteq> S" |
|
6826 and hom: "homotopic_with (\<lambda>x. True) S S (k \<circ> h) id" |
|
6827 "homotopic_with (\<lambda>x. True) T T (h \<circ> k) id" |
|
6828 using assms by (auto simp: homotopy_eqv_def) |
|
6829 have "homotopic_with (\<lambda>f. True) U S (k \<circ> f) (k \<circ> g)" |
|
6830 apply (rule homUS) |
|
6831 using f g k |
|
6832 apply (safe intro!: continuous_on_compose h k f elim!: continuous_on_subset) |
|
6833 apply (force simp: o_def)+ |
|
6834 done |
|
6835 then have "homotopic_with (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (k \<circ> g))" |
|
6836 apply (rule homotopic_with_compose_continuous_left) |
|
6837 apply (simp_all add: h) |
|
6838 done |
|
6839 moreover have "homotopic_with (\<lambda>x. True) U T (h \<circ> k \<circ> f) (id \<circ> f)" |
|
6840 apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) |
|
6841 apply (auto simp: hom f) |
|
6842 done |
|
6843 moreover have "homotopic_with (\<lambda>x. True) U T (h \<circ> k \<circ> g) (id \<circ> g)" |
|
6844 apply (rule homotopic_with_compose_continuous_right [where X=T and Y=T]) |
|
6845 apply (auto simp: hom g) |
|
6846 done |
|
6847 ultimately show "homotopic_with (\<lambda>x. True) U T f g" |
|
6848 apply (simp add: o_assoc) |
|
6849 using homotopic_with_trans homotopic_with_sym by blast |
|
6850 qed |
|
6851 |
|
6852 lemma homotopy_eqv_homotopic_triviality: |
|
6853 fixes S :: "'a::real_normed_vector set" |
|
6854 and T :: "'b::real_normed_vector set" |
|
6855 and U :: "'c::real_normed_vector set" |
|
6856 assumes "S homotopy_eqv T" |
|
6857 shows "(\<forall>f g. continuous_on U f \<and> f ` U \<subseteq> S \<and> |
|
6858 continuous_on U g \<and> g ` U \<subseteq> S |
|
6859 \<longrightarrow> homotopic_with (\<lambda>x. True) U S f g) \<longleftrightarrow> |
|
6860 (\<forall>f g. continuous_on U f \<and> f ` U \<subseteq> T \<and> |
|
6861 continuous_on U g \<and> g ` U \<subseteq> T |
|
6862 \<longrightarrow> homotopic_with (\<lambda>x. True) U T f g)" |
|
6863 apply (rule iffI) |
|
6864 apply (metis assms homotopy_eqv_homotopic_triviality_imp) |
|
6865 by (metis (no_types) assms homotopy_eqv_homotopic_triviality_imp homotopy_eqv_sym) |
|
6866 |
|
6867 lemma homotopy_eqv_cohomotopic_triviality_null_imp: |
|
6868 fixes S :: "'a::real_normed_vector set" |
|
6869 and T :: "'b::real_normed_vector set" |
|
6870 and U :: "'c::real_normed_vector set" |
|
6871 assumes "S homotopy_eqv T" |
|
6872 and f: "continuous_on T f" "f ` T \<subseteq> U" |
|
6873 and homSU: "\<And>f. \<lbrakk>continuous_on S f; f ` S \<subseteq> U\<rbrakk> |
|
6874 \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) S U f (\<lambda>x. c)" |
|
6875 obtains c where "homotopic_with (\<lambda>x. True) T U f (\<lambda>x. c)" |
|
6876 proof - |
|
6877 obtain h k where h: "continuous_on S h" "h ` S \<subseteq> T" |
|
6878 and k: "continuous_on T k" "k ` T \<subseteq> S" |
|
6879 and hom: "homotopic_with (\<lambda>x. True) S S (k \<circ> h) id" |
|
6880 "homotopic_with (\<lambda>x. True) T T (h \<circ> k) id" |
|
6881 using assms by (auto simp: homotopy_eqv_def) |
|
6882 obtain c where "homotopic_with (\<lambda>x. True) S U (f \<circ> h) (\<lambda>x. c)" |
|
6883 apply (rule exE [OF homSU [of "f \<circ> h"]]) |
|
6884 apply (intro continuous_on_compose h) |
|
6885 using h f apply (force elim!: continuous_on_subset)+ |
|
6886 done |
|
6887 then have "homotopic_with (\<lambda>x. True) T U ((f \<circ> h) \<circ> k) ((\<lambda>x. c) \<circ> k)" |
|
6888 apply (rule homotopic_with_compose_continuous_right [where X=S]) |
|
6889 using k by auto |
|
6890 moreover have "homotopic_with (\<lambda>x. True) T U (f \<circ> id) (f \<circ> (h \<circ> k))" |
|
6891 apply (rule homotopic_with_compose_continuous_left [where Y=T]) |
|
6892 apply (simp add: hom homotopic_with_symD) |
|
6893 using f apply auto |
|
6894 done |
|
6895 ultimately show ?thesis |
|
6896 apply (rule_tac c=c in that) |
|
6897 apply (simp add: o_def) |
|
6898 using homotopic_with_trans by blast |
|
6899 qed |
|
6900 |
|
6901 lemma homotopy_eqv_cohomotopic_triviality_null: |
|
6902 fixes S :: "'a::real_normed_vector set" |
|
6903 and T :: "'b::real_normed_vector set" |
|
6904 and U :: "'c::real_normed_vector set" |
|
6905 assumes "S homotopy_eqv T" |
|
6906 shows "(\<forall>f. continuous_on S f \<and> f ` S \<subseteq> U |
|
6907 \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) S U f (\<lambda>x. c))) \<longleftrightarrow> |
|
6908 (\<forall>f. continuous_on T f \<and> f ` T \<subseteq> U |
|
6909 \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) T U f (\<lambda>x. c)))" |
|
6910 apply (rule iffI) |
|
6911 apply (metis assms homotopy_eqv_cohomotopic_triviality_null_imp) |
|
6912 by (metis assms homotopy_eqv_cohomotopic_triviality_null_imp homotopy_eqv_sym) |
|
6913 |
|
6914 lemma homotopy_eqv_homotopic_triviality_null_imp: |
|
6915 fixes S :: "'a::real_normed_vector set" |
|
6916 and T :: "'b::real_normed_vector set" |
|
6917 and U :: "'c::real_normed_vector set" |
|
6918 assumes "S homotopy_eqv T" |
|
6919 and f: "continuous_on U f" "f ` U \<subseteq> T" |
|
6920 and homSU: "\<And>f. \<lbrakk>continuous_on U f; f ` U \<subseteq> S\<rbrakk> |
|
6921 \<Longrightarrow> \<exists>c. homotopic_with (\<lambda>x. True) U S f (\<lambda>x. c)" |
|
6922 shows "\<exists>c. homotopic_with (\<lambda>x. True) U T f (\<lambda>x. c)" |
|
6923 proof - |
|
6924 obtain h k where h: "continuous_on S h" "h ` S \<subseteq> T" |
|
6925 and k: "continuous_on T k" "k ` T \<subseteq> S" |
|
6926 and hom: "homotopic_with (\<lambda>x. True) S S (k \<circ> h) id" |
|
6927 "homotopic_with (\<lambda>x. True) T T (h \<circ> k) id" |
|
6928 using assms by (auto simp: homotopy_eqv_def) |
|
6929 obtain c::'a where "homotopic_with (\<lambda>x. True) U S (k \<circ> f) (\<lambda>x. c)" |
|
6930 apply (rule exE [OF homSU [of "k \<circ> f"]]) |
|
6931 apply (intro continuous_on_compose h) |
|
6932 using k f apply (force elim!: continuous_on_subset)+ |
|
6933 done |
|
6934 then have "homotopic_with (\<lambda>x. True) U T (h \<circ> (k \<circ> f)) (h \<circ> (\<lambda>x. c))" |
|
6935 apply (rule homotopic_with_compose_continuous_left [where Y=S]) |
|
6936 using h by auto |
|
6937 moreover have "homotopic_with (\<lambda>x. True) U T (id \<circ> f) ((h \<circ> k) \<circ> f)" |
|
6938 apply (rule homotopic_with_compose_continuous_right [where X=T]) |
|
6939 apply (simp add: hom homotopic_with_symD) |
|
6940 using f apply auto |
|
6941 done |
|
6942 ultimately show ?thesis |
|
6943 using homotopic_with_trans by (fastforce simp add: o_def) |
|
6944 qed |
|
6945 |
|
6946 lemma homotopy_eqv_homotopic_triviality_null: |
|
6947 fixes S :: "'a::real_normed_vector set" |
|
6948 and T :: "'b::real_normed_vector set" |
|
6949 and U :: "'c::real_normed_vector set" |
|
6950 assumes "S homotopy_eqv T" |
|
6951 shows "(\<forall>f. continuous_on U f \<and> f ` U \<subseteq> S |
|
6952 \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) U S f (\<lambda>x. c))) \<longleftrightarrow> |
|
6953 (\<forall>f. continuous_on U f \<and> f ` U \<subseteq> T |
|
6954 \<longrightarrow> (\<exists>c. homotopic_with (\<lambda>x. True) U T f (\<lambda>x. c)))" |
|
6955 apply (rule iffI) |
|
6956 apply (metis assms homotopy_eqv_homotopic_triviality_null_imp) |
|
6957 by (metis assms homotopy_eqv_homotopic_triviality_null_imp homotopy_eqv_sym) |
|
6958 |
|
6959 lemma homotopy_eqv_contractible_sets: |
|
6960 fixes S :: "'a::real_normed_vector set" |
|
6961 and T :: "'b::real_normed_vector set" |
|
6962 assumes "contractible S" "contractible T" "S = {} \<longleftrightarrow> T = {}" |
|
6963 shows "S homotopy_eqv T" |
|
6964 proof (cases "S = {}") |
|
6965 case True with assms show ?thesis |
|
6966 by (simp add: homeomorphic_imp_homotopy_eqv) |
|
6967 next |
|
6968 case False |
|
6969 with assms obtain a b where "a \<in> S" "b \<in> T" |
|
6970 by auto |
|
6971 then show ?thesis |
|
6972 unfolding homotopy_eqv_def |
|
6973 apply (rule_tac x="\<lambda>x. b" in exI) |
|
6974 apply (rule_tac x="\<lambda>x. a" in exI) |
|
6975 apply (intro assms conjI continuous_on_id' homotopic_into_contractible) |
|
6976 apply (auto simp: o_def continuous_on_const) |
|
6977 done |
|
6978 qed |
|
6979 |
|
6980 lemma homotopy_eqv_empty1 [simp]: |
|
6981 fixes S :: "'a::real_normed_vector set" |
|
6982 shows "S homotopy_eqv ({}::'b::real_normed_vector set) \<longleftrightarrow> S = {}" |
|
6983 apply (rule iffI) |
|
6984 using homotopy_eqv_def apply fastforce |
|
6985 by (simp add: homotopy_eqv_contractible_sets) |
|
6986 |
|
6987 lemma homotopy_eqv_empty2 [simp]: |
|
6988 fixes S :: "'a::real_normed_vector set" |
|
6989 shows "({}::'b::real_normed_vector set) homotopy_eqv S \<longleftrightarrow> S = {}" |
|
6990 by (metis homotopy_eqv_empty1 homotopy_eqv_sym) |
|
6991 |
|
6992 lemma homotopy_eqv_contractibility: |
|
6993 fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
6994 shows "S homotopy_eqv T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)" |
|
6995 unfolding homotopy_eqv_def |
|
6996 by (blast intro: homotopy_dominated_contractibility) |
|
6997 |
|
6998 lemma homotopy_eqv_sing: |
|
6999 fixes S :: "'a::real_normed_vector set" and a :: "'b::real_normed_vector" |
|
7000 shows "S homotopy_eqv {a} \<longleftrightarrow> S \<noteq> {} \<and> contractible S" |
|
7001 proof (cases "S = {}") |
|
7002 case True then show ?thesis |
|
7003 by simp |
|
7004 next |
|
7005 case False then show ?thesis |
|
7006 by (metis contractible_sing empty_not_insert homotopy_eqv_contractibility homotopy_eqv_contractible_sets) |
|
7007 qed |
|
7008 |
|
7009 lemma homeomorphic_contractible_eq: |
|
7010 fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
7011 shows "S homeomorphic T \<Longrightarrow> (contractible S \<longleftrightarrow> contractible T)" |
|
7012 by (simp add: homeomorphic_imp_homotopy_eqv homotopy_eqv_contractibility) |
|
7013 |
|
7014 lemma homeomorphic_contractible: |
|
7015 fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" |
|
7016 shows "\<lbrakk>contractible S; S homeomorphic T\<rbrakk> \<Longrightarrow> contractible T" |
|
7017 by (metis homeomorphic_contractible_eq) |
|
7018 |
|
7019 |
|
7020 subsection%unimportant\<open>Misc other results\<close> |
|
7021 |
|
7022 lemma bounded_connected_Compl_real: |
|
7023 fixes S :: "real set" |
|
7024 assumes "bounded S" and conn: "connected(- S)" |
|
7025 shows "S = {}" |
|
7026 proof - |
|
7027 obtain a b where "S \<subseteq> box a b" |
|
7028 by (meson assms bounded_subset_box_symmetric) |
|
7029 then have "a \<notin> S" "b \<notin> S" |
|
7030 by auto |
|
7031 then have "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> x \<in> - S" |
|
7032 by (meson Compl_iff conn connected_iff_interval) |
|
7033 then show ?thesis |
|
7034 using \<open>S \<subseteq> box a b\<close> by auto |
|
7035 qed |
|
7036 |
|
7037 lemma bounded_connected_Compl_1: |
|
7038 fixes S :: "'a::{euclidean_space} set" |
|
7039 assumes "bounded S" and conn: "connected(- S)" and 1: "DIM('a) = 1" |
|
7040 shows "S = {}" |
|
7041 proof - |
|
7042 have "DIM('a) = DIM(real)" |
|
7043 by (simp add: "1") |
|
7044 then obtain f::"'a \<Rightarrow> real" and g |
|
7045 where "linear f" "\<And>x. norm(f x) = norm x" "\<And>x. g(f x) = x" "\<And>y. f(g y) = y" |
|
7046 by (rule isomorphisms_UNIV_UNIV) blast |
|
7047 with \<open>bounded S\<close> have "bounded (f ` S)" |
|
7048 using bounded_linear_image linear_linear by blast |
|
7049 have "connected (f ` (-S))" |
|
7050 using connected_linear_image assms \<open>linear f\<close> by blast |
|
7051 moreover have "f ` (-S) = - (f ` S)" |
|
7052 apply (rule bij_image_Compl_eq) |
|
7053 apply (auto simp: bij_def) |
|
7054 apply (metis \<open>\<And>x. g (f x) = x\<close> injI) |
|
7055 by (metis UNIV_I \<open>\<And>y. f (g y) = y\<close> image_iff) |
|
7056 finally have "connected (- (f ` S))" |
|
7057 by simp |
|
7058 then have "f ` S = {}" |
|
7059 using \<open>bounded (f ` S)\<close> bounded_connected_Compl_real by blast |
|
7060 then show ?thesis |
|
7061 by blast |
|
7062 qed |
|
7063 |
|
7064 |
|
7065 subsection%unimportant\<open>Some Uncountable Sets\<close> |
|
7066 |
|
7067 lemma uncountable_closed_segment: |
|
7068 fixes a :: "'a::real_normed_vector" |
|
7069 assumes "a \<noteq> b" shows "uncountable (closed_segment a b)" |
|
7070 unfolding path_image_linepath [symmetric] path_image_def |
|
7071 using inj_on_linepath [OF assms] uncountable_closed_interval [of 0 1] |
|
7072 countable_image_inj_on by auto |
|
7073 |
|
7074 lemma uncountable_open_segment: |
|
7075 fixes a :: "'a::real_normed_vector" |
|
7076 assumes "a \<noteq> b" shows "uncountable (open_segment a b)" |
|
7077 by (simp add: assms open_segment_def uncountable_closed_segment uncountable_minus_countable) |
|
7078 |
|
7079 lemma uncountable_convex: |
|
7080 fixes a :: "'a::real_normed_vector" |
|
7081 assumes "convex S" "a \<in> S" "b \<in> S" "a \<noteq> b" |
|
7082 shows "uncountable S" |
|
7083 proof - |
|
7084 have "uncountable (closed_segment a b)" |
|
7085 by (simp add: uncountable_closed_segment assms) |
|
7086 then show ?thesis |
|
7087 by (meson assms convex_contains_segment countable_subset) |
|
7088 qed |
|
7089 |
|
7090 lemma uncountable_ball: |
|
7091 fixes a :: "'a::euclidean_space" |
|
7092 assumes "r > 0" |
|
7093 shows "uncountable (ball a r)" |
|
7094 proof - |
|
7095 have "uncountable (open_segment a (a + r *\<^sub>R (SOME i. i \<in> Basis)))" |
|
7096 by (metis Basis_zero SOME_Basis add_cancel_right_right assms less_le scale_eq_0_iff uncountable_open_segment) |
|
7097 moreover have "open_segment a (a + r *\<^sub>R (SOME i. i \<in> Basis)) \<subseteq> ball a r" |
|
7098 using assms by (auto simp: in_segment algebra_simps dist_norm SOME_Basis) |
|
7099 ultimately show ?thesis |
|
7100 by (metis countable_subset) |
|
7101 qed |
|
7102 |
|
7103 lemma ball_minus_countable_nonempty: |
|
7104 assumes "countable (A :: 'a :: euclidean_space set)" "r > 0" |
|
7105 shows "ball z r - A \<noteq> {}" |
|
7106 proof |
|
7107 assume *: "ball z r - A = {}" |
|
7108 have "uncountable (ball z r - A)" |
|
7109 by (intro uncountable_minus_countable assms uncountable_ball) |
|
7110 thus False by (subst (asm) *) auto |
|
7111 qed |
|
7112 |
|
7113 lemma uncountable_cball: |
|
7114 fixes a :: "'a::euclidean_space" |
|
7115 assumes "r > 0" |
|
7116 shows "uncountable (cball a r)" |
|
7117 using assms countable_subset uncountable_ball by auto |
|
7118 |
|
7119 lemma pairwise_disjnt_countable: |
|
7120 fixes \<N> :: "nat set set" |
|
7121 assumes "pairwise disjnt \<N>" |
|
7122 shows "countable \<N>" |
|
7123 proof - |
|
7124 have "inj_on (\<lambda>X. SOME n. n \<in> X) (\<N> - {{}})" |
|
7125 apply (clarsimp simp add: inj_on_def) |
|
7126 by (metis assms disjnt_insert2 insert_absorb pairwise_def subsetI subset_empty tfl_some) |
|
7127 then show ?thesis |
|
7128 by (metis countable_Diff_eq countable_def) |
|
7129 qed |
|
7130 |
|
7131 lemma pairwise_disjnt_countable_Union: |
|
7132 assumes "countable (\<Union>\<N>)" and pwd: "pairwise disjnt \<N>" |
|
7133 shows "countable \<N>" |
|
7134 proof - |
|
7135 obtain f :: "_ \<Rightarrow> nat" where f: "inj_on f (\<Union>\<N>)" |
|
7136 using assms by blast |
|
7137 then have "pairwise disjnt (\<Union> X \<in> \<N>. {f ` X})" |
|
7138 using assms by (force simp: pairwise_def disjnt_inj_on_iff [OF f]) |
|
7139 then have "countable (\<Union> X \<in> \<N>. {f ` X})" |
|
7140 using pairwise_disjnt_countable by blast |
|
7141 then show ?thesis |
|
7142 by (meson pwd countable_image_inj_on disjoint_image f inj_on_image pairwise_disjnt_countable) |
|
7143 qed |
|
7144 |
|
7145 lemma connected_uncountable: |
|
7146 fixes S :: "'a::metric_space set" |
|
7147 assumes "connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S" |
|
7148 proof - |
|
7149 have "continuous_on S (dist a)" |
|
7150 by (intro continuous_intros) |
|
7151 then have "connected (dist a ` S)" |
|
7152 by (metis connected_continuous_image \<open>connected S\<close>) |
|
7153 then have "closed_segment 0 (dist a b) \<subseteq> (dist a ` S)" |
|
7154 by (simp add: assms closed_segment_subset is_interval_connected_1 is_interval_convex) |
|
7155 then have "uncountable (dist a ` S)" |
|
7156 by (metis \<open>a \<noteq> b\<close> countable_subset dist_eq_0_iff uncountable_closed_segment) |
|
7157 then show ?thesis |
|
7158 by blast |
|
7159 qed |
|
7160 |
|
7161 lemma path_connected_uncountable: |
|
7162 fixes S :: "'a::metric_space set" |
|
7163 assumes "path_connected S" "a \<in> S" "b \<in> S" "a \<noteq> b" shows "uncountable S" |
|
7164 using path_connected_imp_connected assms connected_uncountable by metis |
|
7165 |
|
7166 lemma connected_finite_iff_sing: |
|
7167 fixes S :: "'a::metric_space set" |
|
7168 assumes "connected S" |
|
7169 shows "finite S \<longleftrightarrow> S = {} \<or> (\<exists>a. S = {a})" (is "_ = ?rhs") |
|
7170 proof - |
|
7171 have "uncountable S" if "\<not> ?rhs" |
|
7172 using connected_uncountable assms that by blast |
|
7173 then show ?thesis |
|
7174 using uncountable_infinite by auto |
|
7175 qed |
|
7176 |
|
7177 lemma connected_card_eq_iff_nontrivial: |
|
7178 fixes S :: "'a::metric_space set" |
|
7179 shows "connected S \<Longrightarrow> uncountable S \<longleftrightarrow> \<not>(\<exists>a. S \<subseteq> {a})" |
|
7180 apply (auto simp: countable_finite finite_subset) |
|
7181 by (metis connected_uncountable is_singletonI' is_singleton_the_elem subset_singleton_iff) |
|
7182 |
|
7183 lemma simple_path_image_uncountable: |
|
7184 fixes g :: "real \<Rightarrow> 'a::metric_space" |
|
7185 assumes "simple_path g" |
|
7186 shows "uncountable (path_image g)" |
|
7187 proof - |
|
7188 have "g 0 \<in> path_image g" "g (1/2) \<in> path_image g" |
|
7189 by (simp_all add: path_defs) |
|
7190 moreover have "g 0 \<noteq> g (1/2)" |
|
7191 using assms by (fastforce simp add: simple_path_def) |
|
7192 ultimately show ?thesis |
|
7193 apply (simp add: assms connected_card_eq_iff_nontrivial connected_simple_path_image) |
|
7194 by blast |
|
7195 qed |
|
7196 |
|
7197 lemma arc_image_uncountable: |
|
7198 fixes g :: "real \<Rightarrow> 'a::metric_space" |
|
7199 assumes "arc g" |
|
7200 shows "uncountable (path_image g)" |
|
7201 by (simp add: arc_imp_simple_path assms simple_path_image_uncountable) |
|
7202 |
|
7203 |
|
7204 subsection%unimportant\<open> Some simple positive connection theorems\<close> |
|
7205 |
|
7206 proposition path_connected_convex_diff_countable: |
|
7207 fixes U :: "'a::euclidean_space set" |
|
7208 assumes "convex U" "\<not> collinear U" "countable S" |
|
7209 shows "path_connected(U - S)" |
|
7210 proof (clarsimp simp add: path_connected_def) |
|
7211 fix a b |
|
7212 assume "a \<in> U" "a \<notin> S" "b \<in> U" "b \<notin> S" |
|
7213 let ?m = "midpoint a b" |
|
7214 show "\<exists>g. path g \<and> path_image g \<subseteq> U - S \<and> pathstart g = a \<and> pathfinish g = b" |
|
7215 proof (cases "a = b") |
|
7216 case True |
|
7217 then show ?thesis |
|
7218 by (metis DiffI \<open>a \<in> U\<close> \<open>a \<notin> S\<close> path_component_def path_component_refl) |
|
7219 next |
|
7220 case False |
|
7221 then have "a \<noteq> ?m" "b \<noteq> ?m" |
|
7222 using midpoint_eq_endpoint by fastforce+ |
|
7223 have "?m \<in> U" |
|
7224 using \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>convex U\<close> convex_contains_segment by force |
|
7225 obtain c where "c \<in> U" and nc_abc: "\<not> collinear {a,b,c}" |
|
7226 by (metis False \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>\<not> collinear U\<close> collinear_triples insert_absorb) |
|
7227 have ncoll_mca: "\<not> collinear {?m,c,a}" |
|
7228 by (metis (full_types) \<open>a \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc) |
|
7229 have ncoll_mcb: "\<not> collinear {?m,c,b}" |
|
7230 by (metis (full_types) \<open>b \<noteq> ?m\<close> collinear_3_trans collinear_midpoint insert_commute nc_abc) |
|
7231 have "c \<noteq> ?m" |
|
7232 by (metis collinear_midpoint insert_commute nc_abc) |
|
7233 then have "closed_segment ?m c \<subseteq> U" |
|
7234 by (simp add: \<open>c \<in> U\<close> \<open>?m \<in> U\<close> \<open>convex U\<close> closed_segment_subset) |
|
7235 then obtain z where z: "z \<in> closed_segment ?m c" |
|
7236 and disjS: "(closed_segment a z \<union> closed_segment z b) \<inter> S = {}" |
|
7237 proof - |
|
7238 have False if "closed_segment ?m c \<subseteq> {z. (closed_segment a z \<union> closed_segment z b) \<inter> S \<noteq> {}}" |
|
7239 proof - |
|
7240 have closb: "closed_segment ?m c \<subseteq> |
|
7241 {z \<in> closed_segment ?m c. closed_segment a z \<inter> S \<noteq> {}} \<union> {z \<in> closed_segment ?m c. closed_segment z b \<inter> S \<noteq> {}}" |
|
7242 using that by blast |
|
7243 have *: "countable {z \<in> closed_segment ?m c. closed_segment z u \<inter> S \<noteq> {}}" |
|
7244 if "u \<in> U" "u \<notin> S" and ncoll: "\<not> collinear {?m, c, u}" for u |
|
7245 proof - |
|
7246 have **: False if x1: "x1 \<in> closed_segment ?m c" and x2: "x2 \<in> closed_segment ?m c" |
|
7247 and "x1 \<noteq> x2" "x1 \<noteq> u" |
|
7248 and w: "w \<in> closed_segment x1 u" "w \<in> closed_segment x2 u" |
|
7249 and "w \<in> S" for x1 x2 w |
|
7250 proof - |
|
7251 have "x1 \<in> affine hull {?m,c}" "x2 \<in> affine hull {?m,c}" |
|
7252 using segment_as_ball x1 x2 by auto |
|
7253 then have coll_x1: "collinear {x1, ?m, c}" and coll_x2: "collinear {?m, c, x2}" |
|
7254 by (simp_all add: affine_hull_3_imp_collinear) (metis affine_hull_3_imp_collinear insert_commute) |
|
7255 have "\<not> collinear {x1, u, x2}" |
|
7256 proof |
|
7257 assume "collinear {x1, u, x2}" |
|
7258 then have "collinear {?m, c, u}" |
|
7259 by (metis (full_types) \<open>c \<noteq> ?m\<close> coll_x1 coll_x2 collinear_3_trans insert_commute ncoll \<open>x1 \<noteq> x2\<close>) |
|
7260 with ncoll show False .. |
|
7261 qed |
|
7262 then have "closed_segment x1 u \<inter> closed_segment u x2 = {u}" |
|
7263 by (blast intro!: Int_closed_segment) |
|
7264 then have "w = u" |
|
7265 using closed_segment_commute w by auto |
|
7266 show ?thesis |
|
7267 using \<open>u \<notin> S\<close> \<open>w = u\<close> that(7) by auto |
|
7268 qed |
|
7269 then have disj: "disjoint ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}))" |
|
7270 by (fastforce simp: pairwise_def disjnt_def) |
|
7271 have cou: "countable ((\<Union>z \<in> closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})" |
|
7272 apply (rule pairwise_disjnt_countable_Union [OF _ pairwise_subset [OF disj]]) |
|
7273 apply (rule countable_subset [OF _ \<open>countable S\<close>], auto) |
|
7274 done |
|
7275 define f where "f \<equiv> \<lambda>X. (THE z. z \<in> closed_segment ?m c \<and> X = closed_segment z u \<inter> S)" |
|
7276 show ?thesis |
|
7277 proof (rule countable_subset [OF _ countable_image [OF cou, where f=f]], clarify) |
|
7278 fix x |
|
7279 assume x: "x \<in> closed_segment ?m c" "closed_segment x u \<inter> S \<noteq> {}" |
|
7280 show "x \<in> f ` ((\<Union>z\<in>closed_segment ?m c. {closed_segment z u \<inter> S}) - {{}})" |
|
7281 proof (rule_tac x="closed_segment x u \<inter> S" in image_eqI) |
|
7282 show "x = f (closed_segment x u \<inter> S)" |
|
7283 unfolding f_def |
|
7284 apply (rule the_equality [symmetric]) |
|
7285 using x apply (auto simp: dest: **) |
|
7286 done |
|
7287 qed (use x in auto) |
|
7288 qed |
|
7289 qed |
|
7290 have "uncountable (closed_segment ?m c)" |
|
7291 by (metis \<open>c \<noteq> ?m\<close> uncountable_closed_segment) |
|
7292 then show False |
|
7293 using closb * [OF \<open>a \<in> U\<close> \<open>a \<notin> S\<close> ncoll_mca] * [OF \<open>b \<in> U\<close> \<open>b \<notin> S\<close> ncoll_mcb] |
|
7294 apply (simp add: closed_segment_commute) |
|
7295 by (simp add: countable_subset) |
|
7296 qed |
|
7297 then show ?thesis |
|
7298 by (force intro: that) |
|
7299 qed |
|
7300 show ?thesis |
|
7301 proof (intro exI conjI) |
|
7302 have "path_image (linepath a z +++ linepath z b) \<subseteq> U" |
|
7303 by (metis \<open>a \<in> U\<close> \<open>b \<in> U\<close> \<open>closed_segment ?m c \<subseteq> U\<close> z \<open>convex U\<close> closed_segment_subset contra_subsetD path_image_linepath subset_path_image_join) |
|
7304 with disjS show "path_image (linepath a z +++ linepath z b) \<subseteq> U - S" |
|
7305 by (force simp: path_image_join) |
|
7306 qed auto |
|
7307 qed |
|
7308 qed |
|
7309 |
|
7310 |
|
7311 corollary connected_convex_diff_countable: |
|
7312 fixes U :: "'a::euclidean_space set" |
|
7313 assumes "convex U" "\<not> collinear U" "countable S" |
|
7314 shows "connected(U - S)" |
|
7315 by (simp add: assms path_connected_convex_diff_countable path_connected_imp_connected) |
|
7316 |
|
7317 lemma path_connected_punctured_convex: |
|
7318 assumes "convex S" and aff: "aff_dim S \<noteq> 1" |
|
7319 shows "path_connected(S - {a})" |
|
7320 proof - |
|
7321 consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S \<ge> 2" |
|
7322 using assms aff_dim_geq [of S] by linarith |
|
7323 then show ?thesis |
|
7324 proof cases |
|
7325 assume "aff_dim S = -1" |
|
7326 then show ?thesis |
|
7327 by (metis aff_dim_empty empty_Diff path_connected_empty) |
|
7328 next |
|
7329 assume "aff_dim S = 0" |
|
7330 then show ?thesis |
|
7331 by (metis aff_dim_eq_0 Diff_cancel Diff_empty Diff_insert0 convex_empty convex_imp_path_connected path_connected_singleton singletonD) |
|
7332 next |
|
7333 assume ge2: "aff_dim S \<ge> 2" |
|
7334 then have "\<not> collinear S" |
|
7335 proof (clarsimp simp add: collinear_affine_hull) |
|
7336 fix u v |
|
7337 assume "S \<subseteq> affine hull {u, v}" |
|
7338 then have "aff_dim S \<le> aff_dim {u, v}" |
|
7339 by (metis (no_types) aff_dim_affine_hull aff_dim_subset) |
|
7340 with ge2 show False |
|
7341 by (metis (no_types) aff_dim_2 antisym aff not_numeral_le_zero one_le_numeral order_trans) |
|
7342 qed |
|
7343 then show ?thesis |
|
7344 apply (rule path_connected_convex_diff_countable [OF \<open>convex S\<close>]) |
|
7345 by simp |
|
7346 qed |
|
7347 qed |
|
7348 |
|
7349 lemma connected_punctured_convex: |
|
7350 shows "\<lbrakk>convex S; aff_dim S \<noteq> 1\<rbrakk> \<Longrightarrow> connected(S - {a})" |
|
7351 using path_connected_imp_connected path_connected_punctured_convex by blast |
|
7352 |
|
7353 lemma path_connected_complement_countable: |
|
7354 fixes S :: "'a::euclidean_space set" |
|
7355 assumes "2 \<le> DIM('a)" "countable S" |
|
7356 shows "path_connected(- S)" |
|
7357 proof - |
|
7358 have "path_connected(UNIV - S)" |
|
7359 apply (rule path_connected_convex_diff_countable) |
|
7360 using assms by (auto simp: collinear_aff_dim [of "UNIV :: 'a set"]) |
|
7361 then show ?thesis |
|
7362 by (simp add: Compl_eq_Diff_UNIV) |
|
7363 qed |
|
7364 |
|
7365 proposition path_connected_openin_diff_countable: |
|
7366 fixes S :: "'a::euclidean_space set" |
|
7367 assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" |
|
7368 and "\<not> collinear S" "countable T" |
|
7369 shows "path_connected(S - T)" |
|
7370 proof (clarsimp simp add: path_connected_component) |
|
7371 fix x y |
|
7372 assume xy: "x \<in> S" "x \<notin> T" "y \<in> S" "y \<notin> T" |
|
7373 show "path_component (S - T) x y" |
|
7374 proof (rule connected_equivalence_relation_gen [OF \<open>connected S\<close>, where P = "\<lambda>x. x \<notin> T"]) |
|
7375 show "\<exists>z. z \<in> U \<and> z \<notin> T" if opeU: "openin (subtopology euclidean S) U" and "x \<in> U" for U x |
|
7376 proof - |
|
7377 have "openin (subtopology euclidean (affine hull S)) U" |
|
7378 using opeU ope openin_trans by blast |
|
7379 with \<open>x \<in> U\<close> obtain r where Usub: "U \<subseteq> affine hull S" and "r > 0" |
|
7380 and subU: "ball x r \<inter> affine hull S \<subseteq> U" |
|
7381 by (auto simp: openin_contains_ball) |
|
7382 with \<open>x \<in> U\<close> have x: "x \<in> ball x r \<inter> affine hull S" |
|
7383 by auto |
|
7384 have "\<not> S \<subseteq> {x}" |
|
7385 using \<open>\<not> collinear S\<close> collinear_subset by blast |
|
7386 then obtain x' where "x' \<noteq> x" "x' \<in> S" |
|
7387 by blast |
|
7388 obtain y where y: "y \<noteq> x" "y \<in> ball x r \<inter> affine hull S" |
|
7389 proof |
|
7390 show "x + (r / 2 / norm(x' - x)) *\<^sub>R (x' - x) \<noteq> x" |
|
7391 using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> by auto |
|
7392 show "x + (r / 2 / norm (x' - x)) *\<^sub>R (x' - x) \<in> ball x r \<inter> affine hull S" |
|
7393 using \<open>x' \<noteq> x\<close> \<open>r > 0\<close> \<open>x' \<in> S\<close> x |
|
7394 by (simp add: dist_norm mem_affine_3_minus hull_inc) |
|
7395 qed |
|
7396 have "convex (ball x r \<inter> affine hull S)" |
|
7397 by (simp add: affine_imp_convex convex_Int) |
|
7398 with x y subU have "uncountable U" |
|
7399 by (meson countable_subset uncountable_convex) |
|
7400 then have "\<not> U \<subseteq> T" |
|
7401 using \<open>countable T\<close> countable_subset by blast |
|
7402 then show ?thesis by blast |
|
7403 qed |
|
7404 show "\<exists>U. openin (subtopology euclidean S) U \<and> x \<in> U \<and> |
|
7405 (\<forall>x\<in>U. \<forall>y\<in>U. x \<notin> T \<and> y \<notin> T \<longrightarrow> path_component (S - T) x y)" |
|
7406 if "x \<in> S" for x |
|
7407 proof - |
|
7408 obtain r where Ssub: "S \<subseteq> affine hull S" and "r > 0" |
|
7409 and subS: "ball x r \<inter> affine hull S \<subseteq> S" |
|
7410 using ope \<open>x \<in> S\<close> by (auto simp: openin_contains_ball) |
|
7411 then have conv: "convex (ball x r \<inter> affine hull S)" |
|
7412 by (simp add: affine_imp_convex convex_Int) |
|
7413 have "\<not> aff_dim (affine hull S) \<le> 1" |
|
7414 using \<open>\<not> collinear S\<close> collinear_aff_dim by auto |
|
7415 then have "\<not> collinear (ball x r \<inter> affine hull S)" |
|
7416 apply (simp add: collinear_aff_dim) |
|
7417 by (metis (no_types, hide_lams) aff_dim_convex_Int_open IntI open_ball \<open>0 < r\<close> aff_dim_affine_hull affine_affine_hull affine_imp_convex centre_in_ball empty_iff hull_subset inf_commute subsetCE that) |
|
7418 then have *: "path_connected ((ball x r \<inter> affine hull S) - T)" |
|
7419 by (rule path_connected_convex_diff_countable [OF conv _ \<open>countable T\<close>]) |
|
7420 have ST: "ball x r \<inter> affine hull S - T \<subseteq> S - T" |
|
7421 using subS by auto |
|
7422 show ?thesis |
|
7423 proof (intro exI conjI) |
|
7424 show "x \<in> ball x r \<inter> affine hull S" |
|
7425 using \<open>x \<in> S\<close> \<open>r > 0\<close> by (simp add: hull_inc) |
|
7426 have "openin (subtopology euclidean (affine hull S)) (ball x r \<inter> affine hull S)" |
|
7427 by (subst inf.commute) (simp add: openin_Int_open) |
|
7428 then show "openin (subtopology euclidean S) (ball x r \<inter> affine hull S)" |
|
7429 by (rule openin_subset_trans [OF _ subS Ssub]) |
|
7430 qed (use * path_component_trans in \<open>auto simp: path_connected_component path_component_of_subset [OF ST]\<close>) |
|
7431 qed |
|
7432 qed (use xy path_component_trans in auto) |
|
7433 qed |
|
7434 |
|
7435 corollary connected_openin_diff_countable: |
|
7436 fixes S :: "'a::euclidean_space set" |
|
7437 assumes "connected S" and ope: "openin (subtopology euclidean (affine hull S)) S" |
|
7438 and "\<not> collinear S" "countable T" |
|
7439 shows "connected(S - T)" |
|
7440 by (metis path_connected_imp_connected path_connected_openin_diff_countable [OF assms]) |
|
7441 |
|
7442 corollary path_connected_open_diff_countable: |
|
7443 fixes S :: "'a::euclidean_space set" |
|
7444 assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T" |
|
7445 shows "path_connected(S - T)" |
|
7446 proof (cases "S = {}") |
|
7447 case True |
|
7448 then show ?thesis |
|
7449 by (simp add: path_connected_empty) |
|
7450 next |
|
7451 case False |
|
7452 show ?thesis |
|
7453 proof (rule path_connected_openin_diff_countable) |
|
7454 show "openin (subtopology euclidean (affine hull S)) S" |
|
7455 by (simp add: assms hull_subset open_subset) |
|
7456 show "\<not> collinear S" |
|
7457 using assms False by (simp add: collinear_aff_dim aff_dim_open) |
|
7458 qed (simp_all add: assms) |
|
7459 qed |
|
7460 |
|
7461 corollary connected_open_diff_countable: |
|
7462 fixes S :: "'a::euclidean_space set" |
|
7463 assumes "2 \<le> DIM('a)" "open S" "connected S" "countable T" |
|
7464 shows "connected(S - T)" |
|
7465 by (simp add: assms path_connected_imp_connected path_connected_open_diff_countable) |
|
7466 |
|
7467 |
|
7468 |
|
7469 subsection%unimportant \<open>Self-homeomorphisms shuffling points about\<close> |
|
7470 |
|
7471 subsubsection%unimportant\<open>The theorem \<open>homeomorphism_moving_points_exists\<close>\<close> |
|
7472 |
|
7473 lemma homeomorphism_moving_point_1: |
|
7474 fixes a :: "'a::euclidean_space" |
|
7475 assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" |
|
7476 obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
7477 "f a = u" "\<And>x. x \<in> sphere a r \<Longrightarrow> f x = x" |
|
7478 proof - |
|
7479 have nou: "norm (u - a) < r" and "u \<in> T" |
|
7480 using u by (auto simp: dist_norm norm_minus_commute) |
|
7481 then have "0 < r" |
|
7482 by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) |
|
7483 define f where "f \<equiv> \<lambda>x. (1 - norm(x - a) / r) *\<^sub>R (u - a) + x" |
|
7484 have *: "False" if eq: "x + (norm y / r) *\<^sub>R u = y + (norm x / r) *\<^sub>R u" |
|
7485 and nou: "norm u < r" and yx: "norm y < norm x" for x y and u::'a |
|
7486 proof - |
|
7487 have "x = y + (norm x / r - (norm y / r)) *\<^sub>R u" |
|
7488 using eq by (simp add: algebra_simps) |
|
7489 then have "norm x = norm (y + ((norm x - norm y) / r) *\<^sub>R u)" |
|
7490 by (metis diff_divide_distrib) |
|
7491 also have "\<dots> \<le> norm y + norm(((norm x - norm y) / r) *\<^sub>R u)" |
|
7492 using norm_triangle_ineq by blast |
|
7493 also have "\<dots> = norm y + (norm x - norm y) * (norm u / r)" |
|
7494 using yx \<open>r > 0\<close> |
|
7495 by (simp add: divide_simps) |
|
7496 also have "\<dots> < norm y + (norm x - norm y) * 1" |
|
7497 apply (subst add_less_cancel_left) |
|
7498 apply (rule mult_strict_left_mono) |
|
7499 using nou \<open>0 < r\<close> yx |
|
7500 apply (simp_all add: field_simps) |
|
7501 done |
|
7502 also have "\<dots> = norm x" |
|
7503 by simp |
|
7504 finally show False by simp |
|
7505 qed |
|
7506 have "inj f" |
|
7507 unfolding f_def |
|
7508 proof (clarsimp simp: inj_on_def) |
|
7509 fix x y |
|
7510 assume "(1 - norm (x - a) / r) *\<^sub>R (u - a) + x = |
|
7511 (1 - norm (y - a) / r) *\<^sub>R (u - a) + y" |
|
7512 then have eq: "(x - a) + (norm (y - a) / r) *\<^sub>R (u - a) = (y - a) + (norm (x - a) / r) *\<^sub>R (u - a)" |
|
7513 by (auto simp: algebra_simps) |
|
7514 show "x=y" |
|
7515 proof (cases "norm (x - a) = norm (y - a)") |
|
7516 case True |
|
7517 then show ?thesis |
|
7518 using eq by auto |
|
7519 next |
|
7520 case False |
|
7521 then consider "norm (x - a) < norm (y - a)" | "norm (x - a) > norm (y - a)" |
|
7522 by linarith |
|
7523 then have "False" |
|
7524 proof cases |
|
7525 case 1 show False |
|
7526 using * [OF _ nou 1] eq by simp |
|
7527 next |
|
7528 case 2 with * [OF eq nou] show False |
|
7529 by auto |
|
7530 qed |
|
7531 then show "x=y" .. |
|
7532 qed |
|
7533 qed |
|
7534 then have inj_onf: "inj_on f (cball a r \<inter> T)" |
|
7535 using inj_on_Int by fastforce |
|
7536 have contf: "continuous_on (cball a r \<inter> T) f" |
|
7537 unfolding f_def using \<open>0 < r\<close> by (intro continuous_intros) blast |
|
7538 have fim: "f ` (cball a r \<inter> T) = cball a r \<inter> T" |
|
7539 proof |
|
7540 have *: "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> r" if "norm y \<le> r" "norm u < r" for y u::'a |
|
7541 proof - |
|
7542 have "norm (y + (1 - norm y / r) *\<^sub>R u) \<le> norm y + norm((1 - norm y / r) *\<^sub>R u)" |
|
7543 using norm_triangle_ineq by blast |
|
7544 also have "\<dots> = norm y + abs(1 - norm y / r) * norm u" |
|
7545 by simp |
|
7546 also have "\<dots> \<le> r" |
|
7547 proof - |
|
7548 have "(r - norm u) * (r - norm y) \<ge> 0" |
|
7549 using that by auto |
|
7550 then have "r * norm u + r * norm y \<le> r * r + norm u * norm y" |
|
7551 by (simp add: algebra_simps) |
|
7552 then show ?thesis |
|
7553 using that \<open>0 < r\<close> by (simp add: abs_if field_simps) |
|
7554 qed |
|
7555 finally show ?thesis . |
|
7556 qed |
|
7557 have "f ` (cball a r) \<subseteq> cball a r" |
|
7558 apply (clarsimp simp add: dist_norm norm_minus_commute f_def) |
|
7559 using * by (metis diff_add_eq diff_diff_add diff_diff_eq2 norm_minus_commute nou) |
|
7560 moreover have "f ` T \<subseteq> T" |
|
7561 unfolding f_def using \<open>affine T\<close> \<open>a \<in> T\<close> \<open>u \<in> T\<close> |
|
7562 by (force simp: add.commute mem_affine_3_minus) |
|
7563 ultimately show "f ` (cball a r \<inter> T) \<subseteq> cball a r \<inter> T" |
|
7564 by blast |
|
7565 next |
|
7566 show "cball a r \<inter> T \<subseteq> f ` (cball a r \<inter> T)" |
|
7567 proof (clarsimp simp add: dist_norm norm_minus_commute) |
|
7568 fix x |
|
7569 assume x: "norm (x - a) \<le> r" and "x \<in> T" |
|
7570 have "\<exists>v \<in> {0..1}. ((1 - v) * r - norm ((x - a) - v *\<^sub>R (u - a))) \<bullet> 1 = 0" |
|
7571 by (rule ivt_decreasing_component_on_1) (auto simp: x continuous_intros) |
|
7572 then obtain v where "0\<le>v" "v\<le>1" and v: "(1 - v) * r = norm ((x - a) - v *\<^sub>R (u - a))" |
|
7573 by auto |
|
7574 show "x \<in> f ` (cball a r \<inter> T)" |
|
7575 proof (rule image_eqI) |
|
7576 show "x = f (x - v *\<^sub>R (u - a))" |
|
7577 using \<open>r > 0\<close> v by (simp add: f_def field_simps) |
|
7578 have "x - v *\<^sub>R (u - a) \<in> cball a r" |
|
7579 using \<open>r > 0\<close> v \<open>0 \<le> v\<close> |
|
7580 apply (simp add: field_simps dist_norm norm_minus_commute) |
|
7581 by (metis le_add_same_cancel2 order.order_iff_strict zero_le_mult_iff) |
|
7582 moreover have "x - v *\<^sub>R (u - a) \<in> T" |
|
7583 by (simp add: f_def \<open>affine T\<close> \<open>u \<in> T\<close> \<open>x \<in> T\<close> assms mem_affine_3_minus2) |
|
7584 ultimately show "x - v *\<^sub>R (u - a) \<in> cball a r \<inter> T" |
|
7585 by blast |
|
7586 qed |
|
7587 qed |
|
7588 qed |
|
7589 have "\<exists>g. homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
7590 apply (rule homeomorphism_compact [OF _ contf fim inj_onf]) |
|
7591 apply (simp add: affine_closed compact_Int_closed \<open>affine T\<close>) |
|
7592 done |
|
7593 then show ?thesis |
|
7594 apply (rule exE) |
|
7595 apply (erule_tac f=f in that) |
|
7596 using \<open>r > 0\<close> |
|
7597 apply (simp_all add: f_def dist_norm norm_minus_commute) |
|
7598 done |
|
7599 qed |
|
7600 |
|
7601 corollary%unimportant homeomorphism_moving_point_2: |
|
7602 fixes a :: "'a::euclidean_space" |
|
7603 assumes "affine T" "a \<in> T" and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T" |
|
7604 obtains f g where "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
7605 "f u = v" "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x" |
|
7606 proof - |
|
7607 have "0 < r" |
|
7608 by (metis DiffD1 Diff_Diff_Int ball_eq_empty centre_in_ball not_le u) |
|
7609 obtain f1 g1 where hom1: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f1 g1" |
|
7610 and "f1 a = u" and f1: "\<And>x. x \<in> sphere a r \<Longrightarrow> f1 x = x" |
|
7611 using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u] by blast |
|
7612 obtain f2 g2 where hom2: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f2 g2" |
|
7613 and "f2 a = v" and f2: "\<And>x. x \<in> sphere a r \<Longrightarrow> f2 x = x" |
|
7614 using homeomorphism_moving_point_1 [OF \<open>affine T\<close> \<open>a \<in> T\<close> v] by blast |
|
7615 show ?thesis |
|
7616 proof |
|
7617 show "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) (f2 \<circ> g1) (f1 \<circ> g2)" |
|
7618 by (metis homeomorphism_compose homeomorphism_symD hom1 hom2) |
|
7619 have "g1 u = a" |
|
7620 using \<open>0 < r\<close> \<open>f1 a = u\<close> assms hom1 homeomorphism_apply1 by fastforce |
|
7621 then show "(f2 \<circ> g1) u = v" |
|
7622 by (simp add: \<open>f2 a = v\<close>) |
|
7623 show "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> (f2 \<circ> g1) x = x" |
|
7624 using f1 f2 hom1 homeomorphism_apply1 by fastforce |
|
7625 qed |
|
7626 qed |
|
7627 |
|
7628 |
|
7629 corollary%unimportant homeomorphism_moving_point_3: |
|
7630 fixes a :: "'a::euclidean_space" |
|
7631 assumes "affine T" "a \<in> T" and ST: "ball a r \<inter> T \<subseteq> S" "S \<subseteq> T" |
|
7632 and u: "u \<in> ball a r \<inter> T" and v: "v \<in> ball a r \<inter> T" |
|
7633 obtains f g where "homeomorphism S S f g" |
|
7634 "f u = v" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> ball a r \<inter> T" |
|
7635 proof - |
|
7636 obtain f g where hom: "homeomorphism (cball a r \<inter> T) (cball a r \<inter> T) f g" |
|
7637 and "f u = v" and fid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> f x = x" |
|
7638 using homeomorphism_moving_point_2 [OF \<open>affine T\<close> \<open>a \<in> T\<close> u v] by blast |
|
7639 have gid: "\<And>x. \<lbrakk>x \<in> sphere a r; x \<in> T\<rbrakk> \<Longrightarrow> g x = x" |
|
7640 using fid hom homeomorphism_apply1 by fastforce |
|
7641 define ff where "ff \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then f x else x" |
|
7642 define gg where "gg \<equiv> \<lambda>x. if x \<in> ball a r \<inter> T then g x else x" |
|
7643 show ?thesis |
|
7644 proof |
|
7645 show "homeomorphism S S ff gg" |
|
7646 proof (rule homeomorphismI) |
|
7647 have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) ff" |
|
7648 apply (simp add: ff_def) |
|
7649 apply (rule continuous_on_cases) |
|
7650 using homeomorphism_cont1 [OF hom] |
|
7651 apply (auto simp: affine_closed \<open>affine T\<close> continuous_on_id fid) |
|
7652 done |
|
7653 then show "continuous_on S ff" |
|
7654 apply (rule continuous_on_subset) |
|
7655 using ST by auto |
|
7656 have "continuous_on ((cball a r \<inter> T) \<union> (T - ball a r)) gg" |
|
7657 apply (simp add: gg_def) |
|
7658 apply (rule continuous_on_cases) |
|
7659 using homeomorphism_cont2 [OF hom] |
|
7660 apply (auto simp: affine_closed \<open>affine T\<close> continuous_on_id gid) |
|
7661 done |
|
7662 then show "continuous_on S gg" |
|
7663 apply (rule continuous_on_subset) |
|
7664 using ST by auto |
|
7665 show "ff ` S \<subseteq> S" |
|
7666 proof (clarsimp simp add: ff_def) |
|
7667 fix x |
|
7668 assume "x \<in> S" and x: "dist a x < r" and "x \<in> T" |
|
7669 then have "f x \<in> cball a r \<inter> T" |
|
7670 using homeomorphism_image1 [OF hom] by force |
|
7671 then show "f x \<in> S" |
|
7672 using ST(1) \<open>x \<in> T\<close> gid hom homeomorphism_def x by fastforce |
|
7673 qed |
|
7674 show "gg ` S \<subseteq> S" |
|
7675 proof (clarsimp simp add: gg_def) |
|
7676 fix x |
|
7677 assume "x \<in> S" and x: "dist a x < r" and "x \<in> T" |
|
7678 then have "g x \<in> cball a r \<inter> T" |
|
7679 using homeomorphism_image2 [OF hom] by force |
|
7680 then have "g x \<in> ball a r" |
|
7681 using homeomorphism_apply2 [OF hom] |
|
7682 by (metis Diff_Diff_Int Diff_iff \<open>x \<in> T\<close> cball_def fid le_less mem_Collect_eq mem_ball mem_sphere x) |
|
7683 then show "g x \<in> S" |
|
7684 using ST(1) \<open>g x \<in> cball a r \<inter> T\<close> by force |
|
7685 qed |
|
7686 show "\<And>x. x \<in> S \<Longrightarrow> gg (ff x) = x" |
|
7687 apply (simp add: ff_def gg_def) |
|
7688 using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] |
|
7689 apply auto |
|
7690 apply (metis Int_iff homeomorphism_apply1 [OF hom] fid image_eqI less_eq_real_def mem_cball mem_sphere) |
|
7691 done |
|
7692 show "\<And>x. x \<in> S \<Longrightarrow> ff (gg x) = x" |
|
7693 apply (simp add: ff_def gg_def) |
|
7694 using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] |
|
7695 apply auto |
|
7696 apply (metis Int_iff fid image_eqI less_eq_real_def mem_cball mem_sphere) |
|
7697 done |
|
7698 qed |
|
7699 show "ff u = v" |
|
7700 using u by (auto simp: ff_def \<open>f u = v\<close>) |
|
7701 show "{x. \<not> (ff x = x \<and> gg x = x)} \<subseteq> ball a r \<inter> T" |
|
7702 by (auto simp: ff_def gg_def) |
|
7703 qed |
|
7704 qed |
|
7705 |
|
7706 |
|
7707 proposition%unimportant homeomorphism_moving_point: |
|
7708 fixes a :: "'a::euclidean_space" |
|
7709 assumes ope: "openin (subtopology euclidean (affine hull S)) S" |
|
7710 and "S \<subseteq> T" |
|
7711 and TS: "T \<subseteq> affine hull S" |
|
7712 and S: "connected S" "a \<in> S" "b \<in> S" |
|
7713 obtains f g where "homeomorphism T T f g" "f a = b" |
|
7714 "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" |
|
7715 "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
7716 proof - |
|
7717 have 1: "\<exists>h k. homeomorphism T T h k \<and> h (f d) = d \<and> |
|
7718 {x. \<not> (h x = x \<and> k x = x)} \<subseteq> S \<and> bounded {x. \<not> (h x = x \<and> k x = x)}" |
|
7719 if "d \<in> S" "f d \<in> S" and homfg: "homeomorphism T T f g" |
|
7720 and S: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" |
|
7721 and bo: "bounded {x. \<not> (f x = x \<and> g x = x)}" for d f g |
|
7722 proof (intro exI conjI) |
|
7723 show homgf: "homeomorphism T T g f" |
|
7724 by (metis homeomorphism_symD homfg) |
|
7725 then show "g (f d) = d" |
|
7726 by (meson \<open>S \<subseteq> T\<close> homeomorphism_def subsetD \<open>d \<in> S\<close>) |
|
7727 show "{x. \<not> (g x = x \<and> f x = x)} \<subseteq> S" |
|
7728 using S by blast |
|
7729 show "bounded {x. \<not> (g x = x \<and> f x = x)}" |
|
7730 using bo by (simp add: conj_commute) |
|
7731 qed |
|
7732 have 2: "\<exists>f g. homeomorphism T T f g \<and> f x = f2 (f1 x) \<and> |
|
7733 {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
7734 if "x \<in> S" "f1 x \<in> S" "f2 (f1 x) \<in> S" |
|
7735 and hom: "homeomorphism T T f1 g1" "homeomorphism T T f2 g2" |
|
7736 and sub: "{x. \<not> (f1 x = x \<and> g1 x = x)} \<subseteq> S" "{x. \<not> (f2 x = x \<and> g2 x = x)} \<subseteq> S" |
|
7737 and bo: "bounded {x. \<not> (f1 x = x \<and> g1 x = x)}" "bounded {x. \<not> (f2 x = x \<and> g2 x = x)}" |
|
7738 for x f1 f2 g1 g2 |
|
7739 proof (intro exI conjI) |
|
7740 show homgf: "homeomorphism T T (f2 \<circ> f1) (g1 \<circ> g2)" |
|
7741 by (metis homeomorphism_compose hom) |
|
7742 then show "(f2 \<circ> f1) x = f2 (f1 x)" |
|
7743 by force |
|
7744 show "{x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)} \<subseteq> S" |
|
7745 using sub by force |
|
7746 have "bounded ({x. \<not>(f1 x = x \<and> g1 x = x)} \<union> {x. \<not>(f2 x = x \<and> g2 x = x)})" |
|
7747 using bo by simp |
|
7748 then show "bounded {x. \<not> ((f2 \<circ> f1) x = x \<and> (g1 \<circ> g2) x = x)}" |
|
7749 by (rule bounded_subset) auto |
|
7750 qed |
|
7751 have 3: "\<exists>U. openin (subtopology euclidean S) U \<and> |
|
7752 d \<in> U \<and> |
|
7753 (\<forall>x\<in>U. |
|
7754 \<exists>f g. homeomorphism T T f g \<and> f d = x \<and> |
|
7755 {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> |
|
7756 bounded {x. \<not> (f x = x \<and> g x = x)})" |
|
7757 if "d \<in> S" for d |
|
7758 proof - |
|
7759 obtain r where "r > 0" and r: "ball d r \<inter> affine hull S \<subseteq> S" |
|
7760 by (metis \<open>d \<in> S\<close> ope openin_contains_ball) |
|
7761 have *: "\<exists>f g. homeomorphism T T f g \<and> f d = e \<and> |
|
7762 {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> |
|
7763 bounded {x. \<not> (f x = x \<and> g x = x)}" if "e \<in> S" "e \<in> ball d r" for e |
|
7764 apply (rule homeomorphism_moving_point_3 [of "affine hull S" d r T d e]) |
|
7765 using r \<open>S \<subseteq> T\<close> TS that |
|
7766 apply (auto simp: \<open>d \<in> S\<close> \<open>0 < r\<close> hull_inc) |
|
7767 using bounded_subset by blast |
|
7768 show ?thesis |
|
7769 apply (rule_tac x="S \<inter> ball d r" in exI) |
|
7770 apply (intro conjI) |
|
7771 apply (simp add: openin_open_Int) |
|
7772 apply (simp add: \<open>0 < r\<close> that) |
|
7773 apply (blast intro: *) |
|
7774 done |
|
7775 qed |
|
7776 have "\<exists>f g. homeomorphism T T f g \<and> f a = b \<and> |
|
7777 {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
7778 apply (rule connected_equivalence_relation [OF S], safe) |
|
7779 apply (blast intro: 1 2 3)+ |
|
7780 done |
|
7781 then show ?thesis |
|
7782 using that by auto |
|
7783 qed |
|
7784 |
|
7785 |
|
7786 lemma homeomorphism_moving_points_exists_gen: |
|
7787 assumes K: "finite K" "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" |
|
7788 "pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K" |
|
7789 and "2 \<le> aff_dim S" |
|
7790 and ope: "openin (subtopology euclidean (affine hull S)) S" |
|
7791 and "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" |
|
7792 shows "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and> |
|
7793 {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
7794 using assms |
|
7795 proof (induction K) |
|
7796 case empty |
|
7797 then show ?case |
|
7798 by (force simp: homeomorphism_ident) |
|
7799 next |
|
7800 case (insert i K) |
|
7801 then have xney: "\<And>j. \<lbrakk>j \<in> K; j \<noteq> i\<rbrakk> \<Longrightarrow> x i \<noteq> x j \<and> y i \<noteq> y j" |
|
7802 and pw: "pairwise (\<lambda>i j. x i \<noteq> x j \<and> y i \<noteq> y j) K" |
|
7803 and "x i \<in> S" "y i \<in> S" |
|
7804 and xyS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" |
|
7805 by (simp_all add: pairwise_insert) |
|
7806 obtain f g where homfg: "homeomorphism T T f g" and feq: "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i" |
|
7807 and fg_sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" |
|
7808 and bo_fg: "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
7809 using insert.IH [OF xyS pw] insert.prems by (blast intro: that) |
|
7810 then have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(x i) = y i) \<and> |
|
7811 {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
7812 using insert by blast |
|
7813 have aff_eq: "affine hull (S - y ` K) = affine hull S" |
|
7814 apply (rule affine_hull_Diff) |
|
7815 apply (auto simp: insert) |
|
7816 using \<open>y i \<in> S\<close> insert.hyps(2) xney xyS by fastforce |
|
7817 have f_in_S: "f x \<in> S" if "x \<in> S" for x |
|
7818 using homfg fg_sub homeomorphism_apply1 \<open>S \<subseteq> T\<close> |
|
7819 proof - |
|
7820 have "(f (f x) \<noteq> f x \<or> g (f x) \<noteq> f x) \<or> f x \<in> S" |
|
7821 by (metis \<open>S \<subseteq> T\<close> homfg subsetD homeomorphism_apply1 that) |
|
7822 then show ?thesis |
|
7823 using fg_sub by force |
|
7824 qed |
|
7825 obtain h k where homhk: "homeomorphism T T h k" and heq: "h (f (x i)) = y i" |
|
7826 and hk_sub: "{x. \<not> (h x = x \<and> k x = x)} \<subseteq> S - y ` K" |
|
7827 and bo_hk: "bounded {x. \<not> (h x = x \<and> k x = x)}" |
|
7828 proof (rule homeomorphism_moving_point [of "S - y`K" T "f(x i)" "y i"]) |
|
7829 show "openin (subtopology euclidean (affine hull (S - y ` K))) (S - y ` K)" |
|
7830 by (simp add: aff_eq openin_diff finite_imp_closedin image_subset_iff hull_inc insert xyS) |
|
7831 show "S - y ` K \<subseteq> T" |
|
7832 using \<open>S \<subseteq> T\<close> by auto |
|
7833 show "T \<subseteq> affine hull (S - y ` K)" |
|
7834 using insert by (simp add: aff_eq) |
|
7835 show "connected (S - y ` K)" |
|
7836 proof (rule connected_openin_diff_countable [OF \<open>connected S\<close> ope]) |
|
7837 show "\<not> collinear S" |
|
7838 using collinear_aff_dim \<open>2 \<le> aff_dim S\<close> by force |
|
7839 show "countable (y ` K)" |
|
7840 using countable_finite insert.hyps(1) by blast |
|
7841 qed |
|
7842 show "f (x i) \<in> S - y ` K" |
|
7843 apply (auto simp: f_in_S \<open>x i \<in> S\<close>) |
|
7844 by (metis feq homfg \<open>x i \<in> S\<close> homeomorphism_def \<open>S \<subseteq> T\<close> \<open>i \<notin> K\<close> subsetCE xney xyS) |
|
7845 show "y i \<in> S - y ` K" |
|
7846 using insert.hyps xney by (auto simp: \<open>y i \<in> S\<close>) |
|
7847 qed blast |
|
7848 show ?case |
|
7849 proof (intro exI conjI) |
|
7850 show "homeomorphism T T (h \<circ> f) (g \<circ> k)" |
|
7851 using homfg homhk homeomorphism_compose by blast |
|
7852 show "\<forall>i \<in> insert i K. (h \<circ> f) (x i) = y i" |
|
7853 using feq hk_sub by (auto simp: heq) |
|
7854 show "{x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)} \<subseteq> S" |
|
7855 using fg_sub hk_sub by force |
|
7856 have "bounded ({x. \<not>(f x = x \<and> g x = x)} \<union> {x. \<not>(h x = x \<and> k x = x)})" |
|
7857 using bo_fg bo_hk bounded_Un by blast |
|
7858 then show "bounded {x. \<not> ((h \<circ> f) x = x \<and> (g \<circ> k) x = x)}" |
|
7859 by (rule bounded_subset) auto |
|
7860 qed |
|
7861 qed |
|
7862 |
|
7863 proposition%unimportant homeomorphism_moving_points_exists: |
|
7864 fixes S :: "'a::euclidean_space set" |
|
7865 assumes 2: "2 \<le> DIM('a)" "open S" "connected S" "S \<subseteq> T" "finite K" |
|
7866 and KS: "\<And>i. i \<in> K \<Longrightarrow> x i \<in> S \<and> y i \<in> S" |
|
7867 and pw: "pairwise (\<lambda>i j. (x i \<noteq> x j) \<and> (y i \<noteq> y j)) K" |
|
7868 and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" |
|
7869 obtains f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f(x i) = y i" |
|
7870 "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. (\<not> (f x = x \<and> g x = x))}" |
|
7871 proof (cases "S = {}") |
|
7872 case True |
|
7873 then show ?thesis |
|
7874 using KS homeomorphism_ident that by fastforce |
|
7875 next |
|
7876 case False |
|
7877 then have affS: "affine hull S = UNIV" |
|
7878 by (simp add: affine_hull_open \<open>open S\<close>) |
|
7879 then have ope: "openin (subtopology euclidean (affine hull S)) S" |
|
7880 using \<open>open S\<close> open_openin by auto |
|
7881 have "2 \<le> DIM('a)" by (rule 2) |
|
7882 also have "\<dots> = aff_dim (UNIV :: 'a set)" |
|
7883 by simp |
|
7884 also have "\<dots> \<le> aff_dim S" |
|
7885 by (metis aff_dim_UNIV aff_dim_affine_hull aff_dim_le_DIM affS) |
|
7886 finally have "2 \<le> aff_dim S" |
|
7887 by linarith |
|
7888 then show ?thesis |
|
7889 using homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> KS pw _ ope S] that by fastforce |
|
7890 qed |
|
7891 |
|
7892 subsubsection%unimportant\<open>The theorem \<open>homeomorphism_grouping_points_exists\<close>\<close> |
|
7893 |
|
7894 lemma homeomorphism_grouping_point_1: |
|
7895 fixes a::real and c::real |
|
7896 assumes "a < b" "c < d" |
|
7897 obtains f g where "homeomorphism (cbox a b) (cbox c d) f g" "f a = c" "f b = d" |
|
7898 proof - |
|
7899 define f where "f \<equiv> \<lambda>x. ((d - c) / (b - a)) * x + (c - a * ((d - c) / (b - a)))" |
|
7900 have "\<exists>g. homeomorphism (cbox a b) (cbox c d) f g" |
|
7901 proof (rule homeomorphism_compact) |
|
7902 show "continuous_on (cbox a b) f" |
|
7903 apply (simp add: f_def) |
|
7904 apply (intro continuous_intros) |
|
7905 using assms by auto |
|
7906 have "f ` {a..b} = {c..d}" |
|
7907 unfolding f_def image_affinity_atLeastAtMost |
|
7908 using assms sum_sqs_eq by (auto simp: divide_simps algebra_simps) |
|
7909 then show "f ` cbox a b = cbox c d" |
|
7910 by auto |
|
7911 show "inj_on f (cbox a b)" |
|
7912 unfolding f_def inj_on_def using assms by auto |
|
7913 qed auto |
|
7914 then obtain g where "homeomorphism (cbox a b) (cbox c d) f g" .. |
|
7915 then show ?thesis |
|
7916 proof |
|
7917 show "f a = c" |
|
7918 by (simp add: f_def) |
|
7919 show "f b = d" |
|
7920 using assms sum_sqs_eq [of a b] by (auto simp: f_def divide_simps algebra_simps) |
|
7921 qed |
|
7922 qed |
|
7923 |
|
7924 lemma homeomorphism_grouping_point_2: |
|
7925 fixes a::real and w::real |
|
7926 assumes hom_ab: "homeomorphism (cbox a b) (cbox u v) f1 g1" |
|
7927 and hom_bc: "homeomorphism (cbox b c) (cbox v w) f2 g2" |
|
7928 and "b \<in> cbox a c" "v \<in> cbox u w" |
|
7929 and eq: "f1 a = u" "f1 b = v" "f2 b = v" "f2 c = w" |
|
7930 obtains f g where "homeomorphism (cbox a c) (cbox u w) f g" "f a = u" "f c = w" |
|
7931 "\<And>x. x \<in> cbox a b \<Longrightarrow> f x = f1 x" "\<And>x. x \<in> cbox b c \<Longrightarrow> f x = f2 x" |
|
7932 proof - |
|
7933 have le: "a \<le> b" "b \<le> c" "u \<le> v" "v \<le> w" |
|
7934 using assms by simp_all |
|
7935 then have ac: "cbox a c = cbox a b \<union> cbox b c" and uw: "cbox u w = cbox u v \<union> cbox v w" |
|
7936 by auto |
|
7937 define f where "f \<equiv> \<lambda>x. if x \<le> b then f1 x else f2 x" |
|
7938 have "\<exists>g. homeomorphism (cbox a c) (cbox u w) f g" |
|
7939 proof (rule homeomorphism_compact) |
|
7940 have cf1: "continuous_on (cbox a b) f1" |
|
7941 using hom_ab homeomorphism_cont1 by blast |
|
7942 have cf2: "continuous_on (cbox b c) f2" |
|
7943 using hom_bc homeomorphism_cont1 by blast |
|
7944 show "continuous_on (cbox a c) f" |
|
7945 apply (simp add: f_def) |
|
7946 apply (rule continuous_on_cases_le [OF continuous_on_subset [OF cf1] continuous_on_subset [OF cf2]]) |
|
7947 using le eq apply (force simp: continuous_on_id)+ |
|
7948 done |
|
7949 have "f ` cbox a b = f1 ` cbox a b" "f ` cbox b c = f2 ` cbox b c" |
|
7950 unfolding f_def using eq by force+ |
|
7951 then show "f ` cbox a c = cbox u w" |
|
7952 apply (simp only: ac uw image_Un) |
|
7953 by (metis hom_ab hom_bc homeomorphism_def) |
|
7954 have neq12: "f1 x \<noteq> f2 y" if x: "a \<le> x" "x \<le> b" and y: "b < y" "y \<le> c" for x y |
|
7955 proof - |
|
7956 have "f1 x \<in> cbox u v" |
|
7957 by (metis hom_ab homeomorphism_def image_eqI mem_box_real(2) x) |
|
7958 moreover have "f2 y \<in> cbox v w" |
|
7959 by (metis (full_types) hom_bc homeomorphism_def image_subset_iff mem_box_real(2) not_le not_less_iff_gr_or_eq order_refl y) |
|
7960 moreover have "f2 y \<noteq> f2 b" |
|
7961 by (metis cancel_comm_monoid_add_class.diff_cancel diff_gt_0_iff_gt hom_bc homeomorphism_def le(2) less_imp_le less_numeral_extra(3) mem_box_real(2) order_refl y) |
|
7962 ultimately show ?thesis |
|
7963 using le eq by simp |
|
7964 qed |
|
7965 have "inj_on f1 (cbox a b)" |
|
7966 by (metis (full_types) hom_ab homeomorphism_def inj_onI) |
|
7967 moreover have "inj_on f2 (cbox b c)" |
|
7968 by (metis (full_types) hom_bc homeomorphism_def inj_onI) |
|
7969 ultimately show "inj_on f (cbox a c)" |
|
7970 apply (simp (no_asm) add: inj_on_def) |
|
7971 apply (simp add: f_def inj_on_eq_iff) |
|
7972 using neq12 apply force |
|
7973 done |
|
7974 qed auto |
|
7975 then obtain g where "homeomorphism (cbox a c) (cbox u w) f g" .. |
|
7976 then show ?thesis |
|
7977 apply (rule that) |
|
7978 using eq le by (auto simp: f_def) |
|
7979 qed |
|
7980 |
|
7981 lemma homeomorphism_grouping_point_3: |
|
7982 fixes a::real |
|
7983 assumes cbox_sub: "cbox c d \<subseteq> box a b" "cbox u v \<subseteq> box a b" |
|
7984 and box_ne: "box c d \<noteq> {}" "box u v \<noteq> {}" |
|
7985 obtains f g where "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" |
|
7986 "\<And>x. x \<in> cbox c d \<Longrightarrow> f x \<in> cbox u v" |
|
7987 proof - |
|
7988 have less: "a < c" "a < u" "d < b" "v < b" "c < d" "u < v" "cbox c d \<noteq> {}" |
|
7989 using assms |
|
7990 by (simp_all add: cbox_sub subset_eq) |
|
7991 obtain f1 g1 where 1: "homeomorphism (cbox a c) (cbox a u) f1 g1" |
|
7992 and f1_eq: "f1 a = a" "f1 c = u" |
|
7993 using homeomorphism_grouping_point_1 [OF \<open>a < c\<close> \<open>a < u\<close>] . |
|
7994 obtain f2 g2 where 2: "homeomorphism (cbox c d) (cbox u v) f2 g2" |
|
7995 and f2_eq: "f2 c = u" "f2 d = v" |
|
7996 using homeomorphism_grouping_point_1 [OF \<open>c < d\<close> \<open>u < v\<close>] . |
|
7997 obtain f3 g3 where 3: "homeomorphism (cbox d b) (cbox v b) f3 g3" |
|
7998 and f3_eq: "f3 d = v" "f3 b = b" |
|
7999 using homeomorphism_grouping_point_1 [OF \<open>d < b\<close> \<open>v < b\<close>] . |
|
8000 obtain f4 g4 where 4: "homeomorphism (cbox a d) (cbox a v) f4 g4" and "f4 a = a" "f4 d = v" |
|
8001 and f4_eq: "\<And>x. x \<in> cbox a c \<Longrightarrow> f4 x = f1 x" "\<And>x. x \<in> cbox c d \<Longrightarrow> f4 x = f2 x" |
|
8002 using homeomorphism_grouping_point_2 [OF 1 2] less by (auto simp: f1_eq f2_eq) |
|
8003 obtain f g where fg: "homeomorphism (cbox a b) (cbox a b) f g" "f a = a" "f b = b" |
|
8004 and f_eq: "\<And>x. x \<in> cbox a d \<Longrightarrow> f x = f4 x" "\<And>x. x \<in> cbox d b \<Longrightarrow> f x = f3 x" |
|
8005 using homeomorphism_grouping_point_2 [OF 4 3] less by (auto simp: f4_eq f3_eq f2_eq f1_eq) |
|
8006 show ?thesis |
|
8007 apply (rule that [OF fg]) |
|
8008 using f4_eq f_eq homeomorphism_image1 [OF 2] |
|
8009 apply simp |
|
8010 by (metis atLeastAtMost_iff box_real(1) box_real(2) cbox_sub(1) greaterThanLessThan_iff imageI less_eq_real_def subset_eq) |
|
8011 qed |
|
8012 |
|
8013 |
|
8014 lemma homeomorphism_grouping_point_4: |
|
8015 fixes T :: "real set" |
|
8016 assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T" |
|
8017 obtains f g where "homeomorphism T T f g" |
|
8018 "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S" |
|
8019 "bounded {x. (\<not> (f x = x \<and> g x = x))}" |
|
8020 proof - |
|
8021 obtain c d where "box c d \<noteq> {}" "cbox c d \<subseteq> U" |
|
8022 proof - |
|
8023 obtain u where "u \<in> U" |
|
8024 using \<open>U \<noteq> {}\<close> by blast |
|
8025 then obtain e where "e > 0" "cball u e \<subseteq> U" |
|
8026 using \<open>open U\<close> open_contains_cball by blast |
|
8027 then show ?thesis |
|
8028 by (rule_tac c=u and d="u+e" in that) (auto simp: dist_norm subset_iff) |
|
8029 qed |
|
8030 have "compact K" |
|
8031 by (simp add: \<open>finite K\<close> finite_imp_compact) |
|
8032 obtain a b where "box a b \<noteq> {}" "K \<subseteq> cbox a b" "cbox a b \<subseteq> S" |
|
8033 proof (cases "K = {}") |
|
8034 case True then show ?thesis |
|
8035 using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> that by blast |
|
8036 next |
|
8037 case False |
|
8038 then obtain a b where "a \<in> K" "b \<in> K" |
|
8039 and a: "\<And>x. x \<in> K \<Longrightarrow> a \<le> x" and b: "\<And>x. x \<in> K \<Longrightarrow> x \<le> b" |
|
8040 using compact_attains_inf compact_attains_sup by (metis \<open>compact K\<close>)+ |
|
8041 obtain e where "e > 0" "cball b e \<subseteq> S" |
|
8042 using \<open>open S\<close> open_contains_cball |
|
8043 by (metis \<open>b \<in> K\<close> \<open>K \<subseteq> S\<close> subsetD) |
|
8044 show ?thesis |
|
8045 proof |
|
8046 show "box a (b + e) \<noteq> {}" |
|
8047 using \<open>0 < e\<close> \<open>b \<in> K\<close> a by force |
|
8048 show "K \<subseteq> cbox a (b + e)" |
|
8049 using \<open>0 < e\<close> a b by fastforce |
|
8050 have "a \<in> S" |
|
8051 using \<open>a \<in> K\<close> assms(6) by blast |
|
8052 have "b + e \<in> S" |
|
8053 using \<open>0 < e\<close> \<open>cball b e \<subseteq> S\<close> by (force simp: dist_norm) |
|
8054 show "cbox a (b + e) \<subseteq> S" |
|
8055 using \<open>a \<in> S\<close> \<open>b + e \<in> S\<close> \<open>connected S\<close> connected_contains_Icc by auto |
|
8056 qed |
|
8057 qed |
|
8058 obtain w z where "cbox w z \<subseteq> S" and sub_wz: "cbox a b \<union> cbox c d \<subseteq> box w z" |
|
8059 proof - |
|
8060 have "a \<in> S" "b \<in> S" |
|
8061 using \<open>box a b \<noteq> {}\<close> \<open>cbox a b \<subseteq> S\<close> by auto |
|
8062 moreover have "c \<in> S" "d \<in> S" |
|
8063 using \<open>box c d \<noteq> {}\<close> \<open>cbox c d \<subseteq> U\<close> \<open>U \<subseteq> S\<close> by force+ |
|
8064 ultimately have "min a c \<in> S" "max b d \<in> S" |
|
8065 by linarith+ |
|
8066 then obtain e1 e2 where "e1 > 0" "cball (min a c) e1 \<subseteq> S" "e2 > 0" "cball (max b d) e2 \<subseteq> S" |
|
8067 using \<open>open S\<close> open_contains_cball by metis |
|
8068 then have *: "min a c - e1 \<in> S" "max b d + e2 \<in> S" |
|
8069 by (auto simp: dist_norm) |
|
8070 show ?thesis |
|
8071 proof |
|
8072 show "cbox (min a c - e1) (max b d+ e2) \<subseteq> S" |
|
8073 using * \<open>connected S\<close> connected_contains_Icc by auto |
|
8074 show "cbox a b \<union> cbox c d \<subseteq> box (min a c - e1) (max b d + e2)" |
|
8075 using \<open>0 < e1\<close> \<open>0 < e2\<close> by auto |
|
8076 qed |
|
8077 qed |
|
8078 then |
|
8079 obtain f g where hom: "homeomorphism (cbox w z) (cbox w z) f g" |
|
8080 and "f w = w" "f z = z" |
|
8081 and fin: "\<And>x. x \<in> cbox a b \<Longrightarrow> f x \<in> cbox c d" |
|
8082 using homeomorphism_grouping_point_3 [of a b w z c d] |
|
8083 using \<open>box a b \<noteq> {}\<close> \<open>box c d \<noteq> {}\<close> by blast |
|
8084 have contfg: "continuous_on (cbox w z) f" "continuous_on (cbox w z) g" |
|
8085 using hom homeomorphism_def by blast+ |
|
8086 define f' where "f' \<equiv> \<lambda>x. if x \<in> cbox w z then f x else x" |
|
8087 define g' where "g' \<equiv> \<lambda>x. if x \<in> cbox w z then g x else x" |
|
8088 show ?thesis |
|
8089 proof |
|
8090 have T: "cbox w z \<union> (T - box w z) = T" |
|
8091 using \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> by auto |
|
8092 show "homeomorphism T T f' g'" |
|
8093 proof |
|
8094 have clo: "closedin (subtopology euclidean (cbox w z \<union> (T - box w z))) (T - box w z)" |
|
8095 by (metis Diff_Diff_Int Diff_subset T closedin_def open_box openin_open_Int topspace_euclidean_subtopology) |
|
8096 have "continuous_on (cbox w z \<union> (T - box w z)) f'" "continuous_on (cbox w z \<union> (T - box w z)) g'" |
|
8097 unfolding f'_def g'_def |
|
8098 apply (safe intro!: continuous_on_cases_local contfg continuous_on_id clo) |
|
8099 apply (simp_all add: closed_subset) |
|
8100 using \<open>f w = w\<close> \<open>f z = z\<close> apply force |
|
8101 by (metis \<open>f w = w\<close> \<open>f z = z\<close> hom homeomorphism_def less_eq_real_def mem_box_real(2)) |
|
8102 then show "continuous_on T f'" "continuous_on T g'" |
|
8103 by (simp_all only: T) |
|
8104 show "f' ` T \<subseteq> T" |
|
8105 unfolding f'_def |
|
8106 by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2)) |
|
8107 show "g' ` T \<subseteq> T" |
|
8108 unfolding g'_def |
|
8109 by clarsimp (metis \<open>cbox w z \<subseteq> S\<close> \<open>S \<subseteq> T\<close> subsetD hom homeomorphism_def imageI mem_box_real(2)) |
|
8110 show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x" |
|
8111 unfolding f'_def g'_def |
|
8112 using homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] by fastforce |
|
8113 show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y" |
|
8114 unfolding f'_def g'_def |
|
8115 using homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] by fastforce |
|
8116 qed |
|
8117 show "\<And>x. x \<in> K \<Longrightarrow> f' x \<in> U" |
|
8118 using fin sub_wz \<open>K \<subseteq> cbox a b\<close> \<open>cbox c d \<subseteq> U\<close> by (force simp: f'_def) |
|
8119 show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S" |
|
8120 using \<open>cbox w z \<subseteq> S\<close> by (auto simp: f'_def g'_def) |
|
8121 show "bounded {x. \<not> (f' x = x \<and> g' x = x)}" |
|
8122 apply (rule bounded_subset [of "cbox w z"]) |
|
8123 using bounded_cbox apply blast |
|
8124 apply (auto simp: f'_def g'_def) |
|
8125 done |
|
8126 qed |
|
8127 qed |
|
8128 |
|
8129 proposition%unimportant homeomorphism_grouping_points_exists: |
|
8130 fixes S :: "'a::euclidean_space set" |
|
8131 assumes "open U" "open S" "connected S" "U \<noteq> {}" "finite K" "K \<subseteq> S" "U \<subseteq> S" "S \<subseteq> T" |
|
8132 obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S" |
|
8133 "bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" |
|
8134 proof (cases "2 \<le> DIM('a)") |
|
8135 case True |
|
8136 have TS: "T \<subseteq> affine hull S" |
|
8137 using affine_hull_open assms by blast |
|
8138 have "infinite U" |
|
8139 using \<open>open U\<close> \<open>U \<noteq> {}\<close> finite_imp_not_open by blast |
|
8140 then obtain P where "P \<subseteq> U" "finite P" "card K = card P" |
|
8141 using infinite_arbitrarily_large by metis |
|
8142 then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P" |
|
8143 using \<open>finite K\<close> finite_same_card_bij by blast |
|
8144 obtain f g where "homeomorphism T T f g" "\<And>i. i \<in> K \<Longrightarrow> f (id i) = \<gamma> i" "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> S" "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
8145 proof (rule homeomorphism_moving_points_exists [OF True \<open>open S\<close> \<open>connected S\<close> \<open>S \<subseteq> T\<close> \<open>finite K\<close>]) |
|
8146 show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S" |
|
8147 using \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> \<open>U \<subseteq> S\<close> bij_betwE by blast |
|
8148 show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K" |
|
8149 using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def) |
|
8150 qed (use affine_hull_open assms that in auto) |
|
8151 then show ?thesis |
|
8152 using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp add: intro!: that) |
|
8153 next |
|
8154 case False |
|
8155 with DIM_positive have "DIM('a) = 1" |
|
8156 by (simp add: dual_order.antisym) |
|
8157 then obtain h::"'a \<Rightarrow>real" and j |
|
8158 where "linear h" "linear j" |
|
8159 and noh: "\<And>x. norm(h x) = norm x" and noj: "\<And>y. norm(j y) = norm y" |
|
8160 and hj: "\<And>x. j(h x) = x" "\<And>y. h(j y) = y" |
|
8161 and ranh: "surj h" |
|
8162 using isomorphisms_UNIV_UNIV |
|
8163 by (metis (mono_tags, hide_lams) DIM_real UNIV_eq_I range_eqI) |
|
8164 obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" |
|
8165 and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U" |
|
8166 and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S" |
|
8167 and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
8168 apply (rule homeomorphism_grouping_point_4 [of "h ` U" "h ` S" "h ` K" "h ` T"]) |
|
8169 by (simp_all add: assms image_mono \<open>linear h\<close> open_surjective_linear_image connected_linear_image ranh) |
|
8170 have jf: "j (f (h x)) = x \<longleftrightarrow> f (h x) = h x" for x |
|
8171 by (metis hj) |
|
8172 have jg: "j (g (h x)) = x \<longleftrightarrow> g (h x) = h x" for x |
|
8173 by (metis hj) |
|
8174 have cont_hj: "continuous_on X h" "continuous_on Y j" for X Y |
|
8175 by (simp_all add: \<open>linear h\<close> \<open>linear j\<close> linear_linear linear_continuous_on) |
|
8176 show ?thesis |
|
8177 proof |
|
8178 show "homeomorphism T T (j \<circ> f \<circ> h) (j \<circ> g \<circ> h)" |
|
8179 proof |
|
8180 show "continuous_on T (j \<circ> f \<circ> h)" "continuous_on T (j \<circ> g \<circ> h)" |
|
8181 using hom homeomorphism_def |
|
8182 by (blast intro: continuous_on_compose cont_hj)+ |
|
8183 show "(j \<circ> f \<circ> h) ` T \<subseteq> T" "(j \<circ> g \<circ> h) ` T \<subseteq> T" |
|
8184 by auto (metis (mono_tags, hide_lams) hj(1) hom homeomorphism_def imageE imageI)+ |
|
8185 show "\<And>x. x \<in> T \<Longrightarrow> (j \<circ> g \<circ> h) ((j \<circ> f \<circ> h) x) = x" |
|
8186 using hj hom homeomorphism_apply1 by fastforce |
|
8187 show "\<And>y. y \<in> T \<Longrightarrow> (j \<circ> f \<circ> h) ((j \<circ> g \<circ> h) y) = y" |
|
8188 using hj hom homeomorphism_apply2 by fastforce |
|
8189 qed |
|
8190 show "{x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} \<subseteq> S" |
|
8191 apply (clarsimp simp: jf jg hj) |
|
8192 using sub hj |
|
8193 apply (drule_tac c="h x" in subsetD, force) |
|
8194 by (metis imageE) |
|
8195 have "bounded (j ` {x. (\<not> (f x = x \<and> g x = x))})" |
|
8196 by (rule bounded_linear_image [OF bou]) (use \<open>linear j\<close> linear_conv_bounded_linear in auto) |
|
8197 moreover |
|
8198 have *: "{x. \<not>((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)} = j ` {x. (\<not> (f x = x \<and> g x = x))}" |
|
8199 using hj by (auto simp: jf jg image_iff, metis+) |
|
8200 ultimately show "bounded {x. \<not> ((j \<circ> f \<circ> h) x = x \<and> (j \<circ> g \<circ> h) x = x)}" |
|
8201 by metis |
|
8202 show "\<And>x. x \<in> K \<Longrightarrow> (j \<circ> f \<circ> h) x \<in> U" |
|
8203 using f hj by fastforce |
|
8204 qed |
|
8205 qed |
|
8206 |
|
8207 |
|
8208 proposition%unimportant homeomorphism_grouping_points_exists_gen: |
|
8209 fixes S :: "'a::euclidean_space set" |
|
8210 assumes opeU: "openin (subtopology euclidean S) U" |
|
8211 and opeS: "openin (subtopology euclidean (affine hull S)) S" |
|
8212 and "U \<noteq> {}" "finite K" "K \<subseteq> S" and S: "S \<subseteq> T" "T \<subseteq> affine hull S" "connected S" |
|
8213 obtains f g where "homeomorphism T T f g" "{x. (\<not> (f x = x \<and> g x = x))} \<subseteq> S" |
|
8214 "bounded {x. (\<not> (f x = x \<and> g x = x))}" "\<And>x. x \<in> K \<Longrightarrow> f x \<in> U" |
|
8215 proof (cases "2 \<le> aff_dim S") |
|
8216 case True |
|
8217 have opeU': "openin (subtopology euclidean (affine hull S)) U" |
|
8218 using opeS opeU openin_trans by blast |
|
8219 obtain u where "u \<in> U" "u \<in> S" |
|
8220 using \<open>U \<noteq> {}\<close> opeU openin_imp_subset by fastforce+ |
|
8221 have "infinite U" |
|
8222 apply (rule infinite_openin [OF opeU \<open>u \<in> U\<close>]) |
|
8223 apply (rule connected_imp_perfect_aff_dim [OF \<open>connected S\<close> _ \<open>u \<in> S\<close>]) |
|
8224 using True apply simp |
|
8225 done |
|
8226 then obtain P where "P \<subseteq> U" "finite P" "card K = card P" |
|
8227 using infinite_arbitrarily_large by metis |
|
8228 then obtain \<gamma> where \<gamma>: "bij_betw \<gamma> K P" |
|
8229 using \<open>finite K\<close> finite_same_card_bij by blast |
|
8230 have "\<exists>f g. homeomorphism T T f g \<and> (\<forall>i \<in> K. f(id i) = \<gamma> i) \<and> |
|
8231 {x. \<not> (f x = x \<and> g x = x)} \<subseteq> S \<and> bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
8232 proof (rule homeomorphism_moving_points_exists_gen [OF \<open>finite K\<close> _ _ True opeS S]) |
|
8233 show "\<And>i. i \<in> K \<Longrightarrow> id i \<in> S \<and> \<gamma> i \<in> S" |
|
8234 by (metis id_apply opeU openin_contains_cball subsetCE \<open>P \<subseteq> U\<close> \<open>bij_betw \<gamma> K P\<close> \<open>K \<subseteq> S\<close> bij_betwE) |
|
8235 show "pairwise (\<lambda>i j. id i \<noteq> id j \<and> \<gamma> i \<noteq> \<gamma> j) K" |
|
8236 using \<gamma> by (auto simp: pairwise_def bij_betw_def inj_on_def) |
|
8237 qed |
|
8238 then show ?thesis |
|
8239 using \<gamma> \<open>P \<subseteq> U\<close> bij_betwE by (fastforce simp add: intro!: that) |
|
8240 next |
|
8241 case False |
|
8242 with aff_dim_geq [of S] consider "aff_dim S = -1" | "aff_dim S = 0" | "aff_dim S = 1" by linarith |
|
8243 then show ?thesis |
|
8244 proof cases |
|
8245 assume "aff_dim S = -1" |
|
8246 then have "S = {}" |
|
8247 using aff_dim_empty by blast |
|
8248 then have "False" |
|
8249 using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast |
|
8250 then show ?thesis .. |
|
8251 next |
|
8252 assume "aff_dim S = 0" |
|
8253 then obtain a where "S = {a}" |
|
8254 using aff_dim_eq_0 by blast |
|
8255 then have "K \<subseteq> U" |
|
8256 using \<open>U \<noteq> {}\<close> \<open>K \<subseteq> S\<close> openin_imp_subset [OF opeU] by blast |
|
8257 show ?thesis |
|
8258 apply (rule that [of id id]) |
|
8259 using \<open>K \<subseteq> U\<close> by (auto simp: continuous_on_id intro: homeomorphismI) |
|
8260 next |
|
8261 assume "aff_dim S = 1" |
|
8262 then have "affine hull S homeomorphic (UNIV :: real set)" |
|
8263 by (auto simp: homeomorphic_affine_sets) |
|
8264 then obtain h::"'a\<Rightarrow>real" and j where homhj: "homeomorphism (affine hull S) UNIV h j" |
|
8265 using homeomorphic_def by blast |
|
8266 then have h: "\<And>x. x \<in> affine hull S \<Longrightarrow> j(h(x)) = x" and j: "\<And>y. j y \<in> affine hull S \<and> h(j y) = y" |
|
8267 by (auto simp: homeomorphism_def) |
|
8268 have connh: "connected (h ` S)" |
|
8269 by (meson Topological_Spaces.connected_continuous_image \<open>connected S\<close> homeomorphism_cont1 homeomorphism_of_subsets homhj hull_subset top_greatest) |
|
8270 have hUS: "h ` U \<subseteq> h ` S" |
|
8271 by (meson homeomorphism_imp_open_map homeomorphism_of_subsets homhj hull_subset opeS opeU open_UNIV openin_open_eq) |
|
8272 have opn: "openin (subtopology euclidean (affine hull S)) U \<Longrightarrow> open (h ` U)" for U |
|
8273 using homeomorphism_imp_open_map [OF homhj] by simp |
|
8274 have "open (h ` U)" "open (h ` S)" |
|
8275 by (auto intro: opeS opeU openin_trans opn) |
|
8276 then obtain f g where hom: "homeomorphism (h ` T) (h ` T) f g" |
|
8277 and f: "\<And>x. x \<in> h ` K \<Longrightarrow> f x \<in> h ` U" |
|
8278 and sub: "{x. \<not> (f x = x \<and> g x = x)} \<subseteq> h ` S" |
|
8279 and bou: "bounded {x. \<not> (f x = x \<and> g x = x)}" |
|
8280 apply (rule homeomorphism_grouping_points_exists [of "h ` U" "h ` S" "h ` K" "h ` T"]) |
|
8281 using assms by (auto simp: connh hUS) |
|
8282 have jf: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (f (h x)) = x \<longleftrightarrow> f (h x) = h x" |
|
8283 by (metis h j) |
|
8284 have jg: "\<And>x. x \<in> affine hull S \<Longrightarrow> j (g (h x)) = x \<longleftrightarrow> g (h x) = h x" |
|
8285 by (metis h j) |
|
8286 have cont_hj: "continuous_on T h" "continuous_on Y j" for Y |
|
8287 apply (rule continuous_on_subset [OF _ \<open>T \<subseteq> affine hull S\<close>]) |
|
8288 using homeomorphism_def homhj apply blast |
|
8289 by (meson continuous_on_subset homeomorphism_def homhj top_greatest) |
|
8290 define f' where "f' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> f \<circ> h) x else x" |
|
8291 define g' where "g' \<equiv> \<lambda>x. if x \<in> affine hull S then (j \<circ> g \<circ> h) x else x" |
|
8292 show ?thesis |
|
8293 proof |
|
8294 show "homeomorphism T T f' g'" |
|
8295 proof |
|
8296 have "continuous_on T (j \<circ> f \<circ> h)" |
|
8297 apply (intro continuous_on_compose cont_hj) |
|
8298 using hom homeomorphism_def by blast |
|
8299 then show "continuous_on T f'" |
|
8300 apply (rule continuous_on_eq) |
|
8301 using \<open>T \<subseteq> affine hull S\<close> f'_def by auto |
|
8302 have "continuous_on T (j \<circ> g \<circ> h)" |
|
8303 apply (intro continuous_on_compose cont_hj) |
|
8304 using hom homeomorphism_def by blast |
|
8305 then show "continuous_on T g'" |
|
8306 apply (rule continuous_on_eq) |
|
8307 using \<open>T \<subseteq> affine hull S\<close> g'_def by auto |
|
8308 show "f' ` T \<subseteq> T" |
|
8309 proof (clarsimp simp: f'_def) |
|
8310 fix x assume "x \<in> T" |
|
8311 then have "f (h x) \<in> h ` T" |
|
8312 by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) |
|
8313 then show "j (f (h x)) \<in> T" |
|
8314 using \<open>T \<subseteq> affine hull S\<close> h by auto |
|
8315 qed |
|
8316 show "g' ` T \<subseteq> T" |
|
8317 proof (clarsimp simp: g'_def) |
|
8318 fix x assume "x \<in> T" |
|
8319 then have "g (h x) \<in> h ` T" |
|
8320 by (metis (no_types) hom homeomorphism_def image_subset_iff subset_refl) |
|
8321 then show "j (g (h x)) \<in> T" |
|
8322 using \<open>T \<subseteq> affine hull S\<close> h by auto |
|
8323 qed |
|
8324 show "\<And>x. x \<in> T \<Longrightarrow> g' (f' x) = x" |
|
8325 using h j hom homeomorphism_apply1 by (fastforce simp add: f'_def g'_def) |
|
8326 show "\<And>y. y \<in> T \<Longrightarrow> f' (g' y) = y" |
|
8327 using h j hom homeomorphism_apply2 by (fastforce simp add: f'_def g'_def) |
|
8328 qed |
|
8329 next |
|
8330 show "{x. \<not> (f' x = x \<and> g' x = x)} \<subseteq> S" |
|
8331 apply (clarsimp simp: f'_def g'_def jf jg) |
|
8332 apply (rule imageE [OF subsetD [OF sub]], force) |
|
8333 by (metis h hull_inc) |
|
8334 next |
|
8335 have "compact (j ` closure {x. \<not> (f x = x \<and> g x = x)})" |
|
8336 using bou by (auto simp: compact_continuous_image cont_hj) |
|
8337 then have "bounded (j ` {x. \<not> (f x = x \<and> g x = x)})" |
|
8338 by (rule bounded_closure_image [OF compact_imp_bounded]) |
|
8339 moreover |
|
8340 have *: "{x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x} = j ` {x. (\<not> (f x = x \<and> g x = x))}" |
|
8341 using h j by (auto simp: image_iff; metis) |
|
8342 ultimately have "bounded {x \<in> affine hull S. j (f (h x)) \<noteq> x \<or> j (g (h x)) \<noteq> x}" |
|
8343 by metis |
|
8344 then show "bounded {x. \<not> (f' x = x \<and> g' x = x)}" |
|
8345 by (simp add: f'_def g'_def Collect_mono bounded_subset) |
|
8346 next |
|
8347 show "f' x \<in> U" if "x \<in> K" for x |
|
8348 proof - |
|
8349 have "U \<subseteq> S" |
|
8350 using opeU openin_imp_subset by blast |
|
8351 then have "j (f (h x)) \<in> U" |
|
8352 using f h hull_subset that by fastforce |
|
8353 then show "f' x \<in> U" |
|
8354 using \<open>K \<subseteq> S\<close> S f'_def that by auto |
|
8355 qed |
|
8356 qed |
|
8357 qed |
|
8358 qed |
|
8359 |
|
8360 |
|
8361 subsection\<open>Nullhomotopic mappings\<close> |
|
8362 |
|
8363 text\<open> A mapping out of a sphere is nullhomotopic iff it extends to the ball. |
|
8364 This even works out in the degenerate cases when the radius is \<open>\<le>\<close> 0, and |
|
8365 we also don't need to explicitly assume continuity since it's already implicit |
|
8366 in both sides of the equivalence.\<close> |
|
8367 |
|
8368 lemma nullhomotopic_from_lemma: |
|
8369 assumes contg: "continuous_on (cball a r - {a}) g" |
|
8370 and fa: "\<And>e. 0 < e |
|
8371 \<Longrightarrow> \<exists>d. 0 < d \<and> (\<forall>x. x \<noteq> a \<and> norm(x - a) < d \<longrightarrow> norm(g x - f a) < e)" |
|
8372 and r: "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> f x = g x" |
|
8373 shows "continuous_on (cball a r) f" |
|
8374 proof (clarsimp simp: continuous_on_eq_continuous_within Ball_def) |
|
8375 fix x |
|
8376 assume x: "dist a x \<le> r" |
|
8377 show "continuous (at x within cball a r) f" |
|
8378 proof (cases "x=a") |
|
8379 case True |
|
8380 then show ?thesis |
|
8381 by (metis continuous_within_eps_delta fa dist_norm dist_self r) |
|
8382 next |
|
8383 case False |
|
8384 show ?thesis |
|
8385 proof (rule continuous_transform_within [where f=g and d = "norm(x-a)"]) |
|
8386 have "\<exists>d>0. \<forall>x'\<in>cball a r. |
|
8387 dist x' x < d \<longrightarrow> dist (g x') (g x) < e" if "e>0" for e |
|
8388 proof - |
|
8389 obtain d where "d > 0" |
|
8390 and d: "\<And>x'. \<lbrakk>dist x' a \<le> r; x' \<noteq> a; dist x' x < d\<rbrakk> \<Longrightarrow> |
|
8391 dist (g x') (g x) < e" |
|
8392 using contg False x \<open>e>0\<close> |
|
8393 unfolding continuous_on_iff by (fastforce simp add: dist_commute intro: that) |
|
8394 show ?thesis |
|
8395 using \<open>d > 0\<close> \<open>x \<noteq> a\<close> |
|
8396 by (rule_tac x="min d (norm(x - a))" in exI) |
|
8397 (auto simp: dist_commute dist_norm [symmetric] intro!: d) |
|
8398 qed |
|
8399 then show "continuous (at x within cball a r) g" |
|
8400 using contg False by (auto simp: continuous_within_eps_delta) |
|
8401 show "0 < norm (x - a)" |
|
8402 using False by force |
|
8403 show "x \<in> cball a r" |
|
8404 by (simp add: x) |
|
8405 show "\<And>x'. \<lbrakk>x' \<in> cball a r; dist x' x < norm (x - a)\<rbrakk> |
|
8406 \<Longrightarrow> g x' = f x'" |
|
8407 by (metis dist_commute dist_norm less_le r) |
|
8408 qed |
|
8409 qed |
|
8410 qed |
|
8411 |
|
8412 proposition nullhomotopic_from_sphere_extension: |
|
8413 fixes f :: "'M::euclidean_space \<Rightarrow> 'a::real_normed_vector" |
|
8414 shows "(\<exists>c. homotopic_with (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)) \<longleftrightarrow> |
|
8415 (\<exists>g. continuous_on (cball a r) g \<and> g ` (cball a r) \<subseteq> S \<and> |
|
8416 (\<forall>x \<in> sphere a r. g x = f x))" |
|
8417 (is "?lhs = ?rhs") |
|
8418 proof (cases r "0::real" rule: linorder_cases) |
|
8419 case equal |
|
8420 then show ?thesis |
|
8421 apply (auto simp: homotopic_with) |
|
8422 apply (rule_tac x="\<lambda>x. h (0, a)" in exI) |
|
8423 apply (fastforce simp add:) |
|
8424 using continuous_on_const by blast |
|
8425 next |
|
8426 case greater |
|
8427 let ?P = "continuous_on {x. norm(x - a) = r} f \<and> f ` {x. norm(x - a) = r} \<subseteq> S" |
|
8428 have ?P if ?lhs using that |
|
8429 proof |
|
8430 fix c |
|
8431 assume c: "homotopic_with (\<lambda>x. True) (sphere a r) S f (\<lambda>x. c)" |
|
8432 then have contf: "continuous_on (sphere a r) f" and fim: "f ` sphere a r \<subseteq> S" |
|
8433 by (auto simp: homotopic_with_imp_subset1 homotopic_with_imp_continuous) |
|
8434 show ?P |
|
8435 using contf fim by (auto simp: sphere_def dist_norm norm_minus_commute) |
|
8436 qed |
|
8437 moreover have ?P if ?rhs using that |
|
8438 proof |
|
8439 fix g |
|
8440 assume g: "continuous_on (cball a r) g \<and> g ` cball a r \<subseteq> S \<and> (\<forall>xa\<in>sphere a r. g xa = f xa)" |
|
8441 then |
|
8442 show ?P |
|
8443 apply (safe elim!: continuous_on_eq [OF continuous_on_subset]) |
|
8444 apply (auto simp: dist_norm norm_minus_commute) |
|
8445 by (metis dist_norm image_subset_iff mem_sphere norm_minus_commute sphere_cball subsetCE) |
|
8446 qed |
|
8447 moreover have ?thesis if ?P |
|
8448 proof |
|
8449 assume ?lhs |
|
8450 then obtain c where "homotopic_with (\<lambda>x. True) (sphere a r) S (\<lambda>x. c) f" |
|
8451 using homotopic_with_sym by blast |
|
8452 then obtain h where conth: "continuous_on ({0..1::real} \<times> sphere a r) h" |
|
8453 and him: "h ` ({0..1} \<times> sphere a r) \<subseteq> S" |
|
8454 and h: "\<And>x. h(0, x) = c" "\<And>x. h(1, x) = f x" |
|
8455 by (auto simp: homotopic_with_def) |
|
8456 obtain b1::'M where "b1 \<in> Basis" |
|
8457 using SOME_Basis by auto |
|
8458 have "c \<in> S" |
|
8459 apply (rule him [THEN subsetD]) |
|
8460 apply (rule_tac x = "(0, a + r *\<^sub>R b1)" in image_eqI) |
|
8461 using h greater \<open>b1 \<in> Basis\<close> |
|
8462 apply (auto simp: dist_norm) |
|
8463 done |
|
8464 have uconth: "uniformly_continuous_on ({0..1::real} \<times> (sphere a r)) h" |
|
8465 by (force intro: compact_Times conth compact_uniformly_continuous) |
|
8466 let ?g = "\<lambda>x. h (norm (x - a)/r, |
|
8467 a + (if x = a then r *\<^sub>R b1 else (r / norm(x - a)) *\<^sub>R (x - a)))" |
|
8468 let ?g' = "\<lambda>x. h (norm (x - a)/r, a + (r / norm(x - a)) *\<^sub>R (x - a))" |
|
8469 show ?rhs |
|
8470 proof (intro exI conjI) |
|
8471 have "continuous_on (cball a r - {a}) ?g'" |
|
8472 apply (rule continuous_on_compose2 [OF conth]) |
|
8473 apply (intro continuous_intros) |
|
8474 using greater apply (auto simp: dist_norm norm_minus_commute) |
|
8475 done |
|
8476 then show "continuous_on (cball a r) ?g" |
|
8477 proof (rule nullhomotopic_from_lemma) |
|
8478 show "\<exists>d>0. \<forall>x. x \<noteq> a \<and> norm (x - a) < d \<longrightarrow> norm (?g' x - ?g a) < e" if "0 < e" for e |
|
8479 proof - |
|
8480 obtain d where "0 < d" |
|
8481 and d: "\<And>x x'. \<lbrakk>x \<in> {0..1} \<times> sphere a r; x' \<in> {0..1} \<times> sphere a r; dist x' x < d\<rbrakk> |
|
8482 \<Longrightarrow> dist (h x') (h x) < e" |
|
8483 using uniformly_continuous_onE [OF uconth \<open>0 < e\<close>] by auto |
|
8484 have *: "norm (h (norm (x - a) / r, |
|
8485 a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) < e" |
|
8486 if "x \<noteq> a" "norm (x - a) < r" "norm (x - a) < d * r" for x |
|
8487 proof - |
|
8488 have "norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + r *\<^sub>R b1)) = |
|
8489 norm (h (norm (x - a) / r, a + (r / norm (x - a)) *\<^sub>R (x - a)) - h (0, a + (r / norm (x - a)) *\<^sub>R (x - a)))" |
|
8490 by (simp add: h) |
|
8491 also have "\<dots> < e" |
|
8492 apply (rule d [unfolded dist_norm]) |
|
8493 using greater \<open>0 < d\<close> \<open>b1 \<in> Basis\<close> that |
|
8494 by (auto simp: dist_norm divide_simps) |
|
8495 finally show ?thesis . |
|
8496 qed |
|
8497 show ?thesis |
|
8498 apply (rule_tac x = "min r (d * r)" in exI) |
|
8499 using greater \<open>0 < d\<close> by (auto simp: *) |
|
8500 qed |
|
8501 show "\<And>x. x \<in> cball a r \<and> x \<noteq> a \<Longrightarrow> ?g x = ?g' x" |
|
8502 by auto |
|
8503 qed |
|
8504 next |
|
8505 show "?g ` cball a r \<subseteq> S" |
|
8506 using greater him \<open>c \<in> S\<close> |
|
8507 by (force simp: h dist_norm norm_minus_commute) |
|
8508 next |
|
8509 show "\<forall>x\<in>sphere a r. ?g x = f x" |
|
8510 using greater by (auto simp: h dist_norm norm_minus_commute) |
|
8511 qed |
|
8512 next |
|
8513 assume ?rhs |
|
8514 then obtain g where contg: "continuous_on (cball a r) g" |
|
8515 and gim: "g ` cball a r \<subseteq> S" |
|
8516 and gf: "\<forall>x \<in> sphere a r. g x = f x" |
|
8517 by auto |
|
8518 let ?h = "\<lambda>y. g (a + (fst y) *\<^sub>R (snd y - a))" |
|
8519 have "continuous_on ({0..1} \<times> sphere a r) ?h" |
|
8520 apply (rule continuous_on_compose2 [OF contg]) |
|
8521 apply (intro continuous_intros) |
|
8522 apply (auto simp: dist_norm norm_minus_commute mult_left_le_one_le) |
|
8523 done |
|
8524 moreover |
|
8525 have "?h ` ({0..1} \<times> sphere a r) \<subseteq> S" |
|
8526 by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gim [THEN subsetD]) |
|
8527 moreover |
|
8528 have "\<forall>x\<in>sphere a r. ?h (0, x) = g a" "\<forall>x\<in>sphere a r. ?h (1, x) = f x" |
|
8529 by (auto simp: dist_norm norm_minus_commute mult_left_le_one_le gf) |
|
8530 ultimately |
|
8531 show ?lhs |
|
8532 apply (subst homotopic_with_sym) |
|
8533 apply (rule_tac x="g a" in exI) |
|
8534 apply (auto simp: homotopic_with) |
|
8535 done |
|
8536 qed |
|
8537 ultimately |
|
8538 show ?thesis by meson |
|
8539 qed simp |
|
8540 |
|
8541 end |
|