src/HOL/Analysis/Continuous_Extension.thy
author wenzelm
Sun, 31 Jan 2021 20:44:42 +0100
changeset 73216 60c32e2c5577
parent 71255 4258ee13f5d4
child 76987 4c275405faae
permissions -rw-r--r--
clarified messages;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63992
3aa9837d05c7 updated headers;
wenzelm
parents: 63627
diff changeset
     1
(*  Title:      HOL/Analysis/Continuous_Extension.thy
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
    Authors:    LC Paulson, based on material from HOL Light
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     3
*)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     4
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69286
diff changeset
     5
section \<open>Continuous Extensions of Functions\<close>
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63593
diff changeset
     7
theory Continuous_Extension
66289
2562f151541c Divided Convex_Euclidean_Space.thy in half, creating new theory Starlike
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
     8
imports Starlike
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
begin
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    11
subsection\<open>Partitions of unity subordinate to locally finite open coverings\<close>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    12
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
text\<open>A difference from HOL Light: all summations over infinite sets equal zero,
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
   so the "support" must be made explicit in the summation below!\<close>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68224
diff changeset
    16
proposition subordinate_partition_of_unity:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    17
  fixes S :: "'a::metric_space set"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
  assumes "S \<subseteq> \<Union>\<C>" and opC: "\<And>T. T \<in> \<C> \<Longrightarrow> open T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
      and fin: "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. U \<inter> V \<noteq> {}}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
  obtains F :: "['a set, 'a] \<Rightarrow> real"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
    where "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x \<in> S. 0 \<le> F U x)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    22
      and "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    23
      and "\<And>x. x \<in> S \<Longrightarrow> supp_sum (\<lambda>W. F W x) \<C> = 1"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
      and "\<And>x. x \<in> S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}"
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68224
diff changeset
    25
proof (cases "\<exists>W. W \<in> \<C> \<and> S \<subseteq> W")
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    26
  case True
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
    then obtain W where "W \<in> \<C>" "S \<subseteq> W" by metis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    28
    then show ?thesis
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    29
      by (rule_tac F = "\<lambda>V x. if V = W then 1 else 0" in that) (auto simp: supp_sum_def support_on_def)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    31
  case False
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    32
    have nonneg: "0 \<le> supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>" for x
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    33
      by (simp add: supp_sum_def sum_nonneg)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    34
    have sd_pos: "0 < setdist {x} (S - V)" if "V \<in> \<C>" "x \<in> S" "x \<in> V" for V x
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    35
    proof -
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
    36
      have "closedin (top_of_set S) (S - V)"
69286
nipkow
parents: 68607
diff changeset
    37
        by (simp add: Diff_Diff_Int closedin_def opC openin_open_Int \<open>V \<in> \<C>\<close>)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    38
      with that False  setdist_pos_le [of "{x}" "S - V"]
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    39
      show ?thesis
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    40
        using setdist_gt_0_closedin by fastforce
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
    qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    42
    have ss_pos: "0 < supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>" if "x \<in> S" for x
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
    proof -
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
      obtain U where "U \<in> \<C>" "x \<in> U" using \<open>x \<in> S\<close> \<open>S \<subseteq> \<Union>\<C>\<close>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
        by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
      obtain V where "open V" "x \<in> V" "finite {U \<in> \<C>. U \<inter> V \<noteq> {}}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
        using \<open>x \<in> S\<close> fin by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
      then have *: "finite {A \<in> \<C>. \<not> S \<subseteq> A \<and> x \<notin> closure (S - A)}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
        using closure_def that by (blast intro: rev_finite_subset)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
      have "x \<notin> closure (S - U)"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    51
        using \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> opC open_Int_closure_eq_empty by fastforce
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
      then show ?thesis
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    53
        apply (simp add: setdist_eq_0_sing_1 supp_sum_def support_on_def)
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    54
        apply (rule ordered_comm_monoid_add_class.sum_pos2 [OF *, of U])
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
        using \<open>U \<in> \<C>\<close> \<open>x \<in> U\<close> False
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    56
        apply (auto simp: sd_pos that)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
        done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
    qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
    define F where
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
    60
      "F \<equiv> \<lambda>W x. if x \<in> S then setdist {x} (S - W) / supp_sum (\<lambda>V. setdist {x} (S - V)) \<C> else 0"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
    show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
    proof (rule_tac F = F in that)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
      have "continuous_on S (F U)" if "U \<in> \<C>" for U
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
      proof -
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    65
        have *: "continuous_on S (\<lambda>x. supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>)"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
        proof (clarsimp simp add: continuous_on_eq_continuous_within)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
          fix x assume "x \<in> S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
          then obtain X where "open X" and x: "x \<in> S \<inter> X" and finX: "finite {U \<in> \<C>. U \<inter> X \<noteq> {}}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
            using assms by blast
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
    70
          then have OSX: "openin (top_of_set S) (S \<inter> X)" by blast
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
          have sumeq: "\<And>x. x \<in> S \<inter> X \<Longrightarrow>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
                     (\<Sum>V | V \<in> \<C> \<and> V \<inter> X \<noteq> {}. setdist {x} (S - V))
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    73
                     = supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>"
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    74
            apply (simp add: supp_sum_def)
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    75
            apply (rule sum.mono_neutral_right [OF finX])
63593
bbcb05504fdc HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
hoelzl
parents: 63306
diff changeset
    76
            apply (auto simp: setdist_eq_0_sing_1 support_on_def subset_iff)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
            apply (meson DiffI closure_subset disjoint_iff_not_equal subsetCE)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
            done
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    79
          show "continuous (at x within S) (\<lambda>x. supp_sum (\<lambda>V. setdist {x} (S - V)) \<C>)"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
            apply (rule continuous_transform_within_openin
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
    81
                     [where f = "\<lambda>x. (sum (\<lambda>V. setdist {x} (S - V)) {V \<in> \<C>. V \<inter> X \<noteq> {}})"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
                        and S ="S \<inter> X"])
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
            apply (rule continuous_intros continuous_at_setdist continuous_at_imp_continuous_at_within OSX x)+
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
            apply (simp add: sumeq)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
            done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
        qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
        show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
          apply (simp add: F_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
          apply (rule continuous_intros *)+
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
          using ss_pos apply force
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
          done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
      qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
      moreover have "\<lbrakk>U \<in> \<C>; x \<in> S\<rbrakk> \<Longrightarrow> 0 \<le> F U x" for U x
71172
nipkow
parents: 70817
diff changeset
    94
        using nonneg [of x] by (simp add: F_def field_split_simps)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
      ultimately show "\<And>U. U \<in> \<C> \<Longrightarrow> continuous_on S (F U) \<and> (\<forall>x\<in>S. 0 \<le> F U x)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
        by metis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
    next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
      show "\<And>x U. \<lbrakk>U \<in> \<C>; x \<in> S; x \<notin> U\<rbrakk> \<Longrightarrow> F U x = 0"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
        by (simp add: setdist_eq_0_sing_1 closure_def F_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
    next
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   101
      show "supp_sum (\<lambda>W. F W x) \<C> = 1" if "x \<in> S" for x
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
        using that ss_pos [OF that]
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70136
diff changeset
   103
        by (simp add: F_def field_split_simps supp_sum_divide_distrib [symmetric])
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
    next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
      show "\<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. F U x \<noteq> 0}" if "x \<in> S" for x
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
        using fin [OF that] that
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
        by (fastforce simp: setdist_eq_0_sing_1 closure_def F_def elim!: rev_finite_subset)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
    qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69286
diff changeset
   112
subsection\<open>Urysohn's Lemma for Euclidean Spaces\<close>
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69286
diff changeset
   113
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69286
diff changeset
   114
text \<open>For Euclidean spaces the proof is easy using distances.\<close>
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
lemma Urysohn_both_ne:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   117
  assumes US: "closedin (top_of_set U) S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   118
      and UT: "closedin (top_of_set U) T"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
      and "S \<inter> T = {}" "S \<noteq> {}" "T \<noteq> {}" "a \<noteq> b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
    where "continuous_on U f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
          "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
          "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
proof -
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  have S0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} S = 0 \<longleftrightarrow> x \<in> S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
    using \<open>S \<noteq> {}\<close>  US setdist_eq_0_closedin  by auto
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  have T0: "\<And>x. x \<in> U \<Longrightarrow> setdist {x} T = 0 \<longleftrightarrow> x \<in> T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
    using \<open>T \<noteq> {}\<close>  UT setdist_eq_0_closedin  by auto
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
  have sdpos: "0 < setdist {x} S + setdist {x} T" if "x \<in> U" for x
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
  proof -
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69286
diff changeset
   132
    have "\<not> (setdist {x} S = 0 \<and> setdist {x} T = 0)"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
      using assms by (metis IntI empty_iff setdist_eq_0_closedin that)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
    then show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
      by (metis add.left_neutral add.right_neutral add_pos_pos linorder_neqE_linordered_idom not_le setdist_pos_le)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
  define f where "f \<equiv> \<lambda>x. a + (setdist {x} S / (setdist {x} S + setdist {x} T)) *\<^sub>R (b - a)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
  show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
  proof (rule_tac f = f in that)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
    show "continuous_on U f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
      using sdpos unfolding f_def
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
      by (intro continuous_intros | force)+
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
    show "f x \<in> closed_segment a b" if "x \<in> U" for x
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
        unfolding f_def
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
      apply (simp add: closed_segment_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
      apply (rule_tac x="(setdist {x} S / (setdist {x} S + setdist {x} T))" in exI)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
      using sdpos that apply (simp add: algebra_simps)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
      done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
    show "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
      using S0 \<open>a \<noteq> b\<close> f_def sdpos by force
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
    show "(f x = b \<longleftrightarrow> x \<in> T)" if "x \<in> U" for x
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
    proof -
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
      have "f x = b \<longleftrightarrow> (setdist {x} S / (setdist {x} S + setdist {x} T)) = 1"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
        unfolding f_def
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
        apply (rule iffI)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
         apply (metis  \<open>a \<noteq> b\<close> add_diff_cancel_left' eq_iff_diff_eq_0 pth_1 real_vector.scale_right_imp_eq, force)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
        done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
      also have "...  \<longleftrightarrow> setdist {x} T = 0 \<and> setdist {x} S \<noteq> 0"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
        using sdpos that
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70136
diff changeset
   160
        by (simp add: field_split_simps) linarith
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
      also have "...  \<longleftrightarrow> x \<in> T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
        using \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>S \<inter> T = {}\<close> that
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
        by (force simp: S0 T0)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
      finally show ?thesis .
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
    qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
  qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
proposition Urysohn_local_strong:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   170
  assumes US: "closedin (top_of_set U) S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   171
      and UT: "closedin (top_of_set U) T"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
      and "S \<inter> T = {}" "a \<noteq> b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
    where "continuous_on U f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
          "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
          "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
proof (cases "S = {}")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
  case True show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
  proof (cases "T = {}")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
    case True show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
    proof (rule_tac f = "\<lambda>x. midpoint a b" in that)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
      show "continuous_on U (\<lambda>x. midpoint a b)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
        by (intro continuous_intros)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
      show "midpoint a b \<in> closed_segment a b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
        using csegment_midpoint_subset by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
      show "(midpoint a b = a) = (x \<in> S)" for x
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
        using \<open>S = {}\<close> \<open>a \<noteq> b\<close> by simp
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
      show "(midpoint a b = b) = (x \<in> T)" for x
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
        using \<open>T = {}\<close> \<open>a \<noteq> b\<close> by simp
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
    qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
  next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
    case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
    show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
    proof (cases "T = U")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
      case True with \<open>S = {}\<close> \<open>a \<noteq> b\<close> show ?thesis
71172
nipkow
parents: 70817
diff changeset
   197
        by (rule_tac f = "\<lambda>x. b" in that) (auto)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
    next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
      case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
      with UT closedin_subset obtain c where c: "c \<in> U" "c \<notin> T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
        by fastforce
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
      obtain f where f: "continuous_on U f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
                "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment (midpoint a b) b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
                "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
                "\<And>x. x \<in> U \<Longrightarrow> (f x = b \<longleftrightarrow> x \<in> T)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
        apply (rule Urysohn_both_ne [of U "{c}" T "midpoint a b" "b"])
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
        using c \<open>T \<noteq> {}\<close> assms apply simp_all
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
        done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
      show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
        apply (rule_tac f=f in that)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
        using \<open>S = {}\<close> \<open>T \<noteq> {}\<close> f csegment_midpoint_subset notin_segment_midpoint [OF \<open>a \<noteq> b\<close>]
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
        apply force+
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
        done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
    qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
  qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
  case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
  show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
  proof (cases "T = {}")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
    case True show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
    proof (cases "S = U")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
      case True with \<open>T = {}\<close> \<open>a \<noteq> b\<close> show ?thesis
71172
nipkow
parents: 70817
diff changeset
   223
        by (rule_tac f = "\<lambda>x. a" in that) (auto)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
    next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
      case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
      with US closedin_subset obtain c where c: "c \<in> U" "c \<notin> S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
        by fastforce
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
      obtain f where f: "continuous_on U f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
                "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a (midpoint a b)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
                "\<And>x. x \<in> U \<Longrightarrow> (f x = midpoint a b \<longleftrightarrow> x = c)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
                "\<And>x. x \<in> U \<Longrightarrow> (f x = a \<longleftrightarrow> x \<in> S)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
        apply (rule Urysohn_both_ne [of U S "{c}" a "midpoint a b"])
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
        using c \<open>S \<noteq> {}\<close> assms apply simp_all
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
        apply (metis midpoint_eq_endpoint)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
        done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
      show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
        apply (rule_tac f=f in that)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
        using \<open>S \<noteq> {}\<close> \<open>T = {}\<close> f  \<open>a \<noteq> b\<close>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
        apply simp_all
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
        apply (metis (no_types) closed_segment_commute csegment_midpoint_subset midpoint_sym subset_iff)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
        apply (metis closed_segment_commute midpoint_sym notin_segment_midpoint)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
        done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
    qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
  next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
    case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
    show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
      using Urysohn_both_ne [OF US UT \<open>S \<inter> T = {}\<close> \<open>S \<noteq> {}\<close> \<open>T \<noteq> {}\<close> \<open>a \<noteq> b\<close>] that
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   248
      by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   249
  qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
lemma Urysohn_local:
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   253
  assumes US: "closedin (top_of_set U) S"
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   254
      and UT: "closedin (top_of_set U) T"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
      and "S \<inter> T = {}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
    where "continuous_on U f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
          "\<And>x. x \<in> U \<Longrightarrow> f x \<in> closed_segment a b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
          "\<And>x. x \<in> S \<Longrightarrow> f x = a"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
          "\<And>x. x \<in> T \<Longrightarrow> f x = b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
proof (cases "a = b")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
  case True then show ?thesis
71172
nipkow
parents: 70817
diff changeset
   263
    by (rule_tac f = "\<lambda>x. b" in that) (auto)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
  then show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
    apply (rule Urysohn_local_strong [OF assms])
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
    apply (erule that, assumption)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
    apply (meson US closedin_singleton closedin_trans)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
    apply (meson UT closedin_singleton closedin_trans)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
    done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
lemma Urysohn_strong:
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
  assumes US: "closed S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
      and UT: "closed T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
      and "S \<inter> T = {}" "a \<noteq> b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
    where "continuous_on UNIV f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
          "\<And>x. f x \<in> closed_segment a b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
          "\<And>x. f x = a \<longleftrightarrow> x \<in> S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
          "\<And>x. f x = b \<longleftrightarrow> x \<in> T"
66884
c2128ab11f61 Switching to inverse image and constant_on, plus some new material
paulson <lp15@cam.ac.uk>
parents: 66289
diff changeset
   283
using assms by (auto intro: Urysohn_local_strong [of UNIV S T])
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68224
diff changeset
   285
proposition Urysohn:
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
  assumes US: "closed S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
      and UT: "closed T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
      and "S \<inter> T = {}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
  obtains f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
    where "continuous_on UNIV f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
          "\<And>x. f x \<in> closed_segment a b"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
          "\<And>x. x \<in> S \<Longrightarrow> f x = a"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
          "\<And>x. x \<in> T \<Longrightarrow> f x = b"
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68224
diff changeset
   294
  using assms by (auto intro: Urysohn_local [of UNIV S T a b])
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 69508
diff changeset
   297
subsection\<open>Dugundji's Extension Theorem and Tietze Variants\<close>
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
69518
bf88364c9e94 tuned headers etc, added bib-file
nipkow
parents: 69508
diff changeset
   299
text \<open>See \cite{dugundji}.\<close>
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
71255
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   301
lemma convex_supp_sum:
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   302
  assumes "convex S" and 1: "supp_sum u I = 1"
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   303
      and "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> u i \<and> (u i = 0 \<or> f i \<in> S)"
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   304
    shows "supp_sum (\<lambda>i. u i *\<^sub>R f i) I \<in> S"
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   305
proof -
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   306
  have fin: "finite {i \<in> I. u i \<noteq> 0}"
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   307
    using 1 sum.infinite by (force simp: supp_sum_def support_on_def)
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   308
  then have "supp_sum (\<lambda>i. u i *\<^sub>R f i) I = sum (\<lambda>i. u i *\<^sub>R f i) {i \<in> I. u i \<noteq> 0}"
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   309
    by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def)
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   310
  also have "... \<in> S"
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   311
    using 1 assms by (force simp: supp_sum_def support_on_def intro: convex_sum [OF fin \<open>convex S\<close>])
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   312
  finally show ?thesis .
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   313
qed
4258ee13f5d4 moved lemmas
nipkow
parents: 71172
diff changeset
   314
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68224
diff changeset
   315
theorem Dugundji:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   316
  fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::real_inner"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
  assumes "convex C" "C \<noteq> {}"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   318
      and cloin: "closedin (top_of_set U) S"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
      and contf: "continuous_on S f" and "f ` S \<subseteq> C"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
  obtains g where "continuous_on U g" "g ` U \<subseteq> C"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
                  "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68224
diff changeset
   322
proof (cases "S = {}")
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
  case True then show thesis
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 66884
diff changeset
   324
    apply (rule_tac g="\<lambda>x. SOME y. y \<in> C" in that)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
      apply (rule continuous_intros)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
     apply (meson all_not_in_conv \<open>C \<noteq> {}\<close> image_subsetI someI_ex, simp)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
    done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
  case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
  then have sd_pos: "\<And>x. \<lbrakk>x \<in> U; x \<notin> S\<rbrakk> \<Longrightarrow> 0 < setdist {x} S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
    using setdist_eq_0_closedin [OF cloin] le_less setdist_pos_le by fastforce
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
  define \<B> where "\<B> = {ball x (setdist {x} S / 2) |x. x \<in> U - S}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
  have [simp]: "\<And>T. T \<in> \<B> \<Longrightarrow> open T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
    by (auto simp: \<B>_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
  have USS: "U - S \<subseteq> \<Union>\<B>"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
    by (auto simp: sd_pos \<B>_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
  obtain \<C> where USsub: "U - S \<subseteq> \<Union>\<C>"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
       and nbrhd: "\<And>U. U \<in> \<C> \<Longrightarrow> open U \<and> (\<exists>T. T \<in> \<B> \<and> U \<subseteq> T)"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   339
       and fin: "\<And>x. x \<in> U - S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U. U \<in> \<C> \<and> U \<inter> V \<noteq> {}}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   340
    by (rule paracompact [OF USS]) auto
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
  have "\<exists>v a. v \<in> U \<and> v \<notin> S \<and> a \<in> S \<and>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   342
              T \<subseteq> ball v (setdist {v} S / 2) \<and>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   343
              dist v a \<le> 2 * setdist {v} S" if "T \<in> \<C>" for T
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
  proof -
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   345
    obtain v where v: "T \<subseteq> ball v (setdist {v} S / 2)" "v \<in> U" "v \<notin> S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   346
      using \<open>T \<in> \<C>\<close> nbrhd by (force simp: \<B>_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
    then obtain a where "a \<in> S" "dist v a < 2 * setdist {v} S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
      using setdist_ltE [of "{v}" S "2 * setdist {v} S"]
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
      using False sd_pos by force
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
    with v show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
      apply (rule_tac x=v in exI)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
      apply (rule_tac x=a in exI, auto)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
      done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
  qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
  then obtain \<V> \<A> where
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
    VA: "\<And>T. T \<in> \<C> \<Longrightarrow> \<V> T \<in> U \<and> \<V> T \<notin> S \<and> \<A> T \<in> S \<and>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
              T \<subseteq> ball (\<V> T) (setdist {\<V> T} S / 2) \<and>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
              dist (\<V> T) (\<A> T) \<le> 2 * setdist {\<V> T} S"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
    by metis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
  have sdle: "setdist {\<V> T} S \<le> 2 * setdist {v} S" if "T \<in> \<C>" "v \<in> T" for T v
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
    using setdist_Lipschitz [of "\<V> T" S v] VA [OF \<open>T \<in> \<C>\<close>] \<open>v \<in> T\<close> by auto
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
  have d6: "dist a (\<A> T) \<le> 6 * dist a v" if "T \<in> \<C>" "v \<in> T" "a \<in> S" for T v a
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
  proof -
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
    have "dist (\<V> T) v < setdist {\<V> T} S / 2"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
      using that VA mem_ball by blast
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   366
    also have "\<dots> \<le> setdist {v} S"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
      using sdle [OF \<open>T \<in> \<C>\<close> \<open>v \<in> T\<close>] by simp
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
    also have vS: "setdist {v} S \<le> dist a v"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
      by (simp add: setdist_le_dist setdist_sym \<open>a \<in> S\<close>)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
    finally have VTV: "dist (\<V> T) v < dist a v" .
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
    have VTS: "setdist {\<V> T} S \<le> 2 * dist a v"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
      using sdle that vS by force
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
    have "dist a (\<A> T) \<le> dist a v + dist v (\<V> T) + dist (\<V> T) (\<A> T)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
      by (metis add.commute add_le_cancel_left dist_commute dist_triangle2 dist_triangle_le)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   375
    also have "\<dots> \<le> dist a v + dist a v + dist (\<V> T) (\<A> T)"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
      using VTV by (simp add: dist_commute)
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   377
    also have "\<dots> \<le> 2 * dist a v + 2 * setdist {\<V> T} S"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
      using VA [OF \<open>T \<in> \<C>\<close>] by auto
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
    finally show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
      using VTS by linarith
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
  qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
  obtain H :: "['a set, 'a] \<Rightarrow> real"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
    where Hcont: "\<And>Z. Z \<in> \<C> \<Longrightarrow> continuous_on (U-S) (H Z)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
      and Hge0: "\<And>Z x. \<lbrakk>Z \<in> \<C>; x \<in> U-S\<rbrakk> \<Longrightarrow> 0 \<le> H Z x"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
      and Heq0: "\<And>x Z. \<lbrakk>Z \<in> \<C>; x \<in> U-S; x \<notin> Z\<rbrakk> \<Longrightarrow> H Z x = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   386
      and H1: "\<And>x. x \<in> U-S \<Longrightarrow> supp_sum (\<lambda>W. H W x) \<C> = 1"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   387
      and Hfin: "\<And>x. x \<in> U-S \<Longrightarrow> \<exists>V. open V \<and> x \<in> V \<and> finite {U \<in> \<C>. \<exists>x\<in>V. H U x \<noteq> 0}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   388
    apply (rule subordinate_partition_of_unity [OF USsub _ fin])
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
    using nbrhd by auto
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   390
  define g where "g \<equiv> \<lambda>x. if x \<in> S then f x else supp_sum (\<lambda>T. H T x *\<^sub>R f(\<A> T)) \<C>"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
  show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
  proof (rule that)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
    show "continuous_on U g"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
    proof (clarsimp simp: continuous_on_eq_continuous_within)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
      fix a assume "a \<in> U"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
      show "continuous (at a within U) g"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
      proof (cases "a \<in> S")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
        case True show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
        proof (clarsimp simp add: continuous_within_topological)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
          fix W
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
          assume "open W" "g a \<in> W"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
          then obtain e where "0 < e" and e: "ball (f a) e \<subseteq> W"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
            using openE True g_def by auto
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
          have "continuous (at a within S) f"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
            using True contf continuous_on_eq_continuous_within by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
          then obtain d where "0 < d"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   407
                        and d: "\<And>x. \<lbrakk>x \<in> S; dist x a < d\<rbrakk> \<Longrightarrow> dist (f x) (f a) < e"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   408
            using continuous_within_eps_delta \<open>0 < e\<close> by force
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   409
          have "g y \<in> ball (f a) e" if "y \<in> U" and y: "y \<in> ball a (d / 6)" for y
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
          proof (cases "y \<in> S")
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
            case True
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
            then have "dist (f a) (f y) < e"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
              by (metis ball_divide_subset_numeral dist_commute in_mono mem_ball y d)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
            then show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
              by (simp add: True g_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
          next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
            case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
            have *: "dist (f (\<A> T)) (f a) < e" if "T \<in> \<C>" "H T y \<noteq> 0" for T
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
            proof -
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
              have "y \<in> T"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
                using Heq0 that False \<open>y \<in> U\<close> by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
              have "dist (\<A> T) a < d"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
                using d6 [OF \<open>T \<in> \<C>\<close> \<open>y \<in> T\<close> \<open>a \<in> S\<close>] y
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
                by (simp add: dist_commute mult.commute)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
              then show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
                using VA [OF \<open>T \<in> \<C>\<close>] by (auto simp: d)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
            qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   428
            have "supp_sum (\<lambda>T. H T y *\<^sub>R f (\<A> T)) \<C> \<in> ball (f a) e"
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   429
              apply (rule convex_supp_sum [OF convex_ball])
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
              apply (simp_all add: False H1 Hge0 \<open>y \<in> U\<close>)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
              by (metis dist_commute *)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
            then show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
              by (simp add: False g_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
          qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
          then show "\<exists>A. open A \<and> a \<in> A \<and> (\<forall>y\<in>U. y \<in> A \<longrightarrow> g y \<in> W)"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
            apply (rule_tac x = "ball a (d / 6)" in exI)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
            using e \<open>0 < d\<close> by fastforce
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
        qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
      next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
        case False
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
        obtain N where N: "open N" "a \<in> N"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
                   and finN: "finite {U \<in> \<C>. \<exists>a\<in>N. H U a \<noteq> 0}"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
          using Hfin False \<open>a \<in> U\<close> by auto
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   444
        have oUS: "openin (top_of_set U) (U - S)"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
          using cloin by (simp add: openin_diff)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
        have HcontU: "continuous (at a within U) (H T)" if "T \<in> \<C>" for T
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
          using Hcont [OF \<open>T \<in> \<C>\<close>] False  \<open>a \<in> U\<close> \<open>T \<in> \<C>\<close>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
          apply (simp add: continuous_on_eq_continuous_within continuous_within)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
          apply (rule Lim_transform_within_set)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
          using oUS
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
            apply (force simp: eventually_at openin_contains_ball dist_commute dest!: bspec)+
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
          done
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
        show ?thesis
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
        proof (rule continuous_transform_within_openin [OF _ oUS])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   455
          show "continuous (at a within U) (\<lambda>x. supp_sum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>)"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
          proof (rule continuous_transform_within_openin)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   457
            show "continuous (at a within U)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   458
                    (\<lambda>x. \<Sum>T\<in>{U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
              by (force intro: continuous_intros HcontU)+
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
          next
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   461
            show "openin (top_of_set U) ((U - S) \<inter> N)"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
              using N oUS openin_trans by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
          next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
            show "a \<in> (U - S) \<inter> N" using False \<open>a \<in> U\<close> N by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
          next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
            show "\<And>x. x \<in> (U - S) \<inter> N \<Longrightarrow>
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
                         (\<Sum>T \<in> {U \<in> \<C>. \<exists>x\<in>N. H U x \<noteq> 0}. H T x *\<^sub>R f (\<A> T))
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   468
                         = supp_sum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C>"
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   469
              by (auto simp: supp_sum_def support_on_def
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   470
                       intro: sum.mono_neutral_right [OF finN])
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
          qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
        next
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
          show "a \<in> U - S" using False \<open>a \<in> U\<close> by blast
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
        next
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   475
          show "\<And>x. x \<in> U - S \<Longrightarrow> supp_sum (\<lambda>T. H T x *\<^sub>R f (\<A> T)) \<C> = g x"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
            by (simp add: g_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
        qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
      qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
    qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
    show "g ` U \<subseteq> C"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
      using \<open>f ` S \<subseteq> C\<close> VA
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   482
      by (fastforce simp: g_def Hge0 intro!: convex_supp_sum [OF \<open>convex C\<close>] H1)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
    show "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
      by (simp add: g_def)
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
  qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
qed
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
68607
67bb59e49834 make theorem, corollary, and proposition %important for HOL-Analysis manual
immler
parents: 68224
diff changeset
   489
corollary Tietze:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   490
  fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::real_inner"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
  assumes "continuous_on S f"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   492
    and "closedin (top_of_set U) S"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   493
    and "0 \<le> B"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   494
    and "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> B"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   496
    "\<And>x. x \<in> U \<Longrightarrow> norm(g x) \<le> B"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   497
  using assms by (auto simp: image_subset_iff intro: Dugundji [of "cball 0 B" U S f])
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   499
corollary\<^marker>\<open>tag unimportant\<close> Tietze_closed_interval:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   500
  fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::euclidean_space"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
  assumes "continuous_on S f"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   502
    and "closedin (top_of_set U) S"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   503
    and "cbox a b \<noteq> {}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   504
    and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   506
    "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   507
  apply (rule Dugundji [of "cbox a b" U S f])
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   508
  using assms by auto
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   510
corollary\<^marker>\<open>tag unimportant\<close> Tietze_closed_interval_1:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   511
  fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> real"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
  assumes "continuous_on S f"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   513
    and "closedin (top_of_set U) S"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   514
    and "a \<le> b"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   515
    and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> cbox a b"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   517
    "\<And>x. x \<in> U \<Longrightarrow> g x \<in> cbox a b"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   518
  apply (rule Dugundji [of "cbox a b" U S f])
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   519
  using assms by (auto simp: image_subset_iff)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   521
corollary\<^marker>\<open>tag unimportant\<close> Tietze_open_interval:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   522
  fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::euclidean_space"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
  assumes "continuous_on S f"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   524
    and "closedin (top_of_set U) S"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   525
    and "box a b \<noteq> {}"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   526
    and "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   528
    "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   529
  apply (rule Dugundji [of "box a b" U S f])
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   530
  using assms by auto
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   532
corollary\<^marker>\<open>tag unimportant\<close> Tietze_open_interval_1:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   533
  fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> real"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
  assumes "continuous_on S f"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   535
    and "closedin (top_of_set U) S"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   536
    and "a < b"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   537
    and no: "\<And>x. x \<in> S \<Longrightarrow> f x \<in> box a b"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   539
    "\<And>x. x \<in> U \<Longrightarrow> g x \<in> box a b"
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   540
  apply (rule Dugundji [of "box a b" U S f])
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   541
  using assms by (auto simp: image_subset_iff)
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69922
diff changeset
   543
corollary\<^marker>\<open>tag unimportant\<close> Tietze_unbounded:
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   544
  fixes f :: "'a::{metric_space,second_countable_topology} \<Rightarrow> 'b::real_inner"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
  assumes "continuous_on S f"
69922
4a9167f377b0 new material about topology, etc.; also fixes for yesterday's
paulson <lp15@cam.ac.uk>
parents: 69918
diff changeset
   546
    and "closedin (top_of_set U) S"
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
  obtains g where "continuous_on U g" "\<And>x. x \<in> S \<Longrightarrow> g x = f x"
69918
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   548
  apply (rule Dugundji [of UNIV U S f])
eddcc7c726f3 new material;' strengthened material; moved proofs out of Function_Topology in order to lessen its dependencies
paulson <lp15@cam.ac.uk>
parents: 69518
diff changeset
   549
  using assms by auto
63305
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   550
3b6975875633 Urysohn's lemma, Dugundji extension theorem and many other proofs
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   551
end